@@ -87,9 +87,10 @@ membership of a subgroup's underlying set.
87
87
subgroup, subgroups
88
88
-/
89
89
90
+ open function
90
91
open_locale big_operators
91
92
92
- variables {G : Type *} [group G]
93
+ variables {G G' : Type *} [group G] [group G' ]
93
94
variables {A : Type *} [add_group A]
94
95
95
96
section subgroup_class
@@ -1781,16 +1782,16 @@ instance is_commutative.comm_group [h : H.is_commutative] : comm_group H :=
1781
1782
instance center.is_commutative : (center G).is_commutative :=
1782
1783
⟨⟨λ a b, subtype.ext (b.2 a)⟩⟩
1783
1784
1784
- @[to_additive] instance map_is_commutative {G' : Type *} [group G'] (f : G →* G')
1785
- [H.is_commutative] : (H.map f).is_commutative :=
1785
+ @[to_additive] instance map_is_commutative (f : G →* G') [H.is_commutative] :
1786
+ (H.map f).is_commutative :=
1786
1787
⟨⟨begin
1787
1788
rintros ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩,
1788
1789
rw [subtype.ext_iff, coe_mul, coe_mul, subtype.coe_mk, subtype.coe_mk, ←map_mul, ←map_mul],
1789
1790
exact congr_arg f (subtype.ext_iff.mp (mul_comm ⟨a, ha⟩ ⟨b, hb⟩)),
1790
1791
end ⟩⟩
1791
1792
1792
- @[to_additive] lemma comap_injective_is_commutative {G' : Type *} [group G'] { f : G' →* G}
1793
- (hf : function.injective f) [H.is_commutative] : (H.comap f).is_commutative :=
1793
+ @[to_additive] lemma comap_injective_is_commutative {f : G' →* G} (hf : injective f)
1794
+ [H.is_commutative] : (H.comap f).is_commutative :=
1794
1795
⟨⟨λ a b, subtype.ext begin
1795
1796
have := mul_comm (⟨f a, a.2 ⟩ : H) (⟨f b, b.2 ⟩ : H),
1796
1797
rwa [subtype.ext_iff, coe_mul, coe_mul, coe_mk, coe_mk, ←map_mul, ←map_mul, hf.eq_iff] at this ,
@@ -2405,8 +2406,10 @@ comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.sub
2405
2406
codisjoint (H.subgroup_of (H ⊔ K)) (K.subgroup_of (H ⊔ K)) :=
2406
2407
by { rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self], exacts [le_sup_left, le_sup_right] }
2407
2408
2408
- /-- A subgroup is isomorphic to its image under an injective function -/
2409
- @[to_additive " An additive subgroup is isomorphic to its image under an injective function" ]
2409
+ /-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism,
2410
+ use `mul_equiv.subgroup_map` for better definitional equalities. -/
2411
+ @[to_additive " An additive subgroup is isomorphic to its image under an injective function. If you
2412
+ have an isomorphism, use `add_equiv.add_subgroup_map` for better definitional equalities." ]
2410
2413
noncomputable def equiv_map_of_injective (H : subgroup G)
2411
2414
(f : G →* N) (hf : function.injective f) : H ≃* H.map f :=
2412
2415
{ map_mul' := λ _ _, subtype.ext (f.map_mul _ _), ..equiv.set.image f H hf }
@@ -2764,8 +2767,6 @@ end subgroup
2764
2767
2765
2768
namespace monoid_hom
2766
2769
2767
- variables {G' : Type *} [group G']
2768
-
2769
2770
/-- The `monoid_hom` from the preimage of a subgroup to itself. -/
2770
2771
@[to_additive " the `add_monoid_hom` from the preimage of an additive subgroup to itself." , simps]
2771
2772
def subgroup_comap (f : G →* G') (H' : subgroup G') : H'.comap f →* H' :=
@@ -2784,7 +2785,6 @@ f.submonoid_map_surjective H.to_submonoid
2784
2785
end monoid_hom
2785
2786
2786
2787
namespace mul_equiv
2787
-
2788
2788
variables {H K : subgroup G}
2789
2789
2790
2790
/-- Makes the identity isomorphism from a proof two subgroups of a multiplicative
@@ -2794,20 +2794,30 @@ two subgroups of an additive group are equal."]
2794
2794
def subgroup_congr (h : H = K) : H ≃* K :=
2795
2795
{ map_mul' := λ _ _, rfl, ..equiv.set_congr $ congr_arg _ h }
2796
2796
2797
- /-- A `mul_equiv` `φ` between two groups `G` and `G'` induces a `mul_equiv` between
2798
- a subgroup `H ≤ G` and the subgroup `φ(H) ≤ G'`. -/
2799
- @[to_additive " An `add_equiv` `φ` between two additive groups `G` and `G'` induces an `add_equiv`
2800
- between a subgroup `H ≤ G` and the subgroup `φ(H) ≤ G'`. " ]
2801
- def subgroup_map {G'} [group G'] (e : G ≃* G') (H : subgroup G) :
2802
- H ≃* H.map e.to_monoid_hom :=
2803
- e.submonoid_map H.to_submonoid
2797
+ /-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map,
2798
+ use `subgroup.equiv_map_of_injective`. -/
2799
+ @[to_additive " An additive subgroup is isomorphic to its image under an an isomorphism. If you only
2800
+ have an injective map, use `add_subgroup.equiv_map_of_injective`." ]
2801
+ def subgroup_map (e : G ≃* G') (H : subgroup G) : H ≃* H.map (e : G →* G') :=
2802
+ mul_equiv.submonoid_map (e : G ≃* G') H.to_submonoid
2804
2803
2805
- end mul_equiv
2804
+ @[simp, to_additive]
2805
+ lemma coe_subgroup_map_apply (e : G ≃* G') (H : subgroup G) (g : H) :
2806
+ ((subgroup_map e H g : H.map (e : G →* G')) : G') = e g := rfl
2807
+
2808
+ @[simp, to_additive]
2809
+ lemma subgroup_map_symm_apply (e : G ≃* G') (H : subgroup G) (g : H.map (e : G →* G')) :
2810
+ (e.subgroup_map H).symm g = ⟨e.symm g, set_like.mem_coe.1 $ set.mem_image_equiv.1 g.2 ⟩ := rfl
2806
2811
2807
- -- TODO : ↥(⊤ : subgroup H) ≃* H ?
2812
+ end mul_equiv
2808
2813
2809
2814
namespace subgroup
2810
2815
2816
+ @[simp, to_additive]
2817
+ lemma equiv_map_of_injective_coe_mul_equiv (H : subgroup G) (e : G ≃* G') :
2818
+ H.equiv_map_of_injective (e : G →* G') (equiv_like.injective e) = e.subgroup_map H :=
2819
+ by { ext, refl }
2820
+
2811
2821
variables {C : Type *} [comm_group C] {s t : subgroup C} {x : C}
2812
2822
2813
2823
@[to_additive]
0 commit comments