Documentation

Mathlib.GroupTheory.Perm.Cycle.Factors

Cycle factors of a permutation #

Let β be a Fintype and f : Equiv.Perm β.

cycleOf #

def Equiv.Perm.cycleOf {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) :
Perm α

f.cycleOf x is the cycle of the permutation f to which x belongs.

Equations
Instances For
    theorem Equiv.Perm.cycleOf_apply {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x y : α) :
    (f.cycleOf x) y = if f.SameCycle x y then f y else y
    theorem Equiv.Perm.cycleOf_inv {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) :
    @[simp]
    theorem Equiv.Perm.cycleOf_pow_apply_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) (n : ) :
    (f.cycleOf x ^ n) x = (f ^ n) x
    @[simp]
    theorem Equiv.Perm.cycleOf_zpow_apply_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) (n : ) :
    (f.cycleOf x ^ n) x = (f ^ n) x
    theorem Equiv.Perm.SameCycle.cycleOf_apply {α : Type u_2} {f : Perm α} {x y : α} [DecidableRel f.SameCycle] :
    f.SameCycle x y(f.cycleOf x) y = f y
    theorem Equiv.Perm.cycleOf_apply_of_not_sameCycle {α : Type u_2} {f : Perm α} {x y : α} [DecidableRel f.SameCycle] :
    ¬f.SameCycle x y(f.cycleOf x) y = y
    theorem Equiv.Perm.SameCycle.cycleOf_eq {α : Type u_2} {f : Perm α} {x y : α} [DecidableRel f.SameCycle] (h : f.SameCycle x y) :
    f.cycleOf x = f.cycleOf y
    @[simp]
    theorem Equiv.Perm.cycleOf_apply_apply_zpow_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) (k : ) :
    (f.cycleOf x) ((f ^ k) x) = (f ^ (k + 1)) x
    @[simp]
    theorem Equiv.Perm.cycleOf_apply_apply_pow_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) (k : ) :
    (f.cycleOf x) ((f ^ k) x) = (f ^ (k + 1)) x
    @[simp]
    theorem Equiv.Perm.cycleOf_apply_apply_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) :
    (f.cycleOf x) (f x) = f (f x)
    @[simp]
    theorem Equiv.Perm.cycleOf_apply_self {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) :
    (f.cycleOf x) x = f x
    theorem Equiv.Perm.IsCycle.cycleOf_eq {α : Type u_2} {f : Perm α} {x : α} [DecidableRel f.SameCycle] (hf : f.IsCycle) (hx : f x x) :
    f.cycleOf x = f
    @[simp]
    theorem Equiv.Perm.cycleOf_eq_one_iff {α : Type u_2} {x : α} (f : Perm α) [DecidableRel f.SameCycle] :
    f.cycleOf x = 1 f x = x
    @[simp]
    theorem Equiv.Perm.cycleOf_self_apply {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (x : α) :
    f.cycleOf (f x) = f.cycleOf x
    @[simp]
    theorem Equiv.Perm.cycleOf_self_apply_pow {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (n : ) (x : α) :
    f.cycleOf ((f ^ n) x) = f.cycleOf x
    @[simp]
    theorem Equiv.Perm.cycleOf_self_apply_zpow {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (n : ) (x : α) :
    f.cycleOf ((f ^ n) x) = f.cycleOf x
    theorem Equiv.Perm.IsCycle.cycleOf {α : Type u_2} {f : Perm α} {x : α} [DecidableRel f.SameCycle] [DecidableEq α] (hf : f.IsCycle) :
    f.cycleOf x = if f x = x then 1 else f
    theorem Equiv.Perm.cycleOf_one {α : Type u_2} [DecidableRel (SameCycle 1)] (x : α) :
    cycleOf 1 x = 1
    theorem Equiv.Perm.isCycle_cycleOf {α : Type u_2} {x : α} (f : Perm α) [DecidableRel f.SameCycle] (hx : f x x) :
    theorem Equiv.Perm.pow_mod_orderOf_cycleOf_apply {α : Type u_2} (f : Perm α) [DecidableRel f.SameCycle] (n : ) (x : α) :
    (f ^ (n % orderOf (f.cycleOf x))) x = (f ^ n) x
    theorem Equiv.Perm.cycleOf_mul_of_apply_right_eq_self {α : Type u_2} {f g : Perm α} [DecidableRel f.SameCycle] [DecidableRel (f * g).SameCycle] (h : Commute f g) (x : α) (hx : g x = x) :
    (f * g).cycleOf x = f.cycleOf x
    theorem Equiv.Perm.isCycle_cycleOf_iff {α : Type u_2} {x : α} (f : Perm α) [DecidableRel f.SameCycle] :
    (f.cycleOf x).IsCycle f x x

    x is in the support of f iff Equiv.Perm.cycle_of f x is a cycle.

    @[simp]
    theorem Equiv.Perm.two_le_card_support_cycleOf_iff {α : Type u_2} {f : Perm α} {x : α} [DecidableEq α] [Fintype α] :
    2 (f.cycleOf x).support.card f x x
    @[simp]
    theorem Equiv.Perm.support_cycleOf_nonempty {α : Type u_2} {f : Perm α} {x : α} [DecidableEq α] [Fintype α] :
    theorem Equiv.Perm.mem_support_cycleOf_iff {α : Type u_2} {f : Perm α} {x y : α} [DecidableEq α] [Fintype α] :
    theorem Equiv.Perm.mem_support_cycleOf_iff' {α : Type u_2} {f : Perm α} {x y : α} (hx : f x x) [DecidableEq α] [Fintype α] :
    theorem Equiv.Perm.support_cycleOf_eq_nil_iff {α : Type u_2} {f : Perm α} {x : α} [DecidableEq α] [Fintype α] :
    (f.cycleOf x).support = xf.support
    theorem Equiv.Perm.isCycleOn_support_cycleOf {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) (x : α) :
    theorem Equiv.Perm.SameCycle.exists_pow_eq_of_mem_support {α : Type u_2} {x y : α} {f : Perm α} [DecidableEq α] [Fintype α] (h : f.SameCycle x y) (hx : x f.support) :
    i < (f.cycleOf x).support.card, (f ^ i) x = y
    theorem Equiv.Perm.support_cycleOf_le {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) (x : α) :
    theorem Equiv.Perm.SameCycle.mem_support_iff {α : Type u_2} {x y : α} {f : Perm α} [DecidableEq α] [Fintype α] (h : f.SameCycle x y) :
    theorem Equiv.Perm.pow_mod_card_support_cycleOf_self_apply {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) (n : ) (x : α) :
    (f ^ (n % (f.cycleOf x).support.card)) x = (f ^ n) x
    theorem Equiv.Perm.SameCycle.exists_pow_eq {α : Type u_2} {x y : α} [DecidableEq α] [Fintype α] (f : Perm α) (h : f.SameCycle x y) :
    ∃ (i : ), 0 < i i (f.cycleOf x).support.card + 1 (f ^ i) x = y
    theorem Equiv.Perm.zpow_eq_zpow_on_iff {α : Type u_2} [DecidableEq α] [Fintype α] (g : Perm α) {m n : } {x : α} (hx : g x x) :
    (g ^ m) x = (g ^ n) x m % (g.cycleOf x).support.card = n % (g.cycleOf x).support.card

    cycleFactors #

    def Equiv.Perm.cycleFactorsAux {α : Type u_2} [DecidableEq α] [Fintype α] (l : List α) (f : Perm α) (h : ∀ {x : α}, f x xx l) :
    { pl : List (Perm α) // pl.prod = f (∀ gpl, g.IsCycle) List.Pairwise Disjoint pl }

    Given a list l : List α and a permutation f : Perm α whose nonfixed points are all in l, recursively factors f into cycles.

    Equations
    Instances For
      def Equiv.Perm.cycleFactorsAux.go {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) (l : List α) (g : Perm α) (hg : ∀ {x : α}, g x xx l) (hfg : ∀ {x : α}, g x xf.cycleOf x = g.cycleOf x) :
      { pl : List (Perm α) // pl.prod = g (∀ g'pl, g'.IsCycle) List.Pairwise Disjoint pl }

      The auxiliary of cycleFactorsAux. This functions separates cycles from f instead of g to prevent the process of a cycle gets complex.

      Equations
      Instances For
        theorem Equiv.Perm.mem_list_cycles_iff {α : Type u_4} [Finite α] {l : List (Perm α)} (h1 : σl, σ.IsCycle) (h2 : List.Pairwise Disjoint l) {σ : Perm α} :
        σ l σ.IsCycle ∀ (a : α), σ a aσ a = l.prod a
        theorem Equiv.Perm.list_cycles_perm_list_cycles {α : Type u_4} [Finite α] {l₁ l₂ : List (Perm α)} (h₀ : l₁.prod = l₂.prod) (h₁l₁ : σl₁, σ.IsCycle) (h₁l₂ : σl₂, σ.IsCycle) (h₂l₁ : List.Pairwise Disjoint l₁) (h₂l₂ : List.Pairwise Disjoint l₂) :
        l₁.Perm l₂
        def Equiv.Perm.cycleFactors {α : Type u_2} [Fintype α] [LinearOrder α] (f : Perm α) :
        { l : List (Perm α) // l.prod = f (∀ gl, g.IsCycle) List.Pairwise Disjoint l }

        Factors a permutation f into a list of disjoint cyclic permutations that multiply to f.

        Equations
        Instances For
          def Equiv.Perm.truncCycleFactors {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) :
          Trunc { l : List (Perm α) // l.prod = f (∀ gl, g.IsCycle) List.Pairwise Disjoint l }

          Factors a permutation f into a list of disjoint cyclic permutations that multiply to f, without a linear order.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            def Equiv.Perm.cycleFactorsFinset {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) :

            Factors a permutation f into a Finset of disjoint cyclic permutations that multiply to f.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              theorem Equiv.Perm.cycleFactorsFinset_eq_list_toFinset {α : Type u_2} [DecidableEq α] [Fintype α] {σ : Perm α} {l : List (Perm α)} (hn : l.Nodup) :
              theorem Equiv.Perm.cycleFactorsFinset_eq_finset {α : Type u_2} [DecidableEq α] [Fintype α] {σ : Perm α} {s : Finset (Perm α)} :
              σ.cycleFactorsFinset = s (∀ fs, f.IsCycle) ∃ (h : (↑s).Pairwise Disjoint), s.noncommProd id = σ

              Two cycles of a permutation commute.

              theorem Equiv.Perm.cycleFactorsFinset_mem_commute' {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) {g1 g2 : Perm α} (h1 : g1 f.cycleFactorsFinset) (h2 : g2 f.cycleFactorsFinset) :
              Commute g1 g2

              Two cycles of a permutation commute.

              The product of cycle factors is equal to the original f : perm α.

              theorem Equiv.Perm.mem_cycleFactorsFinset_iff {α : Type u_2} [DecidableEq α] [Fintype α] {f p : Perm α} :
              p f.cycleFactorsFinset p.IsCycle ap.support, p a = f a
              theorem Equiv.Perm.pairwise_disjoint_of_mem_zpowers {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) :
              Pairwise fun (i j : { x : Perm α // x f.cycleFactorsFinset }) => ∀ (x y : Perm α), x Subgroup.zpowers iy Subgroup.zpowers jx.Disjoint y
              theorem Equiv.Perm.pairwise_commute_of_mem_zpowers {α : Type u_2} [DecidableEq α] [Fintype α] (f : Perm α) :
              Pairwise fun (i j : { x : Perm α // x f.cycleFactorsFinset }) => ∀ (x y : Perm α), x Subgroup.zpowers iy Subgroup.zpowers jCommute x y

              Two permutations f g : Perm α have the same cycle factors iff they are the same.

              theorem Equiv.Perm.cycle_is_cycleOf {α : Type u_2} [DecidableEq α] [Fintype α] {f c : Perm α} {a : α} (ha : a c.support) (hc : c f.cycleFactorsFinset) :
              c = f.cycleOf a

              If c is a cycle, a ∈ c.support and c is a cycle of f, then c = f.cycleOf a

              theorem Equiv.Perm.zpow_apply_mem_support_of_mem_cycleFactorsFinset_iff {α : Type u_2} [DecidableEq α] [Fintype α] {g : Perm α} {x : α} {m : } {c : { x : Perm α // x g.cycleFactorsFinset }} :
              (g ^ m) x (↑c).support x (↑c).support

              A permutation c is a cycle of g iff k * c * k⁻¹ is a cycle of k * g * k⁻¹

              theorem Equiv.Perm.commute_of_mem_cycleFactorsFinset_commute {α : Type u_2} [DecidableEq α] [Fintype α] (k g : Perm α) (hk : cg.cycleFactorsFinset, Commute k c) :

              If a permutation commutes with every cycle of g, then it commutes with g

              NB. The converse is false. Commuting with every cycle of g means that we belong to the kernel of the action of Equiv.Perm α on g.cycleFactorsFinset

              The cycles of a permutation commute with it

              theorem Equiv.Perm.mem_support_cycle_of_cycle {α : Type u_2} [DecidableEq α] [Fintype α] {g d c : Perm α} (hc : c g.cycleFactorsFinset) (hd : d g.cycleFactorsFinset) (x : α) :

              If c and d are cycles of g, then d stabilizes the support of c

              theorem Equiv.Perm.mem_cycleFactorsFinset_support {α : Type u_2} [DecidableEq α] [Fintype α] {g c : Perm α} (hc : c g.cycleFactorsFinset) (a : α) :

              If a permutation is a cycle of g, then its support is invariant under g.

              theorem Equiv.Perm.cycle_induction_on {β : Type u_3} [Finite β] (P : Perm βProp) (σ : Perm β) (base_one : P 1) (base_cycles : ∀ (σ : Perm β), σ.IsCycleP σ) (induction_disjoint : ∀ (σ τ : Perm β), σ.Disjoint τσ.IsCycleP σP τP (σ * τ)) :
              P σ
              theorem Equiv.Perm.IsCycle.forall_commute_iff {α : Type u_2} [DecidableEq α] [Fintype α] (g z : Perm α) :
              (∀ cg.cycleFactorsFinset, Commute z c) cg.cycleFactorsFinset, ∃ (hc : ∀ (x : α), x c.support z x c.support), ofSubtype (z.subtypePerm hc) Subgroup.zpowers c

              A permutation restricted to the support of a cycle factor is that cycle factor

              theorem Equiv.Perm.commute_iff_of_mem_cycleFactorsFinset {α : Type u_2} [DecidableEq α] [Fintype α] {g k c : Perm α} (hc : c g.cycleFactorsFinset) :
              Commute k c ∃ (hc' : ∀ (x : α), x c.support k x c.support), k.subtypePerm hc' Subgroup.zpowers (g.subtypePerm )