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

Commit cf8e77c

Browse files
committed
feat(group_theory/subgroup/basic): Map a subgroup under an iso (#17775)
Cross-reference `mul_equiv.subgroup_map` and `subgroup.equiv_map_of_injective`. Add lemmas connecting them.
1 parent d794388 commit cf8e77c

File tree

3 files changed

+51
-29
lines changed

3 files changed

+51
-29
lines changed

src/group_theory/commensurable.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -52,10 +52,10 @@ lemma equivalence : equivalence (@commensurable G _) :=
5252
/--Equivalence of `K/H ⊓ K` with `gKg⁻¹/gHg⁻¹ ⊓ gKg⁻¹`-/
5353
def quot_conj_equiv (H K : subgroup G) (g : conj_act G) :
5454
K ⧸ (H.subgroup_of K) ≃ (g • K).1 ⧸ ((g • H).subgroup_of (g • K)) :=
55-
quotient.congr (K.equiv_smul g).to_equiv (λ a b, by { rw [←quotient.eq', ←quotient.eq',
55+
quotient.congr (K.equiv_smul g).to_equiv (λ a b, by { dsimp, rw [←quotient.eq', ←quotient.eq',
5656
quotient_group.eq', quotient_group.eq', subgroup.mem_subgroup_of, subgroup.mem_subgroup_of,
57-
mul_equiv.to_equiv_eq_coe, mul_equiv.coe_to_equiv, ←mul_equiv.map_inv, ←mul_equiv.map_mul,
58-
subgroup.equiv_smul_apply_coe, subgroup.smul_mem_pointwise_smul_iff] })
57+
←mul_equiv.map_inv, ←mul_equiv.map_mul, subgroup.equiv_smul_apply_coe],
58+
exact subgroup.smul_mem_pointwise_smul_iff.symm })
5959

6060
lemma commensurable_conj {H K : subgroup G} (g : conj_act G) :
6161
commensurable H K ↔ commensurable (g • H) (g • K) :=

src/group_theory/subgroup/basic.lean

Lines changed: 29 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -87,9 +87,10 @@ membership of a subgroup's underlying set.
8787
subgroup, subgroups
8888
-/
8989

90+
open function
9091
open_locale big_operators
9192

92-
variables {G : Type*} [group G]
93+
variables {G G' : Type*} [group G] [group G']
9394
variables {A : Type*} [add_group A]
9495

9596
section subgroup_class
@@ -1781,16 +1782,16 @@ instance is_commutative.comm_group [h : H.is_commutative] : comm_group H :=
17811782
instance center.is_commutative : (center G).is_commutative :=
17821783
⟨⟨λ a b, subtype.ext (b.2 a)⟩⟩
17831784

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 :=
17861787
⟨⟨begin
17871788
rintros ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩,
17881789
rw [subtype.ext_iff, coe_mul, coe_mul, subtype.coe_mk, subtype.coe_mk, ←map_mul, ←map_mul],
17891790
exact congr_arg f (subtype.ext_iff.mp (mul_comm ⟨a, ha⟩ ⟨b, hb⟩)),
17901791
end⟩⟩
17911792

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 :=
17941795
⟨⟨λ a b, subtype.ext begin
17951796
have := mul_comm (⟨f a, a.2⟩ : H) (⟨f b, b.2⟩ : H),
17961797
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
24052406
codisjoint (H.subgroup_of (H ⊔ K)) (K.subgroup_of (H ⊔ K)) :=
24062407
by { rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self], exacts [le_sup_left, le_sup_right] }
24072408

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."]
24102413
noncomputable def equiv_map_of_injective (H : subgroup G)
24112414
(f : G →* N) (hf : function.injective f) : H ≃* H.map f :=
24122415
{ map_mul' := λ _ _, subtype.ext (f.map_mul _ _), ..equiv.set.image f H hf }
@@ -2764,8 +2767,6 @@ end subgroup
27642767

27652768
namespace monoid_hom
27662769

2767-
variables {G' : Type*} [group G']
2768-
27692770
/-- The `monoid_hom` from the preimage of a subgroup to itself. -/
27702771
@[to_additive "the `add_monoid_hom` from the preimage of an additive subgroup to itself.", simps]
27712772
def subgroup_comap (f : G →* G') (H' : subgroup G') : H'.comap f →* H' :=
@@ -2784,7 +2785,6 @@ f.submonoid_map_surjective H.to_submonoid
27842785
end monoid_hom
27852786

27862787
namespace mul_equiv
2787-
27882788
variables {H K : subgroup G}
27892789

27902790
/-- Makes the identity isomorphism from a proof two subgroups of a multiplicative
@@ -2794,20 +2794,30 @@ two subgroups of an additive group are equal."]
27942794
def subgroup_congr (h : H = K) : H ≃* K :=
27952795
{ map_mul' := λ _ _, rfl, ..equiv.set_congr $ congr_arg _ h }
27962796

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
28042803

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
28062811

2807-
-- TODO : ↥(⊤ : subgroup H) ≃* H ?
2812+
end mul_equiv
28082813

28092814
namespace subgroup
28102815

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+
28112821
variables {C : Type*} [comm_group C] {s t : subgroup C} {x : C}
28122822

28132823
@[to_additive]

src/group_theory/submonoid/operations.lean

Lines changed: 19 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -605,8 +605,10 @@ def top_equiv : (⊤ : submonoid M) ≃* M :=
605605
(top_equiv : _ ≃* M).to_monoid_hom = (⊤ : submonoid M).subtype :=
606606
rfl
607607

608-
/-- A submonoid is isomorphic to its image under an injective function -/
609-
@[to_additive "An additive submonoid is isomorphic to its image under an injective function"]
608+
/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism,
609+
use `mul_equiv.submonoid_map` for better definitional equalities. -/
610+
@[to_additive "An additive subgroup is isomorphic to its image under an injective function. If you
611+
have an isomorphism, use `add_equiv.add_submonoid_map` for better definitional equalities."]
610612
noncomputable def equiv_map_of_injective
611613
(f : M →* N) (hf : function.injective f) : S ≃* S.map f :=
612614
{ map_mul' := λ _ _, subtype.ext (f.map_mul _ _), ..equiv.set.image f S hf }
@@ -1055,15 +1057,25 @@ a submonoid `S ≤ M` and the submonoid `φ(S) ≤ N`.
10551057
See `monoid_hom.submonoid_map` for a variant for `monoid_hom`s. -/
10561058
@[to_additive "An `add_equiv` `φ` between two additive monoids `M` and `N` induces an `add_equiv`
10571059
between a submonoid `S ≤ M` and the submonoid `φ(S) ≤ N`. See `add_monoid_hom.add_submonoid_map`
1058-
for a variant for `add_monoid_hom`s.", simps]
1060+
for a variant for `add_monoid_hom`s."]
10591061
def submonoid_map (e : M ≃* N) (S : submonoid M) : S ≃* S.map e.to_monoid_hom :=
1060-
{ to_fun := λ x, ⟨e x, _⟩,
1061-
inv_fun := λ x, ⟨e.symm x, _⟩, -- we restate this for `simps` to avoid `⇑e.symm.to_equiv x`
1062-
..e.to_monoid_hom.submonoid_map S,
1063-
..e.to_equiv.image S }
1062+
{ map_mul' := λ _ _, subtype.ext (map_mul e _ _), ..(e : M ≃ N).image S }
1063+
1064+
@[simp, to_additive]
1065+
lemma coe_submonoid_map_apply (e : M ≃* N) (S : submonoid M) (g : S) :
1066+
((submonoid_map e S g : S.map (e : M →* N)) : N) = e g := rfl
1067+
1068+
@[simp, to_additive add_equiv.add_submonoid_map_symm_apply]
1069+
lemma submonoid_map_symm_apply (e : M ≃* N) (S : submonoid M) (g : S.map (e : M →* N)) :
1070+
(e.submonoid_map S).symm g = ⟨e.symm g, set_like.mem_coe.1 $ set.mem_image_equiv.1 g.2⟩ := rfl
10641071

10651072
end mul_equiv
10661073

1074+
@[simp, to_additive]
1075+
lemma submonoid.equiv_map_of_injective_coe_mul_equiv (e : M ≃* N) :
1076+
S.equiv_map_of_injective (e : M →* N) (equiv_like.injective e) = e.submonoid_map S :=
1077+
by { ext, refl }
1078+
10671079
section actions
10681080
/-! ### Actions by `submonoid`s
10691081

0 commit comments

Comments
 (0)