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

Commit d45a8ac

Browse files
committed
refactor(topology/separation): redefine t0_space (#15046)
1 parent 71b1be6 commit d45a8ac

File tree

4 files changed

+70
-51
lines changed

4 files changed

+70
-51
lines changed

src/algebraic_geometry/prime_spectrum/basic.lean

Lines changed: 6 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -738,9 +738,12 @@ lemma le_iff_specializes (x y : prime_spectrum R) :
738738
x ≤ y ↔ x ⤳ y :=
739739
(le_iff_mem_closure x y).trans specializes_iff_mem_closure.symm
740740

741-
instance : t0_space (prime_spectrum R) :=
742-
by { simp [t0_space_iff_or_not_mem_closure, ← le_iff_mem_closure,
743-
← not_and_distrib, ← le_antisymm_iff, eq_comm] }
741+
/-- `nhds` as an order embedding. -/
742+
@[simps { fully_applied := tt }]
743+
def nhds_order_embedding : prime_spectrum R ↪o filter (prime_spectrum R) :=
744+
order_embedding.of_map_le_iff nhds $ λ a b, (le_iff_specializes a b).symm
745+
746+
instance : t0_space (prime_spectrum R) := ⟨nhds_order_embedding.injective⟩
744747

745748
end order
746749

src/topology/alexandroff.lean

Lines changed: 23 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -325,6 +325,26 @@ lemma dense_embedding_coe [noncompact_space X] :
325325
dense_embedding (coe : X → alexandroff X) :=
326326
{ dense := dense_range_coe, .. open_embedding_coe }
327327

328+
@[simp] lemma specializes_coe {x y : X} : (x : alexandroff X) ⤳ y ↔ x ⤳ y :=
329+
open_embedding_coe.to_inducing.specializes_iff
330+
331+
@[simp] lemma inseparable_coe {x y : X} : inseparable (x : alexandroff X) y ↔ inseparable x y :=
332+
open_embedding_coe.to_inducing.inseparable_iff
333+
334+
lemma not_specializes_infty_coe {x : X} : ¬specializes ∞ (x : alexandroff X) :=
335+
is_closed_infty.not_specializes rfl (coe_ne_infty x)
336+
337+
lemma not_inseparable_infty_coe {x : X} : ¬inseparable ∞ (x : alexandroff X) :=
338+
λ h, not_specializes_infty_coe h.specializes
339+
340+
lemma not_inseparable_coe_infty {x : X} : ¬inseparable (x : alexandroff X) ∞ :=
341+
λ h, not_specializes_infty_coe h.specializes'
342+
343+
lemma inseparable_iff {x y : alexandroff X} :
344+
inseparable x y ↔ x = ∞ ∧ y = ∞ ∨ ∃ x' : X, x = x' ∧ ∃ y' : X, y = y' ∧ inseparable x' y' :=
345+
by induction x using alexandroff.rec; induction y using alexandroff.rec;
346+
simp [not_inseparable_infty_coe, not_inseparable_coe_infty, coe_eq_coe]
347+
328348
/-!
329349
### Compactness and separation properties
330350
@@ -351,13 +371,8 @@ instance : compact_space (alexandroff X) :=
351371
instance [t0_space X] : t0_space (alexandroff X) :=
352372
begin
353373
refine ⟨λ x y hxy, _⟩,
354-
induction x using alexandroff.rec; induction y using alexandroff.rec,
355-
{ exact (hxy rfl).elim },
356-
{ use {∞}ᶜ, simp [is_closed_infty] },
357-
{ use {∞}ᶜ, simp [is_closed_infty] },
358-
{ rcases t0_space.t0 x y (mt coe_eq_coe.mpr hxy) with ⟨U, hUo, hU⟩,
359-
refine ⟨coe '' U, is_open_image_coe.2 hUo, _⟩,
360-
simpa [coe_eq_coe] }
374+
rcases inseparable_iff.1 hxy with ⟨rfl, rfl⟩|⟨x, rfl, y, rfl, h⟩,
375+
exacts [rfl, congr_arg coe h.eq]
361376
end
362377

363378
/-- The one point compactification of a `t1_space` space is a `t1_space`. -/
@@ -366,7 +381,7 @@ instance [t1_space X] : t1_space (alexandroff X) :=
366381
begin
367382
induction z using alexandroff.rec,
368383
{ exact is_closed_infty },
369-
{ simp only [← image_singleton, is_closed_image_coe],
384+
{ rw [← image_singleton, is_closed_image_coe],
370385
exact ⟨is_closed_singleton, is_compact_singleton⟩ }
371386
end }
372387

src/topology/separation.lean

Lines changed: 40 additions & 39 deletions
Original file line numberDiff line numberDiff line change
@@ -154,29 +154,22 @@ lemma union_right {a b c : set α} (ab : separated a b) (ac : separated a c) :
154154

155155
end separated
156156

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. -/
159160
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)
161162

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⟩⟩
169166

170167
lemma t0_space_iff_not_inseparable (α : Type u) [topological_space α] :
171168
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]
177170

178171
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
180173

181174
lemma t0_space_iff_nhds_injective (α : Type u) [topological_space α] :
182175
t0_space α ↔ injective (𝓝 : α → filter α) :=
@@ -188,14 +181,23 @@ lemma nhds_injective [t0_space α] : injective (𝓝 : α → filter α) :=
188181
@[simp] lemma nhds_eq_nhds_iff [t0_space α] {a b : α} : 𝓝 a = 𝓝 b ↔ a = b :=
189182
nhds_injective.eq_iff
190183

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+
191193
/-- Specialization forms a partial order on a t0 topological space. -/
192194
def specialization_order (α : Type*) [topological_space α] [t0_space α] : partial_order α :=
193195
{ .. specialization_preorder α,
194196
.. partial_order.lift (order_dual.to_dual ∘ 𝓝) nhds_injective }
195197

196198
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
199201

200202
theorem minimal_nonempty_closed_subsingleton [t0_space α] {s : set α} (hs : is_closed s)
201203
(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 _)
264266

265267
lemma t0_space_of_injective_of_continuous [topological_space β] {f : α → β}
266268
(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⟩
269270

270271
protected lemma embedding.t0_space [topological_space β] [t0_space β] {f : α → β}
271272
(hf : embedding f) : t0_space α :=
@@ -279,12 +280,11 @@ theorem t0_space_iff_or_not_mem_closure (α : Type u) [topological_space α] :
279280
by simp only [t0_space_iff_not_inseparable, inseparable_iff_mem_closure, not_and_distrib]
280281

281282
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⟩
284284

285285
instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i, t0_space (π i)] :
286286
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
288288

289289
/-- A T₁ space, also known as a Fréchet space, is a topological space
290290
where every singleton set is closed. Equivalently, for every pair
@@ -367,7 +367,8 @@ lemma t1_space_tfae (α : Type u) [topological_space α] :
367367
∀ ⦃x y : α⦄, x ≠ y → ∃ s ∈ 𝓝 x, y ∉ s,
368368
∀ ⦃x y : α⦄, x ≠ y → ∃ (U : set α) (hU : is_open U), x ∈ U ∧ y ∉ U,
369369
∀ ⦃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] :=
371372
begin
372373
tfae_have : 12, from ⟨λ h, h.1, λ h, ⟨h⟩⟩,
373374
tfae_have : 23, by simp only [is_open_compl_iff],
@@ -388,6 +389,9 @@ begin
388389
exacts [is_open_empty, compl_compl s ▸ (@set.finite.is_closed _ _ H _ hs).is_open_compl] },
389390
tfae_have : 42,
390391
from λ h x, (cofinite_topology.is_closed_iff.2 $ or.inr (finite_singleton _)).preimage h,
392+
tfae_have : 210,
393+
{ simp only [← closure_subset_iff_is_closed, specializes_iff_mem_closure, subset_def,
394+
mem_singleton_iff, eq_comm] },
391395
tfae_finish
392396
end
393397

@@ -408,12 +412,21 @@ lemma t1_space_iff_disjoint_pure_nhds : t1_space α ↔ ∀ ⦃x y : α⦄, x
408412
lemma t1_space_iff_disjoint_nhds_pure : t1_space α ↔ ∀ ⦃x y : α⦄, x ≠ y → disjoint (𝓝 x) (pure y) :=
409413
(t1_space_tfae α).out 0 7
410414

415+
lemma t1_space_iff_specializes_imp_eq : t1_space α ↔ ∀ ⦃x y : α⦄, x ⤳ y → x = y :=
416+
(t1_space_tfae α).out 0 9
417+
411418
lemma disjoint_pure_nhds [t1_space α] {x y : α} (h : x ≠ y) : disjoint (pure x) (𝓝 y) :=
412419
t1_space_iff_disjoint_pure_nhds.mp ‹_› h
413420

414421
lemma disjoint_nhds_pure [t1_space α] {x y : α} (h : x ≠ y) : disjoint (𝓝 x) (pure y) :=
415422
t1_space_iff_disjoint_nhds_pure.mp ‹_› h
416423

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+
417430
instance {α : Type*} : t1_space (cofinite_topology α) :=
418431
t1_space_iff_continuous_cofinite_of.mpr continuous_id
419432

@@ -454,12 +467,7 @@ end
454467

455468
lemma t1_space_of_injective_of_continuous [topological_space β] {f : α → β}
456469
(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
463471

464472
protected lemma embedding.t1_space [topological_space β] [t1_space β] {f : α → β}
465473
(hf : embedding f) : t1_space α :=
@@ -477,8 +485,7 @@ instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i
477485
⟨λ f, univ_pi_singleton f ▸ is_closed_set_pi (λ i hi, is_closed_singleton)⟩
478486

479487
@[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⟩
482489

483490
@[simp] lemma compl_singleton_mem_nhds_iff [t1_space α] {x y : α} : {x}ᶜ ∈ 𝓝 y ↔ y ≠ x :=
484491
is_open_compl_singleton.mem_nhds_iff
@@ -498,12 +505,6 @@ hs.induction_on (by simp) $ λ x, by simp
498505
(closure s).subsingleton ↔ s.subsingleton :=
499506
⟨λ h, h.mono subset_closure, λ h, h.closure⟩
500507

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-
507508
lemma is_closed_map_const {α β} [topological_space α] [topological_space β] [t1_space β] {y : β} :
508509
is_closed_map (function.const α y) :=
509510
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 α :=
12631264
begin
12641265
rw t1_space_iff_exists_open,
12651266
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,
12671268
cases h,
12681269
{ exact ⟨U, hU, h⟩ },
12691270
{ obtain ⟨R, hR, hh⟩ := regular_space.regular (is_closed_compl_iff.mpr hU) (not_not.mpr h.1),

src/topology/uniform_space/separation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -186,7 +186,7 @@ end
186186

187187
@[priority 100] -- see Note [lower instance priority]
188188
instance separated_regular [separated_space α] : regular_space α :=
189-
{ t0 := by { haveI := separated_iff_t2.mp ‹_›, exact t1_space.t0_space.t0 },
189+
{ to_t0_space := by { haveI := separated_iff_t2.mp ‹_›, exact t1_space.t0_space },
190190
regular := λs a hs ha,
191191
have sᶜ ∈ 𝓝 a,
192192
from is_open.mem_nhds hs.is_open_compl ha,

0 commit comments

Comments
 (0)