Documentation

Mathlib.GroupTheory.Rank

Rank of a group #

This file defines the rank of a group, namely the minimum size of a generating set.

TODO #

Should we define erank G : ℕ∞ the rank of a not necessarily finitely generated group G, then redefine rank G as (erank G).toNat? Maybe a Cardinal-valued version too?

noncomputable def Group.rank (G : Type u_1) [Group G] [h : FG G] :

The minimum number of generators of a group.

Equations
Instances For
    noncomputable def AddGroup.rank (G : Type u_1) [AddGroup G] [h : FG G] :

    The minimum number of generators of an additive group.

    Equations
    Instances For
      theorem Group.rank_spec (G : Type u_1) [Group G] [h : FG G] :
      ∃ (S : Finset G), S.card = rank G Subgroup.closure S =
      theorem AddGroup.rank_spec (G : Type u_1) [AddGroup G] [h : FG G] :
      ∃ (S : Finset G), S.card = rank G AddSubgroup.closure S =
      theorem Group.rank_le {G : Type u_1} [Group G] [h : FG G] {S : Finset G} (hS : Subgroup.closure S = ) :
      theorem AddGroup.rank_le {G : Type u_1} [AddGroup G] [h : FG G] {S : Finset G} (hS : AddSubgroup.closure S = ) :
      theorem Group.rank_le_of_surjective {G : Type u_1} {H : Type u_2} [Group G] [Group H] [FG G] [FG H] (f : G →* H) (hf : Function.Surjective f) :
      theorem AddGroup.rank_le_of_surjective {G : Type u_1} {H : Type u_2} [AddGroup G] [AddGroup H] [FG G] [FG H] (f : G →+ H) (hf : Function.Surjective f) :
      theorem Group.rank_range_le {G : Type u_1} {H : Type u_2} [Group G] [Group H] [FG G] {f : G →* H} :
      theorem AddGroup.rank_range_le {G : Type u_1} {H : Type u_2} [AddGroup G] [AddGroup H] [FG G] {f : G →+ H} :
      theorem Group.rank_congr {G : Type u_1} {H : Type u_2} [Group G] [Group H] [FG G] [FG H] (e : G ≃* H) :
      rank G = rank H
      theorem AddGroup.rank_congr {G : Type u_1} {H : Type u_2} [AddGroup G] [AddGroup H] [FG G] [FG H] (e : G ≃+ H) :
      rank G = rank H
      theorem Subgroup.rank_congr {G : Type u_1} [Group G] {H K : Subgroup G} [Group.FG H] [Group.FG K] (h : H = K) :
      theorem AddSubgroup.rank_congr {G : Type u_1} [AddGroup G] {H K : AddSubgroup G} [AddGroup.FG H] [AddGroup.FG K] (h : H = K) :