Documentation

Mathlib.GroupTheory.CommutingProbability

Commuting Probability #

This file introduces the commuting probability of finite groups.

Main definitions #

TODO #

def commProb (M : Type u_1) [Mul M] :

The commuting probability of a finite type with a multiplication operation.

Equations
theorem commProb_def (M : Type u_1) [Mul M] :
commProb M = (Nat.card { p : M × M // Commute p.1 p.2 }) / (Nat.card M) ^ 2
theorem commProb_prod (M : Type u_1) [Mul M] (M' : Type u_2) [Mul M'] :
theorem commProb_pi {α : Type u_2} (i : αType u_3) [Fintype α] [(a : α) → Mul (i a)] :
commProb ((a : α) → i a) = a : α, commProb (i a)
theorem commProb_function {α : Type u_2} {β : Type u_3} [Fintype α] [Mul β] :
commProb (αβ) = commProb β ^ Fintype.card α
@[simp]
theorem commProb_eq_zero_of_infinite (M : Type u_1) [Mul M] [Infinite M] :
theorem commProb_pos (M : Type u_1) [Mul M] [Finite M] [h : Nonempty M] :
theorem commProb_le_one (M : Type u_1) [Mul M] [Finite M] :
theorem commProb_eq_one_iff {M : Type u_1} [Mul M] [Finite M] [h : Nonempty M] :
commProb M = 1 Std.Commutative fun (x1 x2 : M) => x1 * x2
theorem commProb_def' (G : Type u_2) [Group G] :
theorem Subgroup.commProb_subgroup_le {G : Type u_2} [Group G] [Finite G] (H : Subgroup G) :
commProb H commProb G * H.index ^ 2
theorem Subgroup.commProb_quotient_le {G : Type u_2} [Group G] [Finite G] (H : Subgroup G) [H.Normal] :
commProb (G H) commProb G * (Nat.card H)
theorem DihedralGroup.commProb_odd {n : } (hn : Odd n) :
commProb (DihedralGroup n) = (n + 3) / (4 * n)
@[irreducible]

A list of Dihedral groups whose product will have commuting probability 1 / n.

Equations
  • One or more equations did not get rendered due to their size.
theorem DihedralGroup.reciprocalFactors_odd {n : } (h1 : n 1) (h2 : Odd n) :
@[reducible, inline]

A finite product of Dihedral groups.

Equations
@[irreducible]

Construction of a group with commuting probability 1 / n.