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

Commit 2e7483d

Browse files
committed
refactor(group_theory/commutator): Move and golf commutator_le (#12599)
This PR golfs `commutator_le` and moves it earlier in the file so that it can be used earlier. This PR will conflict with #12600, so don't merge them simultaneously (bors d+ might be better).
1 parent 09d0f02 commit 2e7483d

File tree

2 files changed

+7
-19
lines changed

2 files changed

+7
-19
lines changed

src/group_theory/commutator.lean

Lines changed: 6 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -62,6 +62,10 @@ lemma commutator_mem_commutator {H₁ H₂ : subgroup G} {g₁ g₂ : G} (h₁ :
6262
⁅g₁, g₂⁆ ∈ ⁅H₁, H₂⁆ :=
6363
subset_closure ⟨g₁, h₁, g₂, h₂, rfl⟩
6464

65+
lemma commutator_le {H₁ H₂ H₃ : subgroup G} :
66+
⁅H₁, H₂⁆ ≤ H₃ ↔ ∀ (g₁ ∈ H₁) (g₂ ∈ H₂), ⁅g₁, g₂⁆ ∈ H₃ :=
67+
H₃.closure_le.trans ⟨λ h a b c d, h ⟨a, b, c, d, rfl⟩, λ h g ⟨a, b, c, d, h_eq⟩, h_eq ▸ h a b c d⟩
68+
6569
instance commutator_normal (H₁ H₂ : subgroup G) [h₁ : H₁.normal]
6670
[h₂ : H₂.normal] : normal ⁅H₁, H₂⁆ :=
6771
begin
@@ -88,17 +92,6 @@ lemma commutator_def' (H₁ H₂ : subgroup G) [H₁.normal] [H₂.normal] :
8892
⁅H₁, H₂⁆ = normal_closure {g | ∃ (g₁ ∈ H₁) (g₂ ∈ H₂), ⁅g₁, g₂⁆ = g} :=
8993
le_antisymm closure_le_normal_closure (normal_closure_le_normal subset_closure)
9094

91-
lemma commutator_le (H₁ H₂ : subgroup G) (K : subgroup G) :
92-
⁅H₁, H₂⁆ ≤ K ↔ ∀ (p ∈ H₁) (q ∈ H₂), ⁅p, q⁆ ∈ K :=
93-
begin
94-
rw [subgroup.commutator, closure_le],
95-
split,
96-
{ intros h p hp q hq,
97-
exact h ⟨p, hp, q, hq, rfl⟩, },
98-
{ rintros h x ⟨p, hp, q, hq, rfl⟩,
99-
exact h p hp q hq, }
100-
end
101-
10295
lemma commutator_comm (H₁ H₂ : subgroup G) : ⁅H₁, H₂⁆ = ⁅H₂, H₁⁆ :=
10396
begin
10497
suffices : ∀ H₁ H₂ : subgroup G, ⁅H₁, H₂⁆ ≤ ⁅H₂, H₁⁆, { exact le_antisymm (this _ _) (this _ _) },
@@ -111,7 +104,7 @@ begin
111104
end
112105

113106
lemma commutator_le_right (H₁ H₂ : subgroup G) [h : normal H₂] : ⁅H₁, H₂⁆ ≤ H₂ :=
114-
(commutator_le _ _ _).mpr (λ g₁ h₁ g₂ h₂, H₂.mul_mem (h.conj_mem g₂ h₂ g₁) (H₂.inv_mem h₂))
107+
commutator_le.mpr (λ g₁ h₁ g₂ h₂, H₂.mul_mem (h.conj_mem g₂ h₂ g₁) (H₂.inv_mem h₂))
115108

116109
lemma commutator_le_left (H₁ H₂ : subgroup G) [h : normal H₁] : ⁅H₁, H₂⁆ ≤ H₁ :=
117110
commutator_comm H₂ H₁ ▸ commutator_le_right H₂ H₁
@@ -159,7 +152,7 @@ See `commutator_pi_pi_of_fintype` for equality given `fintype η`.
159152
lemma commutator_pi_pi_le {η : Type*} {Gs : η → Type*} [∀ i, group (Gs i)]
160153
(H K : Π i, subgroup (Gs i)) :
161154
⁅subgroup.pi set.univ H, subgroup.pi set.univ K⁆ ≤ subgroup.pi set.univ (λ i, ⁅H i, K i⁆) :=
162-
(commutator_le _ _ _).mpr $ λ p hp q hq i hi, commutator_mem_commutator (hp i hi) (hq i hi)
155+
commutator_le.mpr $ λ p hp q hq i hi, commutator_mem_commutator (hp i hi) (hq i hi)
163156

164157
/-- The commutator of a finite direct product is contained in the direct product of the commutators.
165158
-/

src/group_theory/nilpotent.lean

Lines changed: 1 addition & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -319,12 +319,7 @@ end
319319
lemma descending_central_series_ge_lower (H : ℕ → subgroup G)
320320
(hH : is_descending_central_series H) : ∀ n : ℕ, lower_central_series G n ≤ H n
321321
| 0 := hH.1.symm ▸ le_refl ⊤
322-
| (n + 1) := begin
323-
specialize descending_central_series_ge_lower n,
324-
apply (commutator_le _ _ _).2,
325-
intros x hx q _,
326-
exact hH.2 x n (descending_central_series_ge_lower hx) q,
327-
end
322+
| (n + 1) := commutator_le.mpr (λ x hx q _, hH.2 x n (descending_central_series_ge_lower n hx) q)
328323

329324
/-- A group is nilpotent if and only if its lower central series eventually reaches
330325
the trivial subgroup. -/

0 commit comments

Comments
 (0)