@@ -154,29 +154,22 @@ lemma union_right {a b c : set α} (ab : separated a b) (ac : separated a c) :
154
154
155
155
end separated
156
156
157
- /-- A T₀ space, also known as a Kolmogorov space, is a topological space
158
- where for every pair `x ≠ y`, there is an open set containing one but not the other. -/
157
+ /-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair
158
+ `x ≠ y`, there is an open set containing one but not the other. We formulate the definition in terms
159
+ of the `inseparable` relation. -/
159
160
class t0_space (α : Type u) [topological_space α] : Prop :=
160
- (t0 : ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)) )
161
+ (t0 : ∀ ⦃ x y : α⦄, inseparable x y → x = y )
161
162
162
- lemma exists_is_open_xor_mem [t0_space α] {x y : α} (h : x ≠ y) :
163
- ∃ U:set α, is_open U ∧ xor (x ∈ U) (y ∈ U) :=
164
- t0_space.t0 x y h
165
-
166
- lemma t0_space_def (α : Type u) [topological_space α] :
167
- t0_space α ↔ ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)) :=
168
- by { split, apply @t0_space.t0, apply t0_space.mk }
163
+ lemma t0_space_iff_inseparable (α : Type u) [topological_space α] :
164
+ t0_space α ↔ ∀ (x y : α), inseparable x y → x = y :=
165
+ ⟨λ ⟨h⟩, h, λ h, ⟨h⟩⟩
169
166
170
167
lemma t0_space_iff_not_inseparable (α : Type u) [topological_space α] :
171
168
t0_space α ↔ ∀ (x y : α), x ≠ y → ¬inseparable x y :=
172
- by simp only [t0_space_def, xor_iff_not_iff, not_forall, exists_prop, inseparable_iff_forall_open]
173
-
174
- lemma t0_space_iff_inseparable (α : Type u) [topological_space α] :
175
- t0_space α ↔ ∀ (x y : α), inseparable x y → x = y :=
176
- by simp only [t0_space_iff_not_inseparable, ne.def, not_imp_not]
169
+ by simp only [t0_space_iff_inseparable, ne.def, not_imp_not]
177
170
178
171
lemma inseparable.eq [t0_space α] {x y : α} (h : inseparable x y) : x = y :=
179
- (t0_space_iff_inseparable α). 1 ‹_› x y h
172
+ t0_space.t0 h
180
173
181
174
lemma t0_space_iff_nhds_injective (α : Type u) [topological_space α] :
182
175
t0_space α ↔ injective (𝓝 : α → filter α) :=
@@ -188,14 +181,23 @@ lemma nhds_injective [t0_space α] : injective (𝓝 : α → filter α) :=
188
181
@[simp] lemma nhds_eq_nhds_iff [t0_space α] {a b : α} : 𝓝 a = 𝓝 b ↔ a = b :=
189
182
nhds_injective.eq_iff
190
183
184
+ lemma t0_space_iff_exists_is_open_xor_mem (α : Type u) [topological_space α] :
185
+ t0_space α ↔ ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)) :=
186
+ by simp only [t0_space_iff_not_inseparable, xor_iff_not_iff, not_forall, exists_prop,
187
+ inseparable_iff_forall_open]
188
+
189
+ lemma exists_is_open_xor_mem [t0_space α] {x y : α} (h : x ≠ y) :
190
+ ∃ U : set α, is_open U ∧ xor (x ∈ U) (y ∈ U) :=
191
+ (t0_space_iff_exists_is_open_xor_mem α).1 ‹_› x y h
192
+
191
193
/-- Specialization forms a partial order on a t0 topological space. -/
192
194
def specialization_order (α : Type *) [topological_space α] [t0_space α] : partial_order α :=
193
195
{ .. specialization_preorder α,
194
196
.. partial_order.lift (order_dual.to_dual ∘ 𝓝) nhds_injective }
195
197
196
198
instance : t0_space (separation_quotient α) :=
197
- (t0_space_iff_inseparable _). 2 $ λ x' y', quotient.induction_on₂' x' y' $
198
- λ x y h, separation_quotient.mk_eq_mk.2 $ separation_quotient.inducing_mk.inseparable_iff.1 h
199
+ ⟨ λ x' y', quotient.induction_on₂' x' y' $
200
+ λ x y h, separation_quotient.mk_eq_mk.2 $ separation_quotient.inducing_mk.inseparable_iff.1 h⟩
199
201
200
202
theorem minimal_nonempty_closed_subsingleton [t0_space α] {s : set α} (hs : is_closed s)
201
203
(hmin : ∀ t ⊆ s, t.nonempty → is_closed t → t = s) :
@@ -264,8 +266,7 @@ let ⟨x, _, h⟩ := exists_open_singleton_of_open_finite (finite.of_fintype _)
264
266
265
267
lemma t0_space_of_injective_of_continuous [topological_space β] {f : α → β}
266
268
(hf : function.injective f) (hf' : continuous f) [t0_space β] : t0_space α :=
267
- ⟨λ x y hxy, let ⟨U, hU, hxyU⟩ := t0_space.t0 (f x) (f y) (hf.ne hxy) in
268
- ⟨f ⁻¹' U, hU.preimage hf', hxyU⟩⟩
269
+ ⟨λ x y h, hf $ (h.map hf').eq⟩
269
270
270
271
protected lemma embedding.t0_space [topological_space β] [t0_space β] {f : α → β}
271
272
(hf : embedding f) : t0_space α :=
@@ -279,12 +280,11 @@ theorem t0_space_iff_or_not_mem_closure (α : Type u) [topological_space α] :
279
280
by simp only [t0_space_iff_not_inseparable, inseparable_iff_mem_closure, not_and_distrib]
280
281
281
282
instance [topological_space β] [t0_space α] [t0_space β] : t0_space (α × β) :=
282
- (t0_space_iff_inseparable _).2 $
283
- λ x y h, prod.ext (h.map continuous_fst).eq (h.map continuous_snd).eq
283
+ ⟨λ x y h, prod.ext (h.map continuous_fst).eq (h.map continuous_snd).eq⟩
284
284
285
285
instance {ι : Type *} {π : ι → Type *} [Π i, topological_space (π i)] [Π i, t0_space (π i)] :
286
286
t0_space (Π i, π i) :=
287
- (t0_space_iff_inseparable _). 2 $ λ x y h, funext $ λ i, (h.map (continuous_apply i)).eq
287
+ ⟨ λ x y h, funext $ λ i, (h.map (continuous_apply i)).eq⟩
288
288
289
289
/-- A T₁ space, also known as a Fréchet space, is a topological space
290
290
where every singleton set is closed. Equivalently, for every pair
@@ -367,7 +367,8 @@ lemma t1_space_tfae (α : Type u) [topological_space α] :
367
367
∀ ⦃x y : α⦄, x ≠ y → ∃ s ∈ 𝓝 x, y ∉ s,
368
368
∀ ⦃x y : α⦄, x ≠ y → ∃ (U : set α) (hU : is_open U), x ∈ U ∧ y ∉ U,
369
369
∀ ⦃x y : α⦄, x ≠ y → disjoint (𝓝 x) (pure y),
370
- ∀ ⦃x y : α⦄, x ≠ y → disjoint (pure x) (𝓝 y)] :=
370
+ ∀ ⦃x y : α⦄, x ≠ y → disjoint (pure x) (𝓝 y),
371
+ ∀ ⦃x y : α⦄, x ⤳ y → x = y] :=
371
372
begin
372
373
tfae_have : 1 ↔ 2 , from ⟨λ h, h.1 , λ h, ⟨h⟩⟩,
373
374
tfae_have : 2 ↔ 3 , by simp only [is_open_compl_iff],
@@ -388,6 +389,9 @@ begin
388
389
exacts [is_open_empty, compl_compl s ▸ (@set.finite.is_closed _ _ H _ hs).is_open_compl] },
389
390
tfae_have : 4 → 2 ,
390
391
from λ h x, (cofinite_topology.is_closed_iff.2 $ or.inr (finite_singleton _)).preimage h,
392
+ tfae_have : 2 ↔ 10 ,
393
+ { simp only [← closure_subset_iff_is_closed, specializes_iff_mem_closure, subset_def,
394
+ mem_singleton_iff, eq_comm] },
391
395
tfae_finish
392
396
end
393
397
@@ -408,12 +412,21 @@ lemma t1_space_iff_disjoint_pure_nhds : t1_space α ↔ ∀ ⦃x y : α⦄, x
408
412
lemma t1_space_iff_disjoint_nhds_pure : t1_space α ↔ ∀ ⦃x y : α⦄, x ≠ y → disjoint (𝓝 x) (pure y) :=
409
413
(t1_space_tfae α).out 0 7
410
414
415
+ lemma t1_space_iff_specializes_imp_eq : t1_space α ↔ ∀ ⦃x y : α⦄, x ⤳ y → x = y :=
416
+ (t1_space_tfae α).out 0 9
417
+
411
418
lemma disjoint_pure_nhds [t1_space α] {x y : α} (h : x ≠ y) : disjoint (pure x) (𝓝 y) :=
412
419
t1_space_iff_disjoint_pure_nhds.mp ‹_› h
413
420
414
421
lemma disjoint_nhds_pure [t1_space α] {x y : α} (h : x ≠ y) : disjoint (𝓝 x) (pure y) :=
415
422
t1_space_iff_disjoint_nhds_pure.mp ‹_› h
416
423
424
+ lemma specializes.eq [t1_space α] {x y : α} (h : x ⤳ y) : x = y :=
425
+ t1_space_iff_specializes_imp_eq.1 ‹_› h
426
+
427
+ @[simp] lemma specializes_iff_eq [t1_space α] {x y : α} : x ⤳ y ↔ x = y :=
428
+ ⟨specializes.eq, λ h, h ▸ specializes_rfl⟩
429
+
417
430
instance {α : Type *} : t1_space (cofinite_topology α) :=
418
431
t1_space_iff_continuous_cofinite_of.mpr continuous_id
419
432
454
467
455
468
lemma t1_space_of_injective_of_continuous [topological_space β] {f : α → β}
456
469
(hf : function.injective f) (hf' : continuous f) [t1_space β] : t1_space α :=
457
- { t1 :=
458
- begin
459
- intros x,
460
- rw [← function.injective.preimage_image hf {x}, image_singleton],
461
- exact (t1_space.t1 $ f x).preimage hf'
462
- end }
470
+ t1_space_iff_specializes_imp_eq.2 $ λ x y h, hf (h.map hf').eq
463
471
464
472
protected lemma embedding.t1_space [topological_space β] [t1_space β] {f : α → β}
465
473
(hf : embedding f) : t1_space α :=
@@ -477,8 +485,7 @@ instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i
477
485
⟨λ f, univ_pi_singleton f ▸ is_closed_set_pi (λ i hi, is_closed_singleton)⟩
478
486
479
487
@[priority 100 ] -- see Note [lower instance priority]
480
- instance t1_space.t0_space [t1_space α] : t0_space α :=
481
- ⟨λ x y h, ⟨{z | z ≠ y}, is_open_ne, or.inl ⟨h, not_not_intro rfl⟩⟩⟩
488
+ instance t1_space.t0_space [t1_space α] : t0_space α := ⟨λ x y h, h.specializes.eq⟩
482
489
483
490
@[simp] lemma compl_singleton_mem_nhds_iff [t1_space α] {x y : α} : {x}ᶜ ∈ 𝓝 y ↔ y ≠ x :=
484
491
is_open_compl_singleton.mem_nhds_iff
@@ -498,12 +505,6 @@ hs.induction_on (by simp) $ λ x, by simp
498
505
(closure s).subsingleton ↔ s.subsingleton :=
499
506
⟨λ h, h.mono subset_closure, λ h, h.closure⟩
500
507
501
- lemma specializes.eq [t1_space α] {x y : α} (h : x ⤳ y) : x = y :=
502
- by simpa only [specializes_iff_mem_closure, closure_singleton, mem_singleton_iff, eq_comm] using h
503
-
504
- @[simp] lemma specializes_iff_eq [t1_space α] {x y : α} : x ⤳ y ↔ x = y :=
505
- ⟨specializes.eq, λ h, h ▸ specializes_refl _⟩
506
-
507
508
lemma is_closed_map_const {α β} [topological_space α] [topological_space β] [t1_space β] {y : β} :
508
509
is_closed_map (function.const α y) :=
509
510
is_closed_map.of_nonempty $ λ s hs h2s, by simp_rw [h2s.image_const, is_closed_singleton]
@@ -1263,7 +1264,7 @@ instance regular_space.t1_space [regular_space α] : t1_space α :=
1263
1264
begin
1264
1265
rw t1_space_iff_exists_open,
1265
1266
intros x y hxy,
1266
- obtain ⟨U, hU, h⟩ := t0_space.t0 x y hxy,
1267
+ obtain ⟨U, hU, h⟩ := exists_is_open_xor_mem hxy,
1267
1268
cases h,
1268
1269
{ exact ⟨U, hU, h⟩ },
1269
1270
{ obtain ⟨R, hR, hh⟩ := regular_space.regular (is_closed_compl_iff.mpr hU) (not_not.mpr h.1 ),
0 commit comments