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

Commit 83f81ae

Browse files
committed
refactor(topology/algebra/open_subgroup): use set_like (#18585)
1 parent 992efbd commit 83f81ae

File tree

3 files changed

+83
-89
lines changed

3 files changed

+83
-89
lines changed

src/field_theory/krull_topology.lean

Lines changed: 1 addition & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -210,9 +210,7 @@ begin
210210
have h_basis : E.fixing_subgroup.carrier ∈ (gal_group_basis K L) :=
211211
⟨E.fixing_subgroup, ⟨E, ‹_›, rfl⟩, rfl⟩,
212212
have h_nhd := group_filter_basis.mem_nhds_one (gal_group_basis K L) h_basis,
213-
rw mem_nhds_iff at h_nhd,
214-
rcases h_nhd with ⟨U, hU_le, hU_open, h1U⟩,
215-
exact subgroup.is_open_of_one_mem_interior ⟨U, ⟨hU_open, hU_le⟩, h1U⟩,
213+
exact subgroup.is_open_of_mem_nhds _ h_nhd
216214
end
217215

218216
/-- Given a tower of fields `L/E/K`, with `E/K` finite, the subgroup `Gal(L/E) ≤ L ≃ₐ[K] L` is

src/topology/algebra/nonarchimedean/basic.lean

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -127,14 +127,13 @@ lemma mul_subset (U : open_add_subgroup R) :
127127
∃ V : open_add_subgroup R, (V : set R) * V ⊆ U :=
128128
let ⟨V, H⟩ := prod_self_subset (is_open.mem_nhds (is_open.preimage continuous_mul U.is_open)
129129
begin
130-
simpa only [set.mem_preimage, open_add_subgroup.mem_coe, prod.snd_zero, mul_zero]
131-
using U.zero_mem,
130+
simpa only [set.mem_preimage, set_like.mem_coe, prod.snd_zero, mul_zero] using U.zero_mem,
132131
end) in
133132
begin
134133
use V,
135134
rintros v ⟨a, b, ha, hb, hv⟩,
136135
have hy := H (set.mk_mem_prod ha hb),
137-
simp only [set.mem_preimage, open_add_subgroup.mem_coe] at hy,
136+
simp only [set.mem_preimage, set_like.mem_coe] at hy,
138137
rwa hv at hy
139138
end
140139

src/topology/algebra/open_subgroup.lean

Lines changed: 80 additions & 83 deletions
Original file line numberDiff line numberDiff line change
@@ -55,44 +55,39 @@ variables {G : Type*} [group G] [topological_space G]
5555
variables {U V : open_subgroup G} {g : G}
5656

5757
@[to_additive]
58-
instance has_coe_set : has_coe_t (open_subgroup G) (set G) := ⟨λ U, U.1
58+
instance has_coe_subgroup : has_coe_t (open_subgroup G) (subgroup G) := ⟨to_subgroup
5959

6060
@[to_additive]
61-
instance : has_mem G (open_subgroup G) := ⟨λ g U, g ∈ (U : set G)⟩
61+
lemma coe_subgroup_injective : injective (coe : open_subgroup G → subgroup G)
62+
| ⟨_, _⟩ ⟨_, _⟩ rfl := rfl
6263

6364
@[to_additive]
64-
instance has_coe_subgroup : has_coe_t (open_subgroup G) (subgroup G) := ⟨to_subgroup⟩
65+
instance : set_like (open_subgroup G) G :=
66+
{ coe := λ U, U.1,
67+
coe_injective' := λ _ _ h, coe_subgroup_injective $ set_like.ext' h }
68+
69+
@[to_additive]
70+
instance : subgroup_class (open_subgroup G) G :=
71+
{ mul_mem := λ U _ _, U.mul_mem',
72+
one_mem := λ U, U.one_mem',
73+
inv_mem := λ U _, U.inv_mem' }
6574

6675
@[to_additive]
6776
instance has_coe_opens : has_coe_t (open_subgroup G) (opens G) := ⟨λ U, ⟨U, U.is_open'⟩⟩
6877

69-
@[simp, norm_cast, to_additive] lemma mem_coe : g ∈ (U : set G) ↔ g ∈ U := iff.rfl
78+
@[simp, norm_cast, to_additive] lemma coe_coe_opens : ((U : opens G) : set G) = U := rfl
79+
@[simp, norm_cast, to_additive] lemma coe_coe_subgroup : ((U : subgroup G) : set G) = U := rfl
7080
@[simp, norm_cast, to_additive] lemma mem_coe_opens : g ∈ (U : opens G) ↔ g ∈ U := iff.rfl
7181
@[simp, norm_cast, to_additive]
7282
lemma mem_coe_subgroup : g ∈ (U : subgroup G) ↔ g ∈ U := iff.rfl
7383

74-
@[to_additive] lemma coe_injective : injective (coe : open_subgroup G → set G) :=
75-
by { rintros ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨h⟩, congr, }
76-
7784
@[ext, to_additive]
78-
lemma ext (h : ∀ x, x ∈ U ↔ x ∈ V) : (U = V) := coe_injective $ set.ext h
79-
80-
@[to_additive]
81-
lemma ext_iff : (U = V) ↔ (∀ x, x ∈ U ↔ x ∈ V) := ⟨λ h x, h ▸ iff.rfl, ext⟩
85+
lemma ext (h : ∀ x, x ∈ U ↔ x ∈ V) : (U = V) := set_like.ext h
8286

8387
variable (U)
8488
@[to_additive]
8589
protected lemma is_open : is_open (U : set G) := U.is_open'
8690

87-
@[to_additive]
88-
protected lemma one_mem : (1 : G) ∈ U := U.one_mem'
89-
90-
@[to_additive]
91-
protected lemma inv_mem {g : G} (h : g ∈ U) : g⁻¹ ∈ U := U.inv_mem' h
92-
93-
@[to_additive]
94-
protected lemma mul_mem {g₁ g₂ : G} (h₁ : g₁ ∈ U) (h₂ : g₂ ∈ U) : g₁ * g₂ ∈ U := U.mul_mem' h₁ h₂
95-
9691
@[to_additive]
9792
lemma mem_nhds_one : (U : set G) ∈ 𝓝 (1 : G) :=
9893
is_open.mem_nhds U.is_open U.one_mem
@@ -101,6 +96,15 @@ variable {U}
10196
@[to_additive]
10297
instance : has_top (open_subgroup G) := ⟨{ is_open' := is_open_univ, .. (⊤ : subgroup G) }⟩
10398

99+
@[simp, to_additive] lemma mem_top (x : G) : x ∈ (⊤ : open_subgroup G) := trivial
100+
@[simp, norm_cast, to_additive] lemma coe_top : ((⊤ : open_subgroup G) : set G) = set.univ := rfl
101+
102+
@[simp, norm_cast, to_additive]
103+
lemma coe_subgroup_top : ((⊤ : open_subgroup G) : subgroup G) = ⊤ := rfl
104+
105+
@[simp, norm_cast, to_additive]
106+
lemma coe_opens_top : ((⊤ : open_subgroup G) : opens G) = ⊤ := rfl
107+
104108
@[to_additive]
105109
instance : inhabited (open_subgroup G) := ⟨⊤⟩
106110

@@ -109,67 +113,79 @@ lemma is_closed [has_continuous_mul G] (U : open_subgroup G) : is_closed (U : se
109113
begin
110114
apply is_open_compl_iff.1,
111115
refine is_open_iff_forall_mem_open.2 (λ x hx, ⟨(λ y, y * x⁻¹) ⁻¹' U, _, _, _⟩),
112-
{ intros u hux,
113-
simp only [set.mem_preimage, set.mem_compl_iff, mem_coe] at hux hx ⊢,
114-
refine mt (λ hu, _) hx,
116+
{ refine λ u hux hu, hx _,
117+
simp only [set.mem_preimage, set_like.mem_coe] at hux hu ⊢,
115118
convert U.mul_mem (U.inv_mem hux) hu,
116119
simp },
117120
{ exact U.is_open.preimage (continuous_mul_right _) },
118-
{ simp [U.one_mem] }
121+
{ simp [one_mem] }
119122
end
120123

124+
@[to_additive]
125+
lemma is_clopen [has_continuous_mul G] (U : open_subgroup G) : is_clopen (U : set G) :=
126+
⟨U.is_open, U.is_closed⟩
127+
121128
section
122129
variables {H : Type*} [group H] [topological_space H]
123130

124131
/-- The product of two open subgroups as an open subgroup of the product group. -/
125132
@[to_additive "The product of two open subgroups as an open subgroup of the product group."]
126133
def prod (U : open_subgroup G) (V : open_subgroup H) : open_subgroup (G × H) :=
127-
{ carrier := U ×ˢ V,
128-
is_open' := U.is_open.prod V.is_open,
134+
{ is_open' := U.is_open.prod V.is_open,
129135
.. (U : subgroup G).prod (V : subgroup H) }
130136

137+
@[simp, norm_cast, to_additive] lemma coe_prod (U : open_subgroup G) (V : open_subgroup H) :
138+
(U.prod V : set (G × H)) = U ×ˢ V :=
139+
rfl
140+
141+
@[simp, norm_cast, to_additive]
142+
lemma coe_subgroup_prod (U : open_subgroup G) (V : open_subgroup H) :
143+
(U.prod V : subgroup (G × H)) = (U : subgroup G).prod V :=
144+
rfl
145+
131146
end
132147

133148
@[to_additive]
134-
instance : partial_order (open_subgroup G) :=
135-
{ le := λ U V, ∀ ⦃x⦄, x ∈ U → x ∈ V,
136-
.. partial_order.lift (coe : open_subgroup G → set G) coe_injective }
149+
instance : has_inf (open_subgroup G) :=
150+
⟨λ U V, ⟨U ⊓ V, U.is_open.inter V.is_open⟩⟩
151+
152+
@[simp, norm_cast, to_additive] lemma coe_inf : (↑(U ⊓ V) : set G) = (U : set G) ∩ V := rfl
153+
@[simp, norm_cast, to_additive] lemma coe_subgroup_inf : (↑(U ⊓ V) : subgroup G) = ↑U ⊓ ↑V := rfl
154+
@[simp, norm_cast, to_additive] lemma coe_opens_inf : (↑(U ⊓ V) : opens G) = ↑U ⊓ ↑V := rfl
155+
@[simp, to_additive] lemma mem_inf {x} : x ∈ U ⊓ V ↔ x ∈ U ∧ x ∈ V := iff.rfl
137156

138157
@[to_additive]
139158
instance : semilattice_inf (open_subgroup G) :=
140-
{ inf := λ U V, { is_open' := is_open.inter U.is_open V.is_open, .. (U : subgroup G) ⊓ V },
141-
inf_le_left := λ U V, set.inter_subset_left _ _,
142-
inf_le_right := λ U V, set.inter_subset_right _ _,
143-
le_inf := λ U V W hV hW, set.subset_inter hV hW,
144-
..open_subgroup.partial_order }
159+
{ .. set_like.partial_order,
160+
.. set_like.coe_injective.semilattice_inf (coe : open_subgroup G → set G) (λ _ _, rfl) }
145161

146162
@[to_additive]
147163
instance : order_top (open_subgroup G) :=
148164
{ top := ⊤,
149165
le_top := λ U, set.subset_univ _ }
150166

151-
@[simp, norm_cast, to_additive] lemma coe_inf : (↑(U ⊓ V) : set G) = (U : set G) ∩ V := rfl
152-
153-
@[simp, norm_cast, to_additive] lemma coe_subset : (U : set G) ⊆ V ↔ U ≤ V := iff.rfl
154-
155167
@[simp, norm_cast, to_additive] lemma coe_subgroup_le :
156-
(U : subgroup G) ≤ (V : subgroup G) ↔ U ≤ V := iff.rfl
168+
(U : subgroup G) ≤ (V : subgroup G) ↔ U ≤ V :=
169+
iff.rfl
157170

158171
variables {N : Type*} [group N] [topological_space N]
159172

160173
/-- The preimage of an `open_subgroup` along a continuous `monoid` homomorphism
161174
is an `open_subgroup`. -/
162175
@[to_additive "The preimage of an `open_add_subgroup` along a continuous `add_monoid` homomorphism
163176
is an `open_add_subgroup`."]
164-
def comap (f : G →* N)
165-
(hf : continuous f) (H : open_subgroup N) : open_subgroup G :=
177+
def comap (f : G →* N) (hf : continuous f) (H : open_subgroup N) : open_subgroup G :=
166178
{ is_open' := H.is_open.preimage hf,
167179
.. (H : subgroup N).comap f }
168180

169-
@[simp, to_additive]
181+
@[simp, norm_cast, to_additive]
170182
lemma coe_comap (H : open_subgroup N) (f : G →* N) (hf : continuous f) :
171183
(H.comap f hf : set G) = f ⁻¹' H := rfl
172184

185+
@[simp, norm_cast, to_additive]
186+
lemma coe_subgroup_comap (H : open_subgroup N) (f : G →* N) (hf : continuous f) :
187+
(H.comap f hf : subgroup G) = (H : subgroup N).comap f := rfl
188+
173189
@[simp, to_additive]
174190
lemma mem_comap {H : open_subgroup N} {f : G →* N} {hf : continuous f} {x : G} :
175191
x ∈ H.comap f hf ↔ f x ∈ H := iff.rfl
@@ -190,45 +206,30 @@ variables {G : Type*} [group G] [topological_space G] [has_continuous_mul G] (H
190206
lemma is_open_of_mem_nhds {g : G} (hg : (H : set G) ∈ 𝓝 g) :
191207
is_open (H : set G) :=
192208
begin
193-
simp only [is_open_iff_mem_nhds, set_like.mem_coe] at hg ⊢,
194-
intros x hx,
195-
have : filter.tendsto (λ y, y * (x⁻¹ * g)) (𝓝 x) (𝓝 $ x * (x⁻¹ * g)) :=
196-
(continuous_id.mul continuous_const).tendsto _,
197-
rw [mul_inv_cancel_left] at this,
198-
have := filter.mem_map'.1 (this hg),
199-
replace hg : g ∈ H := set_like.mem_coe.1 (mem_of_mem_nhds hg),
200-
simp only [set_like.mem_coe, H.mul_mem_cancel_right (H.mul_mem (H.inv_mem hx) hg)] at this,
201-
exact this
209+
refine is_open_iff_mem_nhds.2 (λ x hx, _),
210+
have hg' : g ∈ H := set_like.mem_coe.1 (mem_of_mem_nhds hg),
211+
have : filter.tendsto (λ y, y * (x⁻¹ * g)) (𝓝 x) (𝓝 g) :=
212+
(continuous_id.mul continuous_const).tendsto' _ _ (mul_inv_cancel_left _ _),
213+
simpa only [set_like.mem_coe, filter.mem_map',
214+
H.mul_mem_cancel_right (H.mul_mem (H.inv_mem hx) hg')] using this hg,
202215
end
203216

204217
@[to_additive]
205-
lemma is_open_of_open_subgroup {U : open_subgroup G} (h : U.1 ≤ H) :
218+
lemma is_open_mono {H₁ H₂ : subgroup G} (h : H₁ ≤ H₂) (h₁ : is_open (H₁ : set G)) :
219+
is_open (H₂ : set G) :=
220+
is_open_of_mem_nhds _ $ filter.mem_of_superset (h₁.mem_nhds $ one_mem H₁) h
221+
222+
@[to_additive]
223+
lemma is_open_of_open_subgroup {U : open_subgroup G} (h : ↑U ≤ H) :
206224
is_open (H : set G) :=
207-
H.is_open_of_mem_nhds (filter.mem_of_superset U.mem_nhds_one h)
225+
is_open_mono h U.is_open
208226

209227
/-- If a subgroup of a topological group has `1` in its interior, then it is open. -/
210228
@[to_additive "If a subgroup of an additive topological group has `0` in its interior, then it is
211229
open."]
212-
lemma is_open_of_one_mem_interior {G : Type*} [group G] [topological_space G]
213-
[topological_group G] {H : subgroup G} (h_1_int : (1 : G) ∈ interior (H : set G)) :
230+
lemma is_open_of_one_mem_interior (h_1_int : (1 : G) ∈ interior (H : set G)) :
214231
is_open (H : set G) :=
215-
begin
216-
have h : 𝓝 1 ≤ filter.principal (H : set G) :=
217-
nhds_le_of_le h_1_int (is_open_interior) (filter.principal_mono.2 interior_subset),
218-
rw is_open_iff_nhds,
219-
intros g hg,
220-
rw (show 𝓝 g = filter.map ⇑(homeomorph.mul_left g) (𝓝 1), by simp),
221-
convert filter.map_mono h,
222-
simp only [homeomorph.coe_mul_left, filter.map_principal, set.image_mul_left,
223-
filter.principal_eq_iff_eq],
224-
ext,
225-
simp [H.mul_mem_cancel_left (H.inv_mem hg)],
226-
end
227-
228-
@[to_additive]
229-
lemma is_open_mono {H₁ H₂ : subgroup G} (h : H₁ ≤ H₂) (h₁ : is_open (H₁ : set G)) :
230-
is_open (H₂ : set G) :=
231-
@is_open_of_open_subgroup _ _ _ _ H₂ { is_open' := h₁, .. H₁ } h
232+
is_open_of_mem_nhds H $ mem_interior_iff_mem_nhds.1 h_1_int
232233

233234
end subgroup
234235

@@ -237,20 +238,16 @@ namespace open_subgroup
237238
variables {G : Type*} [group G] [topological_space G] [has_continuous_mul G]
238239

239240
@[to_additive]
240-
instance : semilattice_sup (open_subgroup G) :=
241-
{ sup := λ U V,
242-
{ is_open' := show is_open (((U : subgroup G) ⊔ V : subgroup G) : set G),
243-
from subgroup.is_open_mono le_sup_left U.is_open,
244-
.. ((U : subgroup G) ⊔ V) },
245-
le_sup_left := λ U V, coe_subgroup_le.1 le_sup_left,
246-
le_sup_right := λ U V, coe_subgroup_le.1 le_sup_right,
247-
sup_le := λ U V W hU hV, coe_subgroup_le.1 (sup_le hU hV),
248-
..open_subgroup.semilattice_inf }
241+
instance : has_sup (open_subgroup G) :=
242+
⟨λ U V, ⟨U ⊔ V, subgroup.is_open_mono (le_sup_left : U.1 ≤ U ⊔ V) U.is_open⟩⟩
243+
244+
@[simp, norm_cast, to_additive]
245+
lemma coe_subgroup_sup (U V : open_subgroup G) : (↑(U ⊔ V) : subgroup G) = ↑U ⊔ ↑V := rfl
249246

250247
@[to_additive]
251248
instance : lattice (open_subgroup G) :=
252-
{ ..open_subgroup.semilattice_sup, ..open_subgroup.semilattice_inf }
253-
249+
{ .. open_subgroup.semilattice_inf,
250+
.. coe_subgroup_injective.semilattice_sup (coe : open_subgroup G → subgroup G) (λ _ _, rfl) }
254251

255252
end open_subgroup
256253

0 commit comments

Comments
 (0)