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

Commit de20a39

Browse files
committed
feat(group_theory/subroup,ring_theory/ideal/operations): lift_of_surjective (#3888)
Surjective homomorphisms behave like quotient maps
1 parent 045b6c7 commit de20a39

File tree

5 files changed

+150
-6
lines changed

5 files changed

+150
-6
lines changed

src/group_theory/abelianization.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -66,7 +66,7 @@ begin
6666
end
6767

6868
def lift : abelianization G →* A :=
69-
quotient_group.lift _ f (λ x h, monoid_hom.mem_ker.2 $ commutator_subset_ker _ h)
69+
quotient_group.lift _ f (λ x h, f.mem_ker.2 $ commutator_subset_ker _ h)
7070

7171
@[simp] lemma lift.of (x : G) : lift f (of x) = f x :=
7272
rfl

src/group_theory/presented_group.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -41,10 +41,10 @@ variable (h : ∀ r ∈ rels, F r = 1)
4141

4242
-- FIXME why is apply_instance needed here? surely this should be a [] argument in `subgroup.normal_closure_le_normal`
4343
lemma closure_rels_subset_ker : subgroup.normal_closure rels ≤ monoid_hom.ker F :=
44-
subgroup.normal_closure_le_normal (by apply_instance) (λ x w, monoid_hom.mem_ker.2 (h x w))
44+
subgroup.normal_closure_le_normal (by apply_instance) (λ x w, (monoid_hom.mem_ker _).2 (h x w))
4545

4646
lemma to_group_eq_one_of_mem_closure : ∀ x ∈ subgroup.normal_closure rels, F x = 1 :=
47-
λ x w, monoid_hom.mem_ker.1 $ closure_rels_subset_ker h w
47+
λ x w, (monoid_hom.mem_ker _).1 $ closure_rels_subset_ker h w
4848

4949
/-- The extension of a map f : α → β that satisfies the given relations to a group homomorphism
5050
from presented_group rels → β. -/

src/group_theory/quotient_group.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -123,7 +123,7 @@ open function monoid_hom
123123
@[to_additive quotient_add_group.ker_lift "The induced map from the quotient by the kernel to the
124124
codomain."]
125125
def ker_lift : quotient (ker φ) →* H :=
126-
lift _ φ $ λ g, mem_ker.mp
126+
lift _ φ $ λ g, φ.mem_ker.mp
127127

128128
@[simp, to_additive quotient_add_group.ker_lift_mk]
129129
lemma ker_lift_mk (g : G) : (ker_lift φ) g = φ g :=
@@ -147,7 +147,7 @@ show a⁻¹ * b ∈ ker φ, by rw [mem_ker,
147147
@[to_additive quotient_add_group.range_ker_lift "The induced map from the quotient by the kernel to
148148
the range."]
149149
def range_ker_lift : quotient (ker φ) →* φ.range :=
150-
lift _ (to_range φ) $ λ g hg, mem_ker.mp $ by rwa to_range_ker
150+
lift _ (to_range φ) $ λ g hg, (mem_ker _).mp $ by rwa to_range_ker
151151

152152
@[to_additive quotient_add_group.range_ker_lift_injective]
153153
lemma range_ker_lift_injective : injective (range_ker_lift φ) :=

src/group_theory/subgroup.lean

Lines changed: 83 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -942,7 +942,7 @@ such that `f x = 0`"]
942942
def ker (f : G →* N) := (⊥ : subgroup N).comap f
943943

944944
@[to_additive]
945-
lemma mem_ker {f : G →* N} {x : G} : x ∈ f.ker ↔ f x = 1 := iff.rfl
945+
lemma mem_ker (f : G →* N) {x : G} : x ∈ f.ker ↔ f x = 1 := iff.rfl
946946

947947
@[to_additive]
948948
lemma comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker := rfl
@@ -995,6 +995,88 @@ le_antisymm
995995

996996
end monoid_hom
997997

998+
namespace monoid_hom
999+
1000+
variables {G₁ G₂ G₃ : Type*} [group G₁] [group G₂] [group G₃]
1001+
variables (f : G₁ →* G₂)
1002+
1003+
/-- `lift_of_surjective f hf g hg` is the unique group homomorphism `φ`
1004+
1005+
* such that `φ.comp f = g` (`lift_of_surjective_comp`),
1006+
* where `f : G₁ →+* G₂` is surjective (`hf`),
1007+
* and `g : G₂ →+* G₃` satisfies `hg : f.ker ≤ g.ker`.
1008+
1009+
See `lift_of_surjective_eq` for the uniqueness lemma.
1010+
1011+
```
1012+
G₁.
1013+
| \
1014+
f | \ g
1015+
| \
1016+
v \⌟
1017+
G₂----> G₃
1018+
∃!φ
1019+
```
1020+
-/
1021+
@[to_additive "`lift_of_surjective f hf g hg` is the unique additive group homomorphism `φ`
1022+
1023+
* such that `φ.comp f = g` (`lift_of_surjective_comp`),
1024+
* where `f : G₁ →+* G₂` is surjective (`hf`),
1025+
* and `g : G₂ →+* G₃` satisfies `hg : f.ker ≤ g.ker`.
1026+
1027+
See `lift_of_surjective_eq` for the uniqueness lemma.
1028+
1029+
```
1030+
G₁.
1031+
| \\
1032+
f | \\ g
1033+
| \\
1034+
v \\
1035+
G₂----> G₃
1036+
∃!φ
1037+
```"]
1038+
noncomputable def lift_of_surjective
1039+
(hf : function.surjective f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) :
1040+
G₂ →* G₃ :=
1041+
{ to_fun := λ b, g (classical.some (hf b)),
1042+
map_one' := hg (classical.some_spec (hf 1)),
1043+
map_mul' :=
1044+
begin
1045+
intros x y,
1046+
rw [← g.map_mul, ← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker],
1047+
apply hg,
1048+
rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one, f.map_mul],
1049+
simp only [classical.some_spec (hf _)],
1050+
end }
1051+
1052+
@[simp, to_additive]
1053+
lemma lift_of_surjective_comp_apply
1054+
(hf : function.surjective f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) (x : G₁) :
1055+
(f.lift_of_surjective hf g hg) (f x) = g x :=
1056+
begin
1057+
dsimp [lift_of_surjective],
1058+
rw [← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker],
1059+
apply hg,
1060+
rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one],
1061+
simp only [classical.some_spec (hf _)],
1062+
end
1063+
1064+
@[simp, to_additive]
1065+
lemma lift_of_surjective_comp (hf : function.surjective f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) :
1066+
(f.lift_of_surjective hf g hg).comp f = g :=
1067+
by { ext, simp only [comp_apply, lift_of_surjective_comp_apply] }
1068+
1069+
@[to_additive]
1070+
lemma eq_lift_of_surjective (hf : function.surjective f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker)
1071+
(h : G₂ →* G₃) (hh : h.comp f = g) :
1072+
h = (f.lift_of_surjective hf g hg) :=
1073+
begin
1074+
ext b, rcases hf b with ⟨a, rfl⟩,
1075+
simp only [← comp_apply, hh, f.lift_of_surjective_comp],
1076+
end
1077+
1078+
end monoid_hom
1079+
9981080
variables {N : Type*} [group N]
9991081

10001082
@[to_additive]

src/ring_theory/ideal/operations.lean

Lines changed: 62 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -875,3 +875,65 @@ instance semimodule_submodule : semimodule (ideal R) (submodule R M) :=
875875
smul_zero := smul_bot }
876876

877877
end submodule
878+
879+
namespace ring_hom
880+
variables {A B C : Type*} [comm_ring A] [comm_ring B] [comm_ring C]
881+
variables (f : A →+* B)
882+
883+
/-- `lift_of_surjective f hf g hg` is the unique ring homomorphism `φ`
884+
885+
* such that `φ.comp f = g` (`lift_of_surjective_comp`),
886+
* where `f : A →+* B` is surjective (`hf`),
887+
* and `g : B →+* C` satisfies `hg : f.ker ≤ g.ker`.
888+
889+
See `lift_of_surjective_eq` for the uniqueness lemma.
890+
891+
```
892+
A .
893+
| \
894+
f | \ g
895+
| \
896+
v \⌟
897+
B ----> C
898+
∃!φ
899+
```
900+
-/
901+
noncomputable def lift_of_surjective
902+
(hf : function.surjective f) (g : A →+* C) (hg : f.ker ≤ g.ker) :
903+
B →+* C :=
904+
{ to_fun := λ b, g (classical.some (hf b)),
905+
map_one' :=
906+
begin
907+
rw [← g.map_one, ← sub_eq_zero, ← g.map_sub, ← g.mem_ker],
908+
apply hg,
909+
rw [f.mem_ker, f.map_sub, sub_eq_zero, f.map_one],
910+
exact classical.some_spec (hf 1)
911+
end,
912+
map_mul' :=
913+
begin
914+
intros x y,
915+
rw [← g.map_mul, ← sub_eq_zero, ← g.map_sub, ← g.mem_ker],
916+
apply hg,
917+
rw [f.mem_ker, f.map_sub, sub_eq_zero, f.map_mul],
918+
simp only [classical.some_spec (hf _)],
919+
end,
920+
.. add_monoid_hom.lift_of_surjective f.to_add_monoid_hom hf g.to_add_monoid_hom hg }
921+
922+
@[simp] lemma lift_of_surjective_comp_apply
923+
(hf : function.surjective f) (g : A →+* C) (hg : f.ker ≤ g.ker) (a : A) :
924+
(f.lift_of_surjective hf g hg) (f a) = g a :=
925+
f.to_add_monoid_hom.lift_of_surjective_comp_apply hf g.to_add_monoid_hom hg a
926+
927+
@[simp] lemma lift_of_surjective_comp (hf : function.surjective f) (g : A →+* C) (hg : f.ker ≤ g.ker) :
928+
(f.lift_of_surjective hf g hg).comp f = g :=
929+
by { ext, simp only [comp_apply, lift_of_surjective_comp_apply] }
930+
931+
lemma eq_lift_of_surjective (hf : function.surjective f) (g : A →+* C) (hg : f.ker ≤ g.ker)
932+
(h : B →+* C) (hh : h.comp f = g) :
933+
h = (f.lift_of_surjective hf g hg) :=
934+
begin
935+
ext b, rcases hf b with ⟨a, rfl⟩,
936+
simp only [← comp_apply, hh, f.lift_of_surjective_comp],
937+
end
938+
939+
end ring_hom

0 commit comments

Comments
 (0)