Documentation

Mathlib.GroupTheory.Perm.Centralizer

Centralizer of a permutation and cardinality of conjugacy classes in the symmetric groups #

Let α : Type with Fintype α (and DecidableEq α). The main goal of this file is to compute the cardinality of conjugacy classes in Equiv.Perm α. Every g : Equiv.Perm α has a cycleType α : Multiset. By Equiv.Perm.isConj_iff_cycleType_eq, two permutations are conjugate in Equiv.Perm α iff their cycle types are equal. To compute the cardinality of the conjugacy classes, we could use a purely combinatorial approach and compute the number of permutations with given cycle type but we resorted to a more algebraic approach based on the study of the centralizer of a permutation g.

Given g : Equiv.Perm α, the conjugacy class of g is the orbit of g under the action ConjAct (Equiv.Perm α), and we use the orbit-stabilizer theorem (MulAction.card_orbit_mul_card_stabilizer_eq_card_group) to reduce the computation to the computation of the centralizer of g, the subgroup of Equiv.Perm α consisting of all permutations which commute with g. It is accessed here as MulAction.stabilizer (ConjAct (Equiv.Perm α)) g and Equiv.Perm.centralizer_eq_comap_stabilizer.

We compute this subgroup as follows.

This is shown by constructing a right inverse Equiv.Perm.Basis.toCentralizer, as established by Equiv.Perm.Basis.toPermHom_apply_toCentralizer.

This allows to give a description of the kernel of Equiv.Perm.OnCycleFactors.toPermHom g as the product of a symmetric group and of a product of cyclic groups. This analysis starts with the morphism Equiv.Perm.OnCycleFactors.θ, its injectivity Equiv.Perm.OnCycleFactors.θ_injective, its range Equiv.Perm.OnCycleFactors.θ_range_eq, and its cardinality Equiv.Perm.OnCycleFactors.θ_range_card.

The canonical morphism from Subgroup.centralizer {g} to the group of permutations of g.cycleFactorsFinset

Equations
Instances For
    theorem Equiv.Perm.OnCycleFactors.centralizer_smul_def {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) (k : (Subgroup.centralizer {g})) (c : { x : Perm α // x g.cycleFactorsFinset }) :
    k c = k * c * k⁻¹,
    @[simp]
    theorem Equiv.Perm.OnCycleFactors.val_centralizer_smul {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) (k : (Subgroup.centralizer {g})) (c : { x : Perm α // x g.cycleFactorsFinset }) :
    ↑(k c) = k * c * k⁻¹
    theorem Equiv.Perm.OnCycleFactors.toPermHom_apply {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) (k : (Subgroup.centralizer {g})) (c : { x : Perm α // x g.cycleFactorsFinset }) :
    ((toPermHom g) k) c = k c
    theorem Equiv.Perm.OnCycleFactors.coe_toPermHom {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) (k : (Subgroup.centralizer {g})) (c : { x : Perm α // x g.cycleFactorsFinset }) :
    (((toPermHom g) k) c) = k * c * (↑k)⁻¹

    The range of Equiv.Perm.OnCycleFactors.toPermHom.

    The equality is proved by Equiv.Perm.OnCycleFactors.range_toPermHom_eq_range_toPermHom'.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      k : Subgroup.centralizer {g} belongs to the kernel of toPermHom g iff it commutes with each cycle of g

      structure Equiv.Perm.Basis {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) :
      Type u_1

      A Basis of a permutation is a choice of an element in each of its cycles

      Instances For
        Equations
        theorem Equiv.Perm.Basis.nonempty {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) :
        theorem Equiv.Perm.Basis.mem_support_self {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (c : { x : Perm α // x g.cycleFactorsFinset }) :
        a c (↑c).support
        theorem Equiv.Perm.Basis.injective {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) :
        theorem Equiv.Perm.Basis.cycleOf_eq {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (c : { x : Perm α // x g.cycleFactorsFinset }) :
        g.cycleOf (a c) = c
        theorem Equiv.Perm.Basis.sameCycle {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) {x : α} (hx : g.cycleOf x g.cycleFactorsFinset) :
        g.SameCycle (a g.cycleOf x, hx) x
        def Equiv.Perm.Basis.ofPermHomFun {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (τ : (OnCycleFactors.range_toPermHom' g)) (x : α) :
        α

        The function that will provide a right inverse toCentralizer to toPermHom

        Equations
        Instances For
          theorem Equiv.Perm.Basis.mem_fixedPoints_or_exists_zpow_eq {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (x : α) :
          x Function.fixedPoints g ∃ (c : { x : Perm α // x g.cycleFactorsFinset }) (_ : x (↑c).support) (m : ), (g ^ m) (a c) = x
          theorem Equiv.Perm.Basis.ofPermHomFun_apply_of_cycleOf_mem {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (τ : (OnCycleFactors.range_toPermHom' g)) {x : α} {c : { x : Perm α // x g.cycleFactorsFinset }} (hx : x (↑c).support) {m : } (hm : (g ^ m) (a c) = x) :
          a.ofPermHomFun τ x = (g ^ m) (a (τ c))
          theorem Equiv.Perm.Basis.ofPermHomFun_apply_mem_support_cycle_iff {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (τ : (OnCycleFactors.range_toPermHom' g)) {x : α} {c : { x : Perm α // x g.cycleFactorsFinset }} :
          a.ofPermHomFun τ x (↑(τ c)).support x (↑c).support
          theorem Equiv.Perm.Basis.ofPermHomFun_commute_zpow_apply {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (τ : (OnCycleFactors.range_toPermHom' g)) (x : α) (j : ) :
          a.ofPermHomFun τ ((g ^ j) x) = (g ^ j) (a.ofPermHomFun τ x)
          theorem Equiv.Perm.Basis.ofPermHomFun_mul {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (σ τ : (OnCycleFactors.range_toPermHom' g)) (x : α) :
          a.ofPermHomFun (σ * τ) x = a.ofPermHomFun σ (a.ofPermHomFun τ x)
          theorem Equiv.Perm.Basis.ofPermHomFun_one {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (x : α) :
          a.ofPermHomFun 1 x = x
          noncomputable def Equiv.Perm.Basis.ofPermHom {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) :

          Given a : g.Basis and a permutation of g.cycleFactorsFinset that preserve the lengths of the cycles, a permutation of α that moves the Basis and commutes with g

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            noncomputable def Equiv.Perm.Basis.toCentralizer {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) :

            Given a : Equiv.Perm.Basis g, we define a right inverse of Equiv.Perm.OnCycleFactors.toPermHom, on range_toPermHom' g

            Equations
            Instances For
              theorem Equiv.Perm.Basis.toCentralizer_apply {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (τ : (OnCycleFactors.range_toPermHom' g)) (x : α) :
              (a.toCentralizer τ) x = a.ofPermHomFun τ x
              theorem Equiv.Perm.Basis.toCentralizer_equivariant {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} (a : g.Basis) (c : { x : Perm α // x g.cycleFactorsFinset }) (τ : (OnCycleFactors.range_toPermHom' g)) :
              a.toCentralizer τ c = τ c
              theorem Equiv.Perm.OnCycleFactors.mem_range_toPermHom_iff {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} {τ : Perm { x : Perm α // x g.cycleFactorsFinset }} :
              τ (toPermHom g).range ∀ (c : { x : Perm α // x g.cycleFactorsFinset }), (↑(τ c)).support.card = (↑c).support.card
              theorem Equiv.Perm.OnCycleFactors.mem_range_toPermHom_iff' {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} {τ : Perm { x : Perm α // x g.cycleFactorsFinset }} :
              τ (toPermHom g).range (fun (c : { x : Perm α // x g.cycleFactorsFinset }) => (↑c).support.card) τ = fun (c : { x : Perm α // x g.cycleFactorsFinset }) => (↑c).support.card

              Unapplied variant of Equiv.Perm.mem_range_toPermHom_iff

              def Equiv.Perm.OnCycleFactors.kerParam {α : Type u_1} [DecidableEq α] [Fintype α] (g : Perm α) :
              Perm (Function.fixedPoints g) × ((c : { x : Perm α // x g.cycleFactorsFinset }) → (Subgroup.zpowers c)) →* Perm α

              The parametrization of the kernel of toPermHom

              Equations
              Instances For
                theorem Equiv.Perm.OnCycleFactors.kerParam_apply {α : Type u_1} [DecidableEq α] [Fintype α] {g : Perm α} {u : Perm (Function.fixedPoints g)} {v : (c : { x : Perm α // x g.cycleFactorsFinset }) → (Subgroup.zpowers c)} {x : α} :
                ((kerParam g) (u, v)) x = if hx : g.cycleOf x g.cycleFactorsFinset then (v g.cycleOf x, hx) x else (ofSubtype u) x

                Cardinality of the centralizer in Equiv.Perm α of a permutation given cycleType

                Cardinality of a conjugacy class in Equiv.Perm α of a given cycleType

                theorem Equiv.Perm.card_of_cycleType_eq_zero_iff (α : Type u_1) [DecidableEq α] [Fintype α] {m : Multiset } :
                (Finset.filter (fun (g : Perm α) => g.cycleType = m) Finset.univ).card = 0 ¬(m.sum Fintype.card α am, 2 a)
                theorem Equiv.Perm.card_of_cycleType (α : Type u_1) [DecidableEq α] [Fintype α] (m : Multiset ) :
                (Finset.filter (fun (g : Perm α) => g.cycleType = m) Finset.univ).card = if m.sum Fintype.card α am, 2 a then (Fintype.card α).factorial / ((Fintype.card α - m.sum).factorial * m.prod * nm.toFinset, (Multiset.count n m).factorial) else 0

                Cardinality of the Finset of Equiv.Perm α of given cycleType