Skip to content

Commit

Permalink
feat(group_theory/nilpotent): add subgroups of nilpotent group are ni…
Browse files Browse the repository at this point in the history
…lpotent (#8854)




Co-authored-by: ineswright <52750254+ineswright@users.noreply.github.com>
  • Loading branch information
ineswright and ineswright committed Aug 28, 2021
1 parent 1aaff8d commit 0b48570
Showing 1 changed file with 28 additions and 1 deletion.
29 changes: 28 additions & 1 deletion src/group_theory/nilpotent.lean
@@ -1,7 +1,7 @@
/-
Copyright (c) 2021 Kevin Buzzard. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kevin Buzzard
Authors: Kevin Buzzard, Ines Wright
-/

import group_theory.general_commutator
Expand Down Expand Up @@ -265,6 +265,11 @@ lemma mem_lower_central_series_succ_iff (n : ℕ) (q : G) :
q ∈ closure {x | ∃ (p ∈ lower_central_series G n) (q ∈ (⊤ : subgroup G)), p * q * p⁻¹ * q⁻¹ = x}
:= iff.rfl

lemma lower_central_series_succ (n : ℕ) :
lower_central_series G (n + 1) =
closure {x | ∃ (p ∈ lower_central_series G n) (q ∈ (⊤ : subgroup G)), p * q * p⁻¹ * q⁻¹ = x} :=
rfl

instance (n : ℕ) : normal (lower_central_series G n) :=
begin
induction n with d hd,
Expand Down Expand Up @@ -308,6 +313,28 @@ begin
use [lower_central_series G, lower_central_series_is_descending_central_series, h] },
end

lemma lower_central_series_map_subtype_le (H : subgroup G) (n : ℕ) :
(lower_central_series H n).map H.subtype ≤ lower_central_series G n :=
begin
induction n with d hd,
{ simp },
{ rw [lower_central_series_succ, lower_central_series_succ, monoid_hom.map_closure],
apply subgroup.closure_mono,
rintros x1 ⟨x2, ⟨x3, hx3, x4, hx4, rfl⟩, rfl⟩,
exact ⟨x3, (hd (mem_map.mpr ⟨x3, hx3, rfl⟩)), x4, by simp⟩ }
end

instance subgroup.is_nilpotent (H : subgroup G) [hG : is_nilpotent G] :
is_nilpotent H :=
begin
rw nilpotent_iff_lower_central_series at *,
rcases hG with ⟨n, hG⟩,
use n,
have := lower_central_series_map_subtype_le H n,
simp only [hG, set_like.le_def, mem_map, forall_apply_eq_imp_iff₂, exists_imp_distrib] at this,
exact eq_bot_iff.mpr (λ x hx, subtype.ext (this x hx)),
end

@[priority 100]
instance is_nilpotent_of_subsingleton [subsingleton G] : is_nilpotent G :=
nilpotent_iff_lower_central_series.20, subsingleton.elim ⊤ ⊥⟩
Expand Down

0 comments on commit 0b48570

Please sign in to comment.