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

Commit bdfa72b

Browse files
committed
feat(group_theory/commutator): Define the set of commutators (#17276)
With some of my upcoming group theory PRs, it is going to start being helpful to have a name for the set of commutators.
1 parent 6fa3fc3 commit bdfa72b

File tree

3 files changed

+31
-9
lines changed

3 files changed

+31
-9
lines changed

src/group_theory/abelianization.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -36,24 +36,24 @@ def commutator : subgroup G :=
3636

3737
lemma commutator_def : commutator G = ⁅(⊤ : subgroup G), ⊤⁆ := rfl
3838

39-
lemma commutator_eq_closure : commutator G = subgroup.closure {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} :=
40-
by simp_rw [commutator, subgroup.commutator_def, subgroup.mem_top, exists_true_left]
39+
lemma commutator_eq_closure : commutator G = subgroup.closure (commutator_set G) :=
40+
by simp [commutator, subgroup.commutator_def, commutator_set]
4141

4242
lemma commutator_eq_normal_closure :
43-
commutator G = subgroup.normal_closure {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} :=
44-
by simp_rw [commutator, subgroup.commutator_def', subgroup.mem_top, exists_true_left]
43+
commutator G = subgroup.normal_closure (commutator_set G) :=
44+
by simp [commutator, subgroup.commutator_def', commutator_set]
4545

4646
instance commutator_characteristic : (commutator G).characteristic :=
4747
subgroup.commutator_characteristic ⊤ ⊤
4848

49-
instance [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] : group.fg (commutator G) :=
49+
instance [finite (commutator_set G)] : group.fg (commutator G) :=
5050
begin
5151
rw commutator_eq_closure,
5252
apply group.closure_finite_fg,
5353
end
5454

55-
lemma rank_commutator_le_card [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] :
56-
group.rank (commutator G) ≤ nat.card {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} :=
55+
lemma rank_commutator_le_card [finite (commutator_set G)] :
56+
group.rank (commutator G) ≤ nat.card (commutator_set G) :=
5757
begin
5858
rw subgroup.rank_congr (commutator_eq_closure G),
5959
apply subgroup.rank_closure_finite_le_nat_card,

src/group_theory/commutator.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -204,3 +204,25 @@ begin
204204
end
205205

206206
end subgroup
207+
208+
variables (G)
209+
210+
/-- The set of commutator elements `⁅g₁, g₂⁆` in `G`. -/
211+
def commutator_set : set G :=
212+
{g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}
213+
214+
lemma commutator_set_def : commutator_set G = {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} := rfl
215+
216+
lemma one_mem_commutator_set : (1 : G) ∈ commutator_set G :=
217+
1, 1, commutator_element_self 1
218+
219+
instance : nonempty (commutator_set G) :=
220+
⟨⟨1, one_mem_commutator_set G⟩⟩
221+
222+
variables {G g}
223+
224+
lemma mem_commutator_set_iff : g ∈ commutator_set G ↔ ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g :=
225+
iff.rfl
226+
227+
lemma commutator_mem_commutator_set : ⁅g₁, g₂⁆ ∈ commutator_set G :=
228+
⟨g₁, g₂, rfl⟩

src/group_theory/schreier.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -146,9 +146,9 @@ variables (G)
146146

147147
/-- If `G` has `n` commutators `[g₁, g₂]`, then `|G'| ∣ [G : Z(G)] ^ ([G : Z(G)] * n + 1)`,
148148
where `G'` denotes the commutator of `G`. -/
149-
lemma card_commutator_dvd_index_center_pow [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] :
149+
lemma card_commutator_dvd_index_center_pow [finite (commutator_set G)] :
150150
nat.card (commutator G) ∣
151-
(center G).index ^ ((center G).index * nat.card {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} + 1) :=
151+
(center G).index ^ ((center G).index * nat.card (commutator_set G) + 1) :=
152152
begin
153153
-- First handle the case when `Z(G)` has infinite index and `[G : Z(G)]` is defined to be `0`
154154
by_cases hG : (center G).index = 0,

0 commit comments

Comments
 (0)