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

Commit 8004fb6

Browse files
committed
chore(topology/algebra/group): move a lemma to group_theory/coset (#4522)
`quotient_group_saturate` didn't use any topology. Move it to `group_theory/coset` and rename to `quotient_group.preimage_image_coe`. Also rename `quotient_group.open_coe` to `quotient_group.is_open_map_coe`
1 parent ce999a8 commit 8004fb6

File tree

3 files changed

+24
-29
lines changed

3 files changed

+24
-29
lines changed

src/group_theory/coset.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -214,6 +214,16 @@ lemma eq_class_eq_left_coset (s : subgroup α) (g : α) :
214214
{x : α | (x : quotient s) = g} = left_coset g s :=
215215
set.ext $ λ z, by { rw [mem_left_coset_iff, set.mem_set_of_eq, eq_comm, quotient_group.eq], simp }
216216

217+
@[to_additive]
218+
lemma preimage_image_coe (N : subgroup α) (s : set α) :
219+
coe ⁻¹' ((coe : α → quotient N) '' s) = ⋃ x : N, (λ y : α, y * x) '' s :=
220+
begin
221+
ext x,
222+
simp only [quotient_group.eq, subgroup.exists, exists_prop, set.mem_preimage, set.mem_Union,
223+
set.mem_image, subgroup.coe_mk, ← eq_inv_mul_iff_mul_eq],
224+
exact ⟨λ ⟨y, hs, hN⟩, ⟨_, hN, y, hs, rfl⟩, λ ⟨z, hN, y, hs, hyz⟩, ⟨y, hs, hyz ▸ hN⟩⟩
225+
end
226+
217227
end quotient_group
218228

219229
namespace subgroup

src/group_theory/quotient_group.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -63,23 +63,23 @@ instance {G : Type*} [comm_group G] (N : subgroup G) : comm_group (quotient N) :
6363

6464
include nN
6565

66+
local notation ` Q ` := quotient N
67+
6668
@[simp, to_additive quotient_add_group.coe_zero]
67-
lemma coe_one : ((1 : G) : quotient N) = 1 := rfl
69+
lemma coe_one : ((1 : G) : Q) = 1 := rfl
6870

6971
@[simp, to_additive quotient_add_group.coe_add]
70-
lemma coe_mul (a b : G) : ((a * b : G) : quotient N) = a * b := rfl
72+
lemma coe_mul (a b : G) : ((a * b : G) : Q) = a * b := rfl
7173

7274
@[simp, to_additive quotient_add_group.coe_neg]
73-
lemma coe_inv (a : G) : ((a⁻¹ : G) : quotient N) = a⁻¹ := rfl
75+
lemma coe_inv (a : G) : ((a⁻¹ : G) : Q) = a⁻¹ := rfl
7476

75-
@[simp] lemma coe_pow (a : G) (n : ℕ) : ((a ^ n : G) : quotient N) = a ^ n :=
77+
@[simp] lemma coe_pow (a : G) (n : ℕ) : ((a ^ n : G) : Q) = a ^ n :=
7678
(mk' N).map_pow a n
7779

78-
@[simp] lemma coe_gpow (a : G) (n : ℤ) : ((a ^ n : G) : quotient N) = a ^ n :=
80+
@[simp] lemma coe_gpow (a : G) (n : ℤ) : ((a ^ n : G) : Q) = a ^ n :=
7981
(mk' N).map_gpow a n
8082

81-
local notation ` Q ` := quotient N
82-
8383
/-- A group homomorphism `φ : G →* H` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a
8484
group homomorphism `G/N →* H`. -/
8585
@[to_additive quotient_add_group.lift "An `add_group` homomorphism `φ : G →+ H` with `N ⊆ ker(φ)`

src/topology/algebra/group.lean

Lines changed: 7 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -198,30 +198,17 @@ variables [topological_space G] [group G] [topological_group G] (N : subgroup G)
198198
@[to_additive]
199199
instance {G : Type u} [group G] [topological_space G] (N : subgroup G) :
200200
topological_space (quotient_group.quotient N) :=
201-
by dunfold quotient_group.quotient; apply_instance
201+
quotient.topological_space
202202

203203
open quotient_group
204-
@[to_additive]
205-
lemma quotient_group_saturate {G : Type u} [group G] (N : subgroup G) (s : set G) :
206-
(coe : G → quotient N) ⁻¹' ((coe : G → quotient N) '' s) = (⋃ x : N, (λ y, y*x.1) '' s) :=
207-
begin
208-
ext x,
209-
simp only [mem_preimage, mem_image, mem_Union, quotient_group.eq],
210-
split,
211-
{ exact assume ⟨a, a_in, h⟩, ⟨⟨_, h⟩, a, a_in, mul_inv_cancel_left _ _⟩ },
212-
{ exact assume ⟨⟨i, hi⟩, a, ha, eq⟩,
213-
⟨a, ha, by { simp only [eq.symm, (mul_assoc _ _ _).symm, inv_mul_cancel_left], exact hi }⟩ }
214-
end
215204

216205
@[to_additive]
217-
lemma quotient_group.open_coe : is_open_map (coe : G → quotient N) :=
206+
lemma quotient_group.is_open_map_coe : is_open_map (coe : G → quotient N) :=
218207
begin
219208
intros s s_op,
220209
change is_open ((coe : G → quotient N) ⁻¹' (coe '' s)),
221-
rw quotient_group_saturate N s,
222-
apply is_open_Union,
223-
rintro ⟨n, _⟩,
224-
exact is_open_map_mul_right n s s_op
210+
rw quotient_group.preimage_image_coe N s,
211+
exact is_open_Union (λ n, is_open_map_mul_right n s s_op)
225212
end
226213

227214
@[to_additive]
@@ -231,11 +218,9 @@ instance topological_group_quotient (n : N.normal) : topological_group (quotient
231218
continuous_quot_mk.comp continuous_mul,
232219
have quot : quotient_map (λ p : G × G, ((p.1:quotient N), (p.2:quotient N))),
233220
{ apply is_open_map.to_quotient_map,
234-
{ exact is_open_map.prod (quotient_group.open_coe N) (quotient_group.open_coe N) },
235-
{ exact (continuous_quot_mk.comp continuous_fst).prod_mk
236-
(continuous_quot_mk.comp continuous_snd) },
237-
{ rintro ⟨⟨x⟩, ⟨y⟩⟩,
238-
exact ⟨(x, y), rfl⟩ } },
221+
{ exact (quotient_group.is_open_map_coe N).prod (quotient_group.is_open_map_coe N) },
222+
{ exact continuous_quot_mk.prod_map continuous_quot_mk },
223+
{ exact (surjective_quot_mk _).prod_map (surjective_quot_mk _) } },
239224
exact (quotient_map.continuous_iff quot).2 cont,
240225
end,
241226
continuous_inv := begin

0 commit comments

Comments
 (0)