Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ef69547

Browse files
committed
feat(group_theory/finiteness): Define the minimum number of generators (#12765)
The PR adds a definition of the minimum number of generators, which will be needed for a statement of Schreier's lemma.
1 parent ee4472b commit ef69547

File tree

1 file changed

+22
-0
lines changed

1 file changed

+22
-0
lines changed

src/group_theory/finiteness.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -236,6 +236,10 @@ lemma group.fg_iff : group.fg G ↔
236236
∃ S : set G, subgroup.closure S = (⊤ : subgroup G) ∧ S.finite :=
237237
⟨λ h, (subgroup.fg_iff ⊤).1 h.out, λ h, ⟨(subgroup.fg_iff ⊤).2 h⟩⟩
238238

239+
@[to_additive] lemma group.fg_iff' :
240+
group.fg G ↔ ∃ n (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤ :=
241+
group.fg_def.trans ⟨λ ⟨S, hS⟩, ⟨S.card, S, rfl, hS⟩, λ ⟨n, S, hn, hS⟩, ⟨S, hS⟩⟩
242+
239243
/-- A group is finitely generated if and only if it is finitely generated as a monoid. -/
240244
@[to_additive add_group.fg_iff_add_monoid.fg "An additive group is finitely generated if and only
241245
if it is finitely generated as an additive monoid."]
@@ -264,6 +268,24 @@ group.fg_iff_monoid.fg.mpr $ @monoid.fg_of_surjective G _ G' _ (group.fg_iff_mon
264268
instance group.fg_range {G' : Type*} [group G'] [group.fg G] (f : G →* G') : group.fg f.range :=
265269
group.fg_of_surjective f.range_restrict_surjective
266270

271+
variables (G)
272+
273+
/-- The minimum number of generators of a group. -/
274+
@[to_additive "The minimum number of generators of an additive group"]
275+
def group.rank [h : group.fg G]
276+
[decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] :=
277+
nat.find (group.fg_iff'.mp h)
278+
279+
@[to_additive] lemma group.rank_spec [h : group.fg G]
280+
[decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] :
281+
∃ S : finset G, S.card = group.rank G ∧ subgroup.closure (S : set G) = ⊤ :=
282+
nat.find_spec (group.fg_iff'.mp h)
283+
284+
@[to_additive] lemma group.rank_le [group.fg G]
285+
[decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)]
286+
{S : finset G} (hS : subgroup.closure (S : set G) = ⊤) : group.rank G ≤ S.card :=
287+
nat.find_le ⟨S, rfl, hS⟩
288+
267289
end group
268290

269291
section quotient_group

0 commit comments

Comments
 (0)