@@ -37,7 +37,6 @@ is equivalent to asking that the uniform structure induced on `s` is separated.
37
37
38
38
* `separation_relation X : set (X × X)`: the separation relation
39
39
* `separated_space X`: a predicate class asserting that `X` is separated
40
- * `is_separated s`: a predicate asserting that `s : set X` is separated
41
40
* `separation_quotient X`: the maximal separated quotient of `X`.
42
41
* `separation_quotient.lift f`: factors a map `f : X → Y` through the separation quotient of `X`.
43
42
* `separation_quotient.map f`: turns a map `f : X → Y` into a map between the separation quotients
@@ -136,6 +135,11 @@ lemma eq_of_forall_symmetric {α : Type*} [uniform_space α] [separated_space α
136
135
(h : ∀ {V}, V ∈ 𝓤 α → symmetric_rel V → (x, y) ∈ V) : x = y :=
137
136
eq_of_uniformity_basis has_basis_symmetric (by simpa [and_imp] using λ _, h)
138
137
138
+ lemma eq_of_cluster_pt_uniformity [separated_space α] {x y : α} (h : cluster_pt (x, y) (𝓤 α)) :
139
+ x = y :=
140
+ eq_of_uniformity_basis uniformity_has_basis_closed $ λ V ⟨hV, hVc⟩,
141
+ is_closed_iff_cluster_pt.1 hVc _ $ h.mono $ le_principal_iff.2 hV
142
+
139
143
lemma id_rel_sub_separation_relation (α : Type *) [uniform_space α] : id_rel ⊆ 𝓢 α :=
140
144
begin
141
145
unfold separation_rel,
@@ -235,86 +239,6 @@ lemma is_closed_range_of_spaced_out {ι} [separated_space α] {V₀ : set (α ×
235
239
is_closed_of_spaced_out V₀_in $
236
240
by { rintro _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ h, exact hf x y (ne_of_apply_ne f h) }
237
241
238
- /-!
239
- ### Separated sets
240
- -/
241
-
242
- /-- A set `s` in a uniform space `α` is separated if the separation relation `𝓢 α`
243
- induces the trivial relation on `s`. -/
244
- def is_separated (s : set α) : Prop := ∀ x y ∈ s, (x, y) ∈ 𝓢 α → x = y
245
-
246
- lemma is_separated_def (s : set α) : is_separated s ↔ ∀ x y ∈ s, (x, y) ∈ 𝓢 α → x = y :=
247
- iff.rfl
248
-
249
- lemma is_separated_def' (s : set α) : is_separated s ↔ (s ×ˢ s) ∩ 𝓢 α ⊆ id_rel :=
250
- begin
251
- rw is_separated_def,
252
- split,
253
- { rintros h ⟨x, y⟩ ⟨⟨x_in, y_in⟩, H⟩,
254
- simp [h x x_in y y_in H] },
255
- { intros h x x_in y y_in xy_in,
256
- rw ← mem_id_rel,
257
- exact h ⟨mk_mem_prod x_in y_in, xy_in⟩ }
258
- end
259
-
260
- lemma is_separated.mono {s t : set α} (hs : is_separated s) (hts : t ⊆ s) : is_separated t :=
261
- λ x hx y hy, hs x (hts hx) y (hts hy)
262
-
263
- lemma univ_separated_iff : is_separated (univ : set α) ↔ separated_space α :=
264
- begin
265
- simp only [is_separated, mem_univ, true_implies_iff, separated_space_iff],
266
- split,
267
- { intro h,
268
- exact subset.antisymm (λ ⟨x, y⟩ xy_in, h x y xy_in) (id_rel_sub_separation_relation α), },
269
- { intros h x y xy_in,
270
- rwa h at xy_in },
271
- end
272
-
273
- lemma is_separated_of_separated_space [separated_space α] (s : set α) : is_separated s :=
274
- begin
275
- rw [is_separated, separated_space.out],
276
- tauto,
277
- end
278
-
279
- lemma is_separated_iff_induced {s : set α} : is_separated s ↔ separated_space s :=
280
- begin
281
- rw separated_space_iff,
282
- change _ ↔ 𝓢 {x // x ∈ s} = _,
283
- rw [separation_rel_comap rfl, is_separated_def'],
284
- split; intro h,
285
- { ext ⟨⟨x, x_in⟩, ⟨y, y_in⟩⟩,
286
- suffices : (x, y) ∈ 𝓢 α ↔ x = y, by simpa only [mem_id_rel],
287
- refine ⟨λ H, h ⟨mk_mem_prod x_in y_in, H⟩, _⟩,
288
- rintro rfl,
289
- exact id_rel_sub_separation_relation α rfl },
290
- { rintros ⟨x, y⟩ ⟨⟨x_in, y_in⟩, hS⟩,
291
- have A : (⟨⟨x, x_in⟩, ⟨y, y_in⟩⟩ : ↥s × ↥s) ∈ prod.map (coe : s → α) (coe : s → α) ⁻¹' 𝓢 α,
292
- from hS,
293
- simpa using h.subset A }
294
- end
295
-
296
- lemma eq_of_uniformity_inf_nhds_of_is_separated {s : set α} (hs : is_separated s) :
297
- ∀ {x y : α}, x ∈ s → y ∈ s → cluster_pt (x, y) (𝓤 α) → x = y :=
298
- begin
299
- intros x y x_in y_in H,
300
- have : ∀ V ∈ 𝓤 α, (x, y) ∈ closure V,
301
- { intros V V_in,
302
- rw mem_closure_iff_cluster_pt,
303
- have : 𝓤 α ≤ 𝓟 V, by rwa le_principal_iff,
304
- exact H.mono this },
305
- apply hs x x_in y y_in,
306
- simpa [separation_rel_eq_inter_closure],
307
- end
308
-
309
- lemma eq_of_uniformity_inf_nhds [separated_space α] :
310
- ∀ {x y : α}, cluster_pt (x, y) (𝓤 α) → x = y :=
311
- begin
312
- have : is_separated (univ : set α),
313
- { rw univ_separated_iff,
314
- assumption },
315
- introv,
316
- simpa using eq_of_uniformity_inf_nhds_of_is_separated this ,
317
- end
318
242
319
243
/-!
320
244
### Separation quotient
@@ -440,11 +364,6 @@ lemma eq_of_separated_of_uniform_continuous [separated_space β] {f : α → β}
440
364
(H : uniform_continuous f) (h : x ≈ y) : f x = f y :=
441
365
separated_def.1 (by apply_instance) _ _ $ separated_of_uniform_continuous H h
442
366
443
- lemma _root_.is_separated.eq_of_uniform_continuous {f : α → β} {x y : α} {s : set β}
444
- (hs : is_separated s) (hxs : f x ∈ s) (hys : f y ∈ s) (H : uniform_continuous f) (h : x ≈ y) :
445
- f x = f y :=
446
- (is_separated_def _).mp hs _ hxs _ hys $ λ _ h', h _ (H h')
447
-
448
367
/-- The maximal separated quotient of a uniform space `α`. -/
449
368
def separation_quotient (α : Type *) [uniform_space α] := quotient (separation_setoid α)
450
369
@@ -519,10 +438,4 @@ separated_def.2 $ assume x y H, prod.ext
519
438
(eq_of_separated_of_uniform_continuous uniform_continuous_fst H)
520
439
(eq_of_separated_of_uniform_continuous uniform_continuous_snd H)
521
440
522
- lemma _root_.is_separated.prod {s : set α} {t : set β} (hs : is_separated s) (ht : is_separated t) :
523
- is_separated (s ×ˢ t) :=
524
- (is_separated_def _).mpr $ λ x hx y hy H, prod.ext
525
- (hs.eq_of_uniform_continuous hx.1 hy.1 uniform_continuous_fst H)
526
- (ht.eq_of_uniform_continuous hx.2 hy.2 uniform_continuous_snd H)
527
-
528
441
end uniform_space
0 commit comments