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

Commit abfaf8d

Browse files
Michael-Howesmergify[bot]
authored andcommitted
refactor(group_theory/abelianization): simplify abelianization (#1126)
* feat(group_theory/conjugates) : define conjugates define group conjugates and normal closure * feat(algebra/order_functions): generalize strict_mono.monotone (#1022) * trying to merge * feat(group_theory\presented_group): define presented groups Presented groups are defined as a quotient of a free group by the normal subgroup the relations generate. * feat(group_theory\presented_group): define presented groups Presented groups are defined as a quotient of a free group by the normal subgroup the relations generate * Update src/group_theory/presented_group.lean Co-Authored-By: Keeley Hoek <keeley@hoek.io> * Uniqueness of extension * Tidied up to_group.unique * Removed unnecessary line * Changed naming * refactor(group_theory/abelianization): simplify abelianization The commutator of a group was previously defined using lists. Now it is defined using `normal_closure`. This change simplifies some of the proofs
1 parent bd2f35f commit abfaf8d

File tree

1 file changed

+14
-31
lines changed

1 file changed

+14
-31
lines changed

src/group_theory/abelianization.lean

Lines changed: 14 additions & 31 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
/-
22
Copyright (c) 2018 Kenny Lau. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Kenny Lau
4+
Authors: Kenny Lau, Michael Howes
55
66
The functor Grp → Ab which is the left adjoint
77
of the forgetful functor Ab → Grp.
@@ -15,25 +15,10 @@ universes u v
1515
variables (α : Type u) [group α]
1616

1717
def commutator : set α :=
18-
{ x | ∃ L : list α, (∀ z ∈ L, ∃ p q, p * q * p⁻¹ * q⁻¹ = z) ∧ L.prod = x }
18+
group.normal_closure {x | ∃ p q, p * q * p⁻¹ * q⁻¹ = x}
1919

2020
instance : normal_subgroup (commutator α) :=
21-
{ one_mem := ⟨[], by simp⟩,
22-
mul_mem := λ x y ⟨L1, HL1, HP1⟩ ⟨L2, HL2, HP2⟩,
23-
⟨L1 ++ L2, list.forall_mem_append.2 ⟨HL1, HL2⟩, by simp*⟩,
24-
inv_mem := λ x ⟨L, HL, HP⟩, ⟨L.reverse.map has_inv.inv,
25-
λ x hx, let ⟨y, h3, h4⟩ := list.exists_of_mem_map hx in
26-
let ⟨p, q, h5⟩ := HL y (list.mem_reverse.1 h3) in
27-
⟨q, p, by rw [← h4, ← h5]; simp [mul_assoc]⟩,
28-
by rw ← HP; from list.rec_on L (by simp) (λ hd tl ih,
29-
by rw [list.reverse_cons, list.map_append, list.prod_append, ih]; simp)⟩,
30-
normal := λ x ⟨L, HL, HP⟩ g, ⟨L.map $ λ z, g * z * g⁻¹,
31-
λ x hx, let ⟨y, h3, h4⟩ := list.exists_of_mem_map hx in
32-
let ⟨p, q, h5⟩ := HL y h3 in
33-
⟨g * p * g⁻¹, g * q * g⁻¹,
34-
by rw [← h4, ← h5]; simp [mul_assoc]⟩,
35-
by rw ← HP; from list.rec_on L (by simp) (λ hd tl ih,
36-
by rw [list.map_cons, list.prod_cons, ih]; simp [mul_assoc])⟩, }
21+
group.normal_closure.is_normal
3722

3823
def abelianization : Type u :=
3924
quotient_group.quotient $ commutator α
@@ -43,11 +28,9 @@ namespace abelianization
4328
local attribute [instance] quotient_group.left_rel normal_subgroup.to_is_subgroup
4429

4530
instance : comm_group (abelianization α) :=
46-
{ mul_comm := λ x y, quotient.induction_on₂ x y $ λ m n,
47-
quotient.sound $ ⟨[n⁻¹*m⁻¹*n*m],
48-
by simp; refine ⟨n⁻¹, m⁻¹, _⟩; simp,
49-
by simp [mul_assoc]⟩,
50-
.. quotient_group.group _ }
31+
{ mul_comm := λ x y, quotient.induction_on₂ x y $ λ a b, quotient.sound
32+
(group.subset_normal_closure ⟨b⁻¹,a⁻¹, by simp [mul_inv_rev, inv_inv, mul_assoc]⟩),
33+
.. quotient_group.group _}
5134

5235
variable {α}
5336

@@ -61,24 +44,24 @@ section lift
6144

6245
variables {β : Type v} [comm_group β] (f : α → β) [is_group_hom f]
6346

47+
lemma commutator_subset_ker : commutator α ⊆ is_group_hom.ker f :=
48+
group.normal_closure_subset (λ x ⟨p,q,w⟩, (is_group_hom.mem_ker f).2
49+
(by {rw ←w, simp [is_group_hom.map_mul f, is_group_hom.map_inv f, mul_comm]}))
50+
6451
def lift : abelianization α → β :=
65-
quotient_group.lift _ f $ λ x ⟨L, HL, hx⟩,
66-
hx ▸ list.rec_on L (λ _, is_group_hom.map_one f) (λ hd tl HL ih,
67-
by rw [list.forall_mem_cons] at ih;
68-
rcases ih with ⟨⟨p, q, hpq⟩, ih⟩;
69-
specialize HL ih; rw [list.prod_cons, is_group_hom.map_mul f, ← hpq, HL];
70-
simp [is_group_hom.map_mul f, is_group_hom.map_inv f, mul_comm]) HL
52+
quotient_group.lift _ f (λ x h, (is_group_hom.mem_ker f).1 (commutator_subset_ker f h))
7153

7254
instance lift.is_group_hom : is_group_hom (lift f) :=
7355
quotient_group.is_group_hom_quotient_lift _ _ _
7456

75-
@[simp] lemma lift.of (x : α) : lift f (of x) = f x := rfl
57+
@[simp] lemma lift.of (x : α) : lift f (of x) = f x :=
58+
rfl
7659

7760
theorem lift.unique
7861
(g : abelianization α → β) [is_group_hom g]
7962
(hg : ∀ x, g (of x) = f x) {x} :
8063
g x = lift f x :=
81-
quotient.induction_on x $ λ m, hg m
64+
quotient_group.induction_on x hg
8265

8366
end lift
8467

0 commit comments

Comments
 (0)