Documentation

Mathlib.GroupTheory.Perm.Fin

Permutations of Fin n #

Permutations of Fin (n + 1) are equivalent to fixing a single Fin (n + 1) and permuting the remaining with a Perm (Fin n). The fixed Fin (n + 1) is swapped with 0.

Equations
Instances For
    @[simp]
    @[simp]
    theorem Equiv.Perm.decomposeFin_symm_apply_zero {n : } (p : Fin (n + 1)) (e : Perm (Fin n)) :
    @[simp]
    theorem Equiv.Perm.decomposeFin_symm_apply_succ {n : } (e : Perm (Fin n)) (p : Fin (n + 1)) (x : Fin n) :
    (decomposeFin.symm (p, e)) x.succ = (swap 0 p) (e x).succ
    @[simp]
    theorem Equiv.Perm.decomposeFin_symm_apply_one {n : } (e : Perm (Fin (n + 1))) (p : Fin (n + 2)) :
    (decomposeFin.symm (p, e)) 1 = (swap 0 p) (e 0).succ
    @[simp]
    theorem Equiv.Perm.decomposeFin.symm_sign {n : } (p : Fin (n + 1)) (e : Perm (Fin n)) :
    sign (decomposeFin.symm (p, e)) = (if p = 0 then 1 else -1) * sign e

    The set of all permutations of Fin (n + 1) can be constructed by augmenting the set of permutations of Fin n by each element of Fin (n + 1) in turn.

    cycleRange section #

    Define the permutations Fin.cycleRange i, the cycle (0 1 2 ... i).

    @[simp]
    theorem sign_finRotate (n : ) :
    Equiv.Perm.sign (finRotate (n + 1)) = (-1) ^ n
    @[simp]
    @[simp]
    theorem cycleType_finRotate {n : } :
    (finRotate (n + 2)).cycleType = {n + 2}
    def Fin.cycleRange {n : } (i : Fin n) :

    Fin.cycleRange i is the cycle (0 1 2 ... i) leaving (i+1 ... (n-1)) unchanged.

    Equations
    Instances For
      theorem Fin.cycleRange_of_gt {n : } {i j : Fin n} (h : i < j) :
      theorem Fin.cycleRange_of_le {n : } {i j : Fin n} [NeZero n] (h : i j) :
      j.cycleRange i = if i = j then 0 else i + 1
      theorem Fin.coe_cycleRange_of_le {n : } {i j : Fin n} (h : i j) :
      (j.cycleRange i) = if i = j then 0 else i + 1
      theorem Fin.cycleRange_of_lt {n : } {i j : Fin n} [NeZero n] (h : i < j) :
      j.cycleRange i = i + 1
      theorem Fin.coe_cycleRange_of_lt {n : } {i j : Fin n} (h : i < j) :
      (j.cycleRange i) = i + 1
      theorem Fin.cycleRange_of_eq {n : } {i j : Fin n} [NeZero n] (h : i = j) :
      @[simp]
      theorem Fin.cycleRange_self {n : } [NeZero n] (i : Fin n) :
      theorem Fin.cycleRange_apply {n : } [NeZero n] (i j : Fin n) :
      i.cycleRange j = if j < i then j + 1 else if j = i then 0 else j
      @[simp]
      theorem Fin.cycleRange_zero (n : ) [NeZero n] :
      @[simp]
      @[simp]
      theorem Fin.cycleRange_mk_zero {n : } (h : 0 < n) :
      @[simp]
      theorem Fin.sign_cycleRange {n : } (i : Fin n) :
      @[simp]
      theorem Fin.succAbove_cycleRange {n : } (i j : Fin n) :
      @[simp]
      theorem Fin.cycleRange_succAbove {n : } (i : Fin (n + 1)) (j : Fin n) :
      @[simp]
      theorem Fin.cycleRange_symm_zero {n : } [NeZero n] (i : Fin n) :
      @[simp]
      theorem Fin.cycleRange_symm_succ {n : } (i : Fin (n + 1)) (j : Fin n) :
      @[simp]
      theorem Fin.insertNth_apply_cycleRange_symm {n : } {α : Type u_1} (p : Fin (n + 1)) (a : α) (x : Fin nα) (j : Fin (n + 1)) :
      p.insertNth a x ((Equiv.symm p.cycleRange) j) = cons a x j
      @[simp]
      theorem Fin.insertNth_comp_cycleRange_symm {n : } {α : Type u_1} (p : Fin (n + 1)) (a : α) (x : Fin nα) :
      @[simp]
      theorem Fin.cons_apply_cycleRange {n : } {α : Type u_1} (a : α) (x : Fin nα) (p j : Fin (n + 1)) :
      cons a x (p.cycleRange j) = p.insertNth a x j
      @[simp]
      theorem Fin.cons_comp_cycleRange {n : } {α : Type u_1} (a : α) (x : Fin nα) (p : Fin (n + 1)) :
      cons a x p.cycleRange = p.insertNth a x
      theorem Fin.isCycle_cycleRange {n : } {i : Fin n} [NeZero n] (h0 : i 0) :
      @[simp]
      theorem Fin.cycleType_cycleRange {n : } {i : Fin n} [NeZero n] (h0 : i 0) :

      The permutation cycleIcc #

      In this section, we define the permutation cycleIcc i j, which is the cycle (i i+1 .... j) leaving (0 ... i-1) and (j+1 ... n-1) unchanged when i ≤ j and returning the dummy value id when i > j. In other words, it rotates elements in [i, j] one step to the right.

      theorem Fin.instNeZeroNatHSubVal_mathlib {n : } {i : Fin n} :
      NeZero (n - i)
      def Fin.cycleIcc {n : } (i j : Fin n) :

      cycleIcc i j is the cycle (i i+1 ... j) leaving (0 ... i-1) and (j+1 ... n-1) unchanged when i < j and returning the dummy value id when i > j. In other words, it rotates elements in [i, j] one step to the right.

      Equations
      Instances For
        @[simp]
        theorem Fin.cycleIcc_def_le {n : } {i j : Fin n} (hij : i j) :
        @[simp]
        theorem Fin.cycleIcc_def_gt {n : } {i j : Fin n} (hij : i < j) :
        j.cycleIcc i = 1
        @[simp]
        theorem Fin.cycleIcc_def_gt' {n : } {i j : Fin n} (hij : ¬j i) :
        j.cycleIcc i = 1
        theorem Fin.cycleIcc_of_lt {n : } {i j k : Fin n} (h : k < i) :
        (i.cycleIcc j) k = k
        theorem Fin.cycleIcc_to_cycleRange {n : } {i j k : Fin n} (hij : i j) (kin : k Set.range (natAdd_castLEEmb )) :
        theorem Fin.cycleIcc_of_gt {n : } {i j k : Fin n} (h : j < k) :
        (i.cycleIcc j) k = k
        @[simp]
        theorem Fin.cycleIcc_of_le_of_le {n : } {i j k : Fin n} (hik : i k) (hkj : k j) [NeZero n] :
        (i.cycleIcc j) k = if k = j then i else k + 1
        theorem Fin.cycleIcc_of_ge_of_lt {n : } {i j k : Fin n} (hik : i k) (hkj : k < j) [NeZero n] :
        (i.cycleIcc j) k = k + 1
        theorem Fin.cycleIcc_of_last {n : } {i j : Fin n} (hij : i j) [NeZero n] :
        (i.cycleIcc j) j = i
        theorem Fin.cycleIcc_eq {n : } {i : Fin n} [NeZero n] :
        i.cycleIcc i = 1
        @[simp]
        theorem Fin.cycleIcc_ge {n : } {i j : Fin n} (hij : i j) [NeZero n] :
        j.cycleIcc i = 1
        theorem Fin.sign_cycleIcc_of_le {n : } {i j : Fin n} (hij : i j) :
        Equiv.Perm.sign (i.cycleIcc j) = (-1) ^ (j - i)
        theorem Fin.sign_cycleIcc_of_ge {n : } {i j : Fin n} (hij : i j) :
        theorem Fin.isCycle_cycleIcc {n : } {i j : Fin n} (hij : i < j) :
        theorem Fin.cycleType_cycleIcc_of_lt {n : } {i j : Fin n} (hij : i < j) :
        (i.cycleIcc j).cycleType = {j - i + 1}
        theorem Fin.cycleType_cycleIcc_of_ge {n : } {i j : Fin n} (hij : i j) [NeZero n] :
        theorem Fin.cycleIcc.trans {n : } {i j k : Fin n} [NeZero n] (hij : i j) (hjk : j k) :
        (i.cycleIcc j) (j.cycleIcc k) = (i.cycleIcc k)
        theorem Fin.cycleIcc.trans_left_one {n : } {i j k : Fin n} [NeZero n] (hij : i j) :
        (j.cycleIcc i) (i.cycleIcc k) = (i.cycleIcc k)
        theorem Fin.cycleIcc.trans_right_one {n : } {i j k : Fin n} [NeZero n] (hjk : j k) :
        (i.cycleIcc k) (k.cycleIcc j) = (i.cycleIcc k)
        theorem Equiv.Perm.sign_eq_prod_prod_Iio {n : } (σ : Perm (Fin n)) :
        sign σ = j : Fin n, iFinset.Iio j, if σ i < σ j then 1 else -1
        theorem Equiv.Perm.sign_eq_prod_prod_Ioi {n : } (σ : Perm (Fin n)) :
        sign σ = i : Fin n, jFinset.Ioi i, if σ i < σ j then 1 else -1
        theorem Equiv.Perm.prod_Iio_comp_eq_sign_mul_prod {n : } {R : Type u_1} [CommRing R] (σ : Perm (Fin n)) {f : Fin nFin nR} (hf : ∀ (i j : Fin n), f i j = -f j i) :
        j : Fin n, iFinset.Iio j, f (σ i) (σ j) = (sign σ) * j : Fin n, iFinset.Iio j, f i j
        theorem Equiv.Perm.prod_Ioi_comp_eq_sign_mul_prod {n : } {R : Type u_1} [CommRing R] (σ : Perm (Fin n)) {f : Fin nFin nR} (hf : ∀ (i j : Fin n), f i j = -f j i) :
        i : Fin n, jFinset.Ioi i, f (σ i) (σ j) = (sign σ) * i : Fin n, jFinset.Ioi i, f i j