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

Commit 33c6eea

Browse files
committed
chore(topology/separation): rename separated to separated_nhds (#16604)
E.g., Wikipedia uses "separated" for `disjoint (closure s) t ∧ disjoint s (closure t)`.
1 parent dc1ac24 commit 33c6eea

File tree

1 file changed

+40
-43
lines changed

1 file changed

+40
-43
lines changed

src/topology/separation.lean

Lines changed: 40 additions & 43 deletions
Original file line numberDiff line numberDiff line change
@@ -11,12 +11,13 @@ import topology.inseparable
1111
/-!
1212
# Separation properties of topological spaces.
1313
14-
This file defines the predicate `separated`, and common separation axioms
14+
This file defines the predicate `separated_nhds`, and common separation axioms
1515
(under the Kolmogorov classification).
1616
1717
## Main definitions
1818
19-
* `separated`: Two `set`s are separated if they are contained in disjoint open sets.
19+
* `separated_nhds`: Two `set`s are separated by neighbourhoods if they are contained in disjoint
20+
open sets.
2021
* `t0_space`: A T₀/Kolmogorov space is a space where, for every two points `x ≠ y`,
2122
there is an open set that contains one, but not the other.
2223
* `t1_space`: A T₁/Fréchet space is a space where every singleton set is closed.
@@ -52,7 +53,7 @@ This file defines the predicate `separated`, and common separation axioms
5253
* `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter.
5354
* `t2_iff_is_closed_diagonal`: A space is T₂ iff the `diagonal` of `α` (that is, the set of all
5455
points of the form `(a, a) : α × α`) is closed under the product topology.
55-
* `finset_disjoint_finset_opens_of_t2`: Any two disjoint finsets are `separated`.
56+
* `finset_disjoint_finset_opens_of_t2`: Any two disjoint finsets are `separated_nhds`.
5657
* Most topological constructions preserve Hausdorffness;
5758
these results are part of the typeclass inference system (e.g. `embedding.t2_space`)
5859
* `set.eq_on.closure`: If two functions are equal on some set `s`, they are equal on its closure.
@@ -99,62 +100,60 @@ variables {α : Type u} {β : Type v} [topological_space α]
99100
section separation
100101

101102
/--
102-
`separated` is a predicate on pairs of sub`set`s of a topological space. It holds if the two
103+
`separated_nhds` is a predicate on pairs of sub`set`s of a topological space. It holds if the two
103104
sub`set`s are contained in disjoint open sets.
104105
-/
105-
def separated : set α → set α → Prop :=
106+
def separated_nhds : set α → set α → Prop :=
106107
λ (s t : set α), ∃ U V : (set α), (is_open U) ∧ is_open V ∧
107108
(s ⊆ U) ∧ (t ⊆ V) ∧ disjoint U V
108109

109-
lemma separated_iff_disjoint {s t : set α} :
110-
separated s t ↔ disjoint (𝓝ˢ s) (𝓝ˢ t) :=
111-
by simp only [(has_basis_nhds_set s).disjoint_iff (has_basis_nhds_set t), separated, exists_prop,
112-
← exists_and_distrib_left, and.assoc, and.comm, and.left_comm]
110+
lemma separated_nhds_iff_disjoint {s t : set α} :
111+
separated_nhds s t ↔ disjoint (𝓝ˢ s) (𝓝ˢ t) :=
112+
by simp only [(has_basis_nhds_set s).disjoint_iff (has_basis_nhds_set t), separated_nhds,
113+
exists_prop, ← exists_and_distrib_left, and.assoc, and.comm, and.left_comm]
113114

114-
namespace separated
115+
namespace separated_nhds
115116

116-
open separated
117+
variables {s s₁ s₂ t t₁ t₂ u : set α}
117118

118-
@[symm] lemma symm {s t : set α} : separated s t → separated t s :=
119+
@[symm] lemma symm : separated_nhds s t → separated_nhds t s :=
119120
λ ⟨U, V, oU, oV, aU, bV, UV⟩, ⟨V, U, oV, oU, bV, aU, disjoint.symm UV⟩
120121

121-
lemma comm (s t : set α) : separated s t ↔ separated t s :=
122-
⟨symm, symm⟩
122+
lemma comm (s t : set α) : separated_nhds s t ↔ separated_nhds t s := ⟨symm, symm⟩
123123

124-
lemma preimage [topological_space β] {f : α → β} {s t : set β} (h : separated s t)
125-
(hf : continuous f) : separated (f ⁻¹' s) (f ⁻¹' t) :=
124+
lemma preimage [topological_space β] {f : α → β} {s t : set β} (h : separated_nhds s t)
125+
(hf : continuous f) : separated_nhds (f ⁻¹' s) (f ⁻¹' t) :=
126126
let ⟨U, V, oU, oV, sU, tV, UV⟩ := h in
127127
⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV,
128128
UV.preimage f⟩
129129

130-
protected lemma disjoint {s t : set α} (h : separated s t) : disjoint s t :=
130+
protected lemma disjoint (h : separated_nhds s t) : disjoint s t :=
131131
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in hd.mono hsU htV
132132

133-
lemma disjoint_closure_left {s t : set α} (h : separated s t) : disjoint (closure s) t :=
133+
lemma disjoint_closure_left (h : separated_nhds s t) : disjoint (closure s) t :=
134134
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h
135135
in (hd.closure_left hV).mono (closure_mono hsU) htV
136136

137-
lemma disjoint_closure_right {s t : set α} (h : separated s t) : disjoint s (closure t) :=
137+
lemma disjoint_closure_right (h : separated_nhds s t) : disjoint s (closure t) :=
138138
h.symm.disjoint_closure_left.symm
139139

140-
lemma empty_right (a : set α) : separated a ∅ :=
140+
lemma empty_right (s : set α) : separated_nhds s ∅ :=
141141
⟨_, _, is_open_univ, is_open_empty, λ a h, mem_univ a, λ a h, by cases h, disjoint_empty _⟩
142142

143-
lemma empty_left (a : set α) : separateda :=
143+
lemma empty_left (s : set α) : separated_nhdss :=
144144
(empty_right _).symm
145145

146-
lemma mono {s₁ s₂ t₁ t₂ : set α} (h : separated s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) :
147-
separated s₁ t₁ :=
146+
lemma mono (h : separated_nhds s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : separated_nhds s₁ t₁ :=
148147
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩
149148

150-
lemma union_left {a b c : set α} : separated a c → separated b cseparated (ab) c :=
151-
by simpa only [separated_iff_disjoint, nhds_set_union, disjoint_sup_left] using and.intro
149+
lemma union_left : separated_nhds s u → separated_nhds t useparated_nhds (st) u :=
150+
by simpa only [separated_nhds_iff_disjoint, nhds_set_union, disjoint_sup_left] using and.intro
152151

153-
lemma union_right {a b c : set α} (ab : separated a b) (ac : separated a c) :
154-
separated a (bc) :=
155-
(ab.symm.union_left ac.symm).symm
152+
lemma union_right (ht : separated_nhds s t) (hu : separated_nhds s u) :
153+
separated_nhds s (tu) :=
154+
(ht.symm.union_left hu.symm).symm
156155

157-
end separated
156+
end separated_nhds
158157

159158
/-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair
160159
`x ≠ y`, there is an open set containing one but not the other. We formulate the definition in terms
@@ -873,10 +872,10 @@ t2_iff_is_closed_diagonal.mp ‹_›
873872

874873
section separated
875874

876-
open separated finset
875+
open separated_nhds finset
877876

878877
lemma finset_disjoint_finset_opens_of_t2 [t2_space α] :
879-
∀ (s t : finset α), disjoint s t → separated (s : set α) t :=
878+
∀ (s t : finset α), disjoint s t → separated_nhds (s : set α) t :=
880879
begin
881880
refine induction_on_union _ (λ a b hi d, (hi d.symm).symm) (λ a d, empty_right a) (λ a b ab, _) _,
882881
{ obtain ⟨U, V, oU, oV, aU, bV, UV⟩ := t2_separation (finset.disjoint_singleton.1 ab),
@@ -888,7 +887,7 @@ begin
888887
end
889888

890889
lemma point_disjoint_finset_opens_of_t2 [t2_space α] {x : α} {s : finset α} (h : x ∉ s) :
891-
separated ({x} : set α) s :=
890+
separated_nhds ({x} : set α) s :=
892891
by exact_mod_cast finset_disjoint_finset_opens_of_t2 {x} s (finset.disjoint_singleton_left.mpr h)
893892

894893
end separated
@@ -1135,8 +1134,8 @@ lemma function.left_inverse.closed_embedding [t2_space α] {f : α → β} {g :
11351134

11361135
lemma compact_compact_separated [t2_space α] {s t : set α}
11371136
(hs : is_compact s) (ht : is_compact t) (hst : disjoint s t) :
1138-
∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v :=
1139-
by simp only [prod_subset_compl_diagonal_iff_disjoint.symm] at ⊢ hst;
1137+
separated_nhds s t :=
1138+
by simp only [separated_nhds, prod_subset_compl_diagonal_iff_disjoint.symm] at ⊢ hst;
11401139
exact generalized_tube_lemma hs ht is_closed_diagonal.is_open_compl hst
11411140

11421141
/-- In a `t2_space`, every compact set is closed. -/
@@ -1526,12 +1525,11 @@ section normality
15261525
omits T₂), is one in which for every pair of disjoint closed sets `C` and `D`,
15271526
there exist disjoint open sets containing `C` and `D` respectively. -/
15281527
class normal_space (α : Type u) [topological_space α] extends t1_space α : Prop :=
1529-
(normal : ∀ s t : set α, is_closed s → is_closed t → disjoint s t →
1530-
∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v)
1528+
(normal : ∀ s t : set α, is_closed s → is_closed t → disjoint s t → separated_nhds s t)
15311529

15321530
theorem normal_separation [normal_space α] {s t : set α}
15331531
(H1 : is_closed s) (H2 : is_closed t) (H3 : disjoint s t) :
1534-
∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v :=
1532+
separated_nhds s t :=
15351533
normal_space.normal s t H1 H2 H3
15361534

15371535
theorem normal_exists_closure_subset [normal_space α] {s t : set α} (hs : is_closed s)
@@ -1562,11 +1560,10 @@ protected lemma closed_embedding.normal_space [topological_space β] [normal_spa
15621560
normal :=
15631561
begin
15641562
intros s t hs ht hst,
1565-
rcases normal_space.normal (f '' s) (f '' t) (hf.is_closed_map s hs) (hf.is_closed_map t ht)
1566-
(disjoint_image_of_injective hf.inj hst) with ⟨u, v, hu, hv, hsu, htv, huv⟩,
1567-
rw image_subset_iff at hsu htv,
1568-
exact ⟨f ⁻¹' u, f ⁻¹' v, hu.preimage hf.continuous, hv.preimage hf.continuous,
1569-
hsu, htv, huv.preimage f⟩
1563+
have H : separated_nhds (f '' s) (f '' t),
1564+
from normal_space.normal (f '' s) (f '' t) (hf.is_closed_map s hs) (hf.is_closed_map t ht)
1565+
(disjoint_image_of_injective hf.inj hst),
1566+
exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _)
15701567
end }
15711568

15721569
variable (α)
@@ -1652,7 +1649,7 @@ instance [t5_space α] {p : α → Prop} : t5_space {x // p x} := embedding_subt
16521649
/-- A `T₅` space is a `T₄` space. -/
16531650
@[priority 100] -- see Note [lower instance priority]
16541651
instance t5_space.to_normal_space [t5_space α] : normal_space α :=
1655-
⟨λ s t hs ht hd, separated_iff_disjoint.2 $
1652+
⟨λ s t hs ht hd, separated_nhds_iff_disjoint.2 $
16561653
completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq])⟩
16571654

16581655
end completely_normal

0 commit comments

Comments
 (0)