Separated neighbourhoods #
This file defines the predicates SeparatedNhds
and HasSeparatingCover
, which are used in
formulating separation axioms for topological spaces.
Main definitions #
SeparatedNhds
: TwoSet
s are separated by neighbourhoods if they are contained in disjoint open sets.HasSeparatingCover
: A set has a countable cover that can be used withhasSeparatingCovers_iff_separatedNhds
to witness when twoSet
s haveSeparatedNhds
.
References #
SeparatedNhds
is a predicate on pairs of subSet
s of a topological space. It holds if the two
subSet
s are contained in disjoint open sets.
Equations
Instances For
theorem
separatedNhds_iff_disjoint
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
:
SeparatedNhds s t ↔ Disjoint (nhdsSet s) (nhdsSet t)
theorem
SeparatedNhds.disjoint_nhdsSet
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
:
SeparatedNhds s t → Disjoint (nhdsSet s) (nhdsSet t)
Alias of the forward direction of separatedNhds_iff_disjoint
.
theorem
hasSeparatingCovers_iff_separatedNhds
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
:
HasSeparatingCover s t ∧ HasSeparatingCover t s ↔ SeparatedNhds s t
Used to prove that a regular topological space with Lindelöf topology is a normal space, and a perfectly normal space is a completely normal space.
theorem
HasSeparatingCover.mono
{X : Type u_1}
[TopologicalSpace X]
{s₁ s₂ t₁ t₂ : Set X}
(sc_st : HasSeparatingCover s₂ t₂)
(s_sub : s₁ ⊆ s₂)
(t_sub : t₁ ⊆ t₂)
:
HasSeparatingCover s₁ t₁
theorem
SeparatedNhds.symm
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
:
SeparatedNhds s t → SeparatedNhds t s
theorem
SeparatedNhds.comm
{X : Type u_1}
[TopologicalSpace X]
(s t : Set X)
:
SeparatedNhds s t ↔ SeparatedNhds t s
theorem
SeparatedNhds.preimage
{X : Type u_1}
{Y : Type u_2}
[TopologicalSpace X]
[TopologicalSpace Y]
{f : X → Y}
{s t : Set Y}
(h : SeparatedNhds s t)
(hf : Continuous f)
:
SeparatedNhds (f ⁻¹' s) (f ⁻¹' t)
theorem
SeparatedNhds.disjoint
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
(h : SeparatedNhds s t)
:
Disjoint s t
theorem
SeparatedNhds.disjoint_closure_left
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
(h : SeparatedNhds s t)
:
theorem
SeparatedNhds.disjoint_closure_right
{X : Type u_1}
[TopologicalSpace X]
{s t : Set X}
(h : SeparatedNhds s t)
:
@[simp]
@[simp]
theorem
SeparatedNhds.mono
{X : Type u_1}
[TopologicalSpace X]
{s₁ s₂ t₁ t₂ : Set X}
(h : SeparatedNhds s₂ t₂)
(hs : s₁ ⊆ s₂)
(ht : t₁ ⊆ t₂)
:
SeparatedNhds s₁ t₁
theorem
SeparatedNhds.union_left
{X : Type u_1}
[TopologicalSpace X]
{s t u : Set X}
:
SeparatedNhds s u → SeparatedNhds t u → SeparatedNhds (s ∪ t) u
theorem
SeparatedNhds.union_right
{X : Type u_1}
[TopologicalSpace X]
{s t u : Set X}
(ht : SeparatedNhds s t)
(hu : SeparatedNhds s u)
:
SeparatedNhds s (t ∪ u)