@@ -11,12 +11,13 @@ import topology.inseparable
11
11
/-!
12
12
# Separation properties of topological spaces.
13
13
14
- This file defines the predicate `separated `, and common separation axioms
14
+ This file defines the predicate `separated_nhds `, and common separation axioms
15
15
(under the Kolmogorov classification).
16
16
17
17
## Main definitions
18
18
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.
20
21
* `t0_space`: A T₀/Kolmogorov space is a space where, for every two points `x ≠ y`,
21
22
there is an open set that contains one, but not the other.
22
23
* `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
52
53
* `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter.
53
54
* `t2_iff_is_closed_diagonal`: A space is T₂ iff the `diagonal` of `α` (that is, the set of all
54
55
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 `.
56
57
* Most topological constructions preserve Hausdorffness;
57
58
these results are part of the typeclass inference system (e.g. `embedding.t2_space`)
58
59
* `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 α]
99
100
section separation
100
101
101
102
/--
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
103
104
sub`set`s are contained in disjoint open sets.
104
105
-/
105
- def separated : set α → set α → Prop :=
106
+ def separated_nhds : set α → set α → Prop :=
106
107
λ (s t : set α), ∃ U V : (set α), (is_open U) ∧ is_open V ∧
107
108
(s ⊆ U) ∧ (t ⊆ V) ∧ disjoint U V
108
109
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]
113
114
114
- namespace separated
115
+ namespace separated_nhds
115
116
116
- open separated
117
+ variables {s s₁ s₂ t t₁ t₂ u : set α}
117
118
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 :=
119
120
λ ⟨U, V, oU, oV, aU, bV, UV⟩, ⟨V, U, oV, oU, bV, aU, disjoint.symm UV⟩
120
121
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⟩
123
123
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) :=
126
126
let ⟨U, V, oU, oV, sU, tV, UV⟩ := h in
127
127
⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV,
128
128
UV.preimage f⟩
129
129
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 :=
131
131
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in hd.mono hsU htV
132
132
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 :=
134
134
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h
135
135
in (hd.closure_left hV).mono (closure_mono hsU) htV
136
136
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) :=
138
138
h.symm.disjoint_closure_left.symm
139
139
140
- lemma empty_right (a : set α) : separated a ∅ :=
140
+ lemma empty_right (s : set α) : separated_nhds s ∅ :=
141
141
⟨_, _, is_open_univ, is_open_empty, λ a h, mem_univ a, λ a h, by cases h, disjoint_empty _⟩
142
142
143
- lemma empty_left (a : set α) : separated ∅ a :=
143
+ lemma empty_left (s : set α) : separated_nhds ∅ s :=
144
144
(empty_right _).symm
145
145
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₁ :=
148
147
let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩
149
148
150
- lemma union_left {a b c : set α} : separated a c → separated b c → separated (a ∪ b) 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 u → separated_nhds (s ∪ t) u :=
150
+ by simpa only [separated_nhds_iff_disjoint , nhds_set_union, disjoint_sup_left] using and.intro
152
151
153
- lemma union_right {a b c : set α} (ab : separated a b ) (ac : separated a c ) :
154
- separated a (b ∪ c ) :=
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 (t ∪ u ) :=
154
+ (ht .symm.union_left hu .symm).symm
156
155
157
- end separated
156
+ end separated_nhds
158
157
159
158
/-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair
160
159
`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 ‹_›
873
872
874
873
section separated
875
874
876
- open separated finset
875
+ open separated_nhds finset
877
876
878
877
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 :=
880
879
begin
881
880
refine induction_on_union _ (λ a b hi d, (hi d.symm).symm) (λ a d, empty_right a) (λ a b ab, _) _,
882
881
{ obtain ⟨U, V, oU, oV, aU, bV, UV⟩ := t2_separation (finset.disjoint_singleton.1 ab),
@@ -888,7 +887,7 @@ begin
888
887
end
889
888
890
889
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 :=
892
891
by exact_mod_cast finset_disjoint_finset_opens_of_t2 {x} s (finset.disjoint_singleton_left.mpr h)
893
892
894
893
end separated
@@ -1135,8 +1134,8 @@ lemma function.left_inverse.closed_embedding [t2_space α] {f : α → β} {g :
1135
1134
1136
1135
lemma compact_compact_separated [t2_space α] {s t : set α}
1137
1136
(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;
1140
1139
exact generalized_tube_lemma hs ht is_closed_diagonal.is_open_compl hst
1141
1140
1142
1141
/-- In a `t2_space`, every compact set is closed. -/
@@ -1526,12 +1525,11 @@ section normality
1526
1525
omits T₂), is one in which for every pair of disjoint closed sets `C` and `D`,
1527
1526
there exist disjoint open sets containing `C` and `D` respectively. -/
1528
1527
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)
1531
1529
1532
1530
theorem normal_separation [normal_space α] {s t : set α}
1533
1531
(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 :=
1535
1533
normal_space.normal s t H1 H2 H3
1536
1534
1537
1535
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
1562
1560
normal :=
1563
1561
begin
1564
1562
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 _ _)
1570
1567
end }
1571
1568
1572
1569
variable (α)
@@ -1652,7 +1649,7 @@ instance [t5_space α] {p : α → Prop} : t5_space {x // p x} := embedding_subt
1652
1649
/-- A `T₅` space is a `T₄` space. -/
1653
1650
@[priority 100 ] -- see Note [lower instance priority]
1654
1651
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 $
1656
1653
completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq])⟩
1657
1654
1658
1655
end completely_normal
0 commit comments