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.
If
h : Subgroup.centralizer {g}
, then the action ofConjAct.toConjAct h
by conjugation onEquiv.Perm α
stabilizesg.cycleFactorsFinset
. That induces an action ofSubgroup.centralizer {g}
ong.cycleFactorsFinset
which is defined viaEquiv.Perm.OnCycleFactors.subMulActionOnCycleFactors
This action defines a group morphism
Equiv.Perm.OnCycleFactors.toPermHom g
fromSubgroup.centralizer {g}
toEquiv.Perm g.cycleFactorsFinset
.Equiv.Perm.OnCycleFactors.range_toPermHom'
is the subgroup ofEquiv.Perm g.cycleFactorsFinset
consisting of permutations that preserve the cardinality of the support.Equiv.Perm.OnCycleFactors.range_toPermHom_eq_range_toPermHom' shows that the range of
Equiv.Perm.OnCycleFactors.toPermHom gis the subgroup
toPermHom_range' gof
Equiv.Perm g.cycleFactorsFinset`.
This is shown by constructing a right inverse
Equiv.Perm.Basis.toCentralizer
, as established by
Equiv.Perm.Basis.toPermHom_apply_toCentralizer
.
Equiv.Perm.OnCycleFactors.nat_card_range_toPermHom
computes the cardinality of(Equiv.Perm.OnCycleFactors.toPermHom g).range
as a product of factorials.Equiv.Perm.OnCycleFactors.mem_ker_toPermHom_iff
proves thatk : Subgroup.centralizer {g}
belongs to the kernel ofEquiv.Perm.OnCycleFactors.toPermHom g
if and only if it commutes with each cycle ofg
. This is equivalent to the conjunction of two properties:k
preserves the set of fixed points ofg
;- on each cycle
c
,k
acts as a power of that cycle.
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
.
Equiv.Perm.nat_card_centralizer g
computes the cardinality of the centralizer ofg
.Equiv.Perm.card_isConj_mul_eq g
computes the cardinality of the conjugacy class ofg
.We now can compute the cardinality of the set of permutations with given cycle type. The condition for this cardinality to be zero is given by
Equiv.Perm.card_of_cycleType_eq_zero_iff
which is itself derived fromEquiv.Perm.exists_with_cycleType_iff
.Equiv.Perm.card_of_cycleType_mul_eq m
andEquiv.Perm.card_of_cycleType m
compute this cardinality.
The action by conjugation of Subgroup.centralizer {g}
on the cycles of a given permutation
Equations
Instances For
The conjugation action of Subgroup.centralizer {g}
on g.cycleFactorsFinset
Equations
Instances For
The canonical morphism from Subgroup.centralizer {g}
to the group of permutations of g.cycleFactorsFinset
Equations
Instances For
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
A Basis
of a permutation is a choice of an element in each of its cycles
A choice of elements in each cycle
For each cycle, the chosen element belongs to the cycle
Instances For
Equations
- g.instDFunLikeBasisSubtypeMemFinsetCycleFactorsFinset = { coe := fun (a : g.Basis) => a.toFun, coe_injective' := ⋯ }
The function that will provide a right inverse toCentralizer
to toPermHom
Equations
- a.ofPermHomFun τ x = if hx : g.cycleOf x ∈ g.cycleFactorsFinset then (g ^ Nat.find ⋯) (a (↑τ ⟨g.cycleOf x, hx⟩)) else x
Instances For
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
Given a : Equiv.Perm.Basis g
,
we define a right inverse of Equiv.Perm.OnCycleFactors.toPermHom
,
on range_toPermHom' g
Equations
- a.toCentralizer = { toFun := fun (τ : ↥(Equiv.Perm.OnCycleFactors.range_toPermHom' g)) => ⟨a.ofPermHom τ, ⋯⟩, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Unapplied variant of Equiv.Perm.mem_range_toPermHom_iff
Computes the range of Equiv.Perm.toPermHom g
The parametrization of the kernel of toPermHom
Equations
Instances For
Cardinality of the centralizer in Equiv.Perm α
of a permutation given cycleType
Cardinality of a conjugacy class in Equiv.Perm α
of a given cycleType
Cardinality of the Finset
of Equiv.Perm α
of given cycleType