@@ -55,44 +55,39 @@ variables {G : Type*} [group G] [topological_space G]
55
55
variables {U V : open_subgroup G} {g : G}
56
56
57
57
@[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 ⟩
59
59
60
60
@[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
62
63
63
64
@[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' }
65
74
66
75
@[to_additive]
67
76
instance has_coe_opens : has_coe_t (open_subgroup G) (opens G) := ⟨λ U, ⟨U, U.is_open'⟩⟩
68
77
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
70
80
@[simp, norm_cast, to_additive] lemma mem_coe_opens : g ∈ (U : opens G) ↔ g ∈ U := iff.rfl
71
81
@[simp, norm_cast, to_additive]
72
82
lemma mem_coe_subgroup : g ∈ (U : subgroup G) ↔ g ∈ U := iff.rfl
73
83
74
- @[to_additive] lemma coe_injective : injective (coe : open_subgroup G → set G) :=
75
- by { rintros ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨h⟩, congr, }
76
-
77
84
@[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
82
86
83
87
variable (U)
84
88
@[to_additive]
85
89
protected lemma is_open : is_open (U : set G) := U.is_open'
86
90
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
-
96
91
@[to_additive]
97
92
lemma mem_nhds_one : (U : set G) ∈ 𝓝 (1 : G) :=
98
93
is_open.mem_nhds U.is_open U.one_mem
@@ -101,6 +96,15 @@ variable {U}
101
96
@[to_additive]
102
97
instance : has_top (open_subgroup G) := ⟨{ is_open' := is_open_univ, .. (⊤ : subgroup G) }⟩
103
98
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
+
104
108
@[to_additive]
105
109
instance : inhabited (open_subgroup G) := ⟨⊤⟩
106
110
@@ -109,67 +113,79 @@ lemma is_closed [has_continuous_mul G] (U : open_subgroup G) : is_closed (U : se
109
113
begin
110
114
apply is_open_compl_iff.1 ,
111
115
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 ⊢,
115
118
convert U.mul_mem (U.inv_mem hux) hu,
116
119
simp },
117
120
{ exact U.is_open.preimage (continuous_mul_right _) },
118
- { simp [U. one_mem] }
121
+ { simp [one_mem] }
119
122
end
120
123
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
+
121
128
section
122
129
variables {H : Type *} [group H] [topological_space H]
123
130
124
131
/-- The product of two open subgroups as an open subgroup of the product group. -/
125
132
@[to_additive " The product of two open subgroups as an open subgroup of the product group." ]
126
133
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,
129
135
.. (U : subgroup G).prod (V : subgroup H) }
130
136
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
+
131
146
end
132
147
133
148
@[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
137
156
138
157
@[to_additive]
139
158
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) }
145
161
146
162
@[to_additive]
147
163
instance : order_top (open_subgroup G) :=
148
164
{ top := ⊤,
149
165
le_top := λ U, set.subset_univ _ }
150
166
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
-
155
167
@[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
157
170
158
171
variables {N : Type *} [group N] [topological_space N]
159
172
160
173
/-- The preimage of an `open_subgroup` along a continuous `monoid` homomorphism
161
174
is an `open_subgroup`. -/
162
175
@[to_additive " The preimage of an `open_add_subgroup` along a continuous `add_monoid` homomorphism
163
176
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 :=
166
178
{ is_open' := H.is_open.preimage hf,
167
179
.. (H : subgroup N).comap f }
168
180
169
- @[simp, to_additive]
181
+ @[simp, norm_cast, to_additive]
170
182
lemma coe_comap (H : open_subgroup N) (f : G →* N) (hf : continuous f) :
171
183
(H.comap f hf : set G) = f ⁻¹' H := rfl
172
184
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
+
173
189
@[simp, to_additive]
174
190
lemma mem_comap {H : open_subgroup N} {f : G →* N} {hf : continuous f} {x : G} :
175
191
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
190
206
lemma is_open_of_mem_nhds {g : G} (hg : (H : set G) ∈ 𝓝 g) :
191
207
is_open (H : set G) :=
192
208
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,
202
215
end
203
216
204
217
@[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) :
206
224
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
208
226
209
227
/-- If a subgroup of a topological group has `1` in its interior, then it is open. -/
210
228
@[to_additive " If a subgroup of an additive topological group has `0` in its interior, then it is
211
229
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)) :
214
231
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
232
233
233
234
end subgroup
234
235
@@ -237,20 +238,16 @@ namespace open_subgroup
237
238
variables {G : Type *} [group G] [topological_space G] [has_continuous_mul G]
238
239
239
240
@[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
249
246
250
247
@[to_additive]
251
248
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) }
254
251
255
252
end open_subgroup
256
253
0 commit comments