Separation properties of topological spaces #
This file defines some of the weaker separation axioms (under the Kolmogorov classification),
notably T₀, R₀, T₁ and R₁ spaces. For T₂ (Hausdorff) spaces and other stronger
conditions, see the file Topology/Separation/Hausdorff.lean
.
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
.T0Space
: A T₀/Kolmogorov space is a space where, for every two pointsx ≠ y
, there is an open set that contains one, but not the other.R0Space
: An R₀ space (sometimes called a symmetric space) is a topological space such that theSpecializes
relation is symmetric.T1Space
: A T₁/Fréchet space is a space where every singleton set is closed. This is equivalent to, for every pairx ≠ y
, there existing an open set containingx
but noty
(t1Space_iff_exists_open
shows that these conditions are equivalent.) T₁ implies T₀ and R₀.R1Space
: An R₁/preregular space is a space where any two topologically distinguishable points have disjoint neighbourhoods. R₁ implies R₀.
Note that mathlib
adopts the modern convention that m ≤ n
if and only if T_m → T_n
, but
occasionally the literature swaps definitions for e.g. T₃ and regular.
Main results #
T₀ spaces #
IsClosed.exists_closed_singleton
: Given a closed setS
in a compact T₀ space, there is somex ∈ S
such that{x}
is closed.exists_isOpen_singleton_of_isOpen_finite
: Given an open finite setS
in a T₀ space, there is somex ∈ S
such that{x}
is open.
T₁ spaces #
isClosedMap_const
: The constant map is a closed map.Finite.instDiscreteTopology
: A finite T₁ space must have the discrete topology.
References #
A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair
x ≠ y
, there is an open set containing one but not the other. We formulate the definition in terms
of the Inseparable
relation.
- t0 ⦃x y : X⦄ : Inseparable x y → x = y
Two inseparable points in a T₀ space are equal.
Instances
A topology inducing map from a T₀ space is injective.
Alias of Topology.IsInducing.injective
.
A topology inducing map from a T₀ space is injective.
A topology inducing map from a T₀ space is a topological embedding.
Alias of Topology.IsInducing.isEmbedding
.
A topology inducing map from a T₀ space is a topological embedding.
Alias of Topology.IsInducing.isEmbedding
.
A topology inducing map from a T₀ space is a topological embedding.
Alias of isEmbedding_iff_isInducing
.
Alias of isEmbedding_iff_isInducing
.
Specialization forms a partial order on a t0 topological space.
Equations
Instances For
Given a closed set S
in a compact T₀ space, there is some x ∈ S
such that {x}
is
closed.
Given an open finite set S
in a T₀ space, there is some x ∈ S
such that {x}
is open.
Alias of Topology.IsEmbedding.t0Space
.
A topological space is called an R₀ space, if Specializes
relation is symmetric.
In other words, given two points x y : X
,
if every neighborhood of y
contains x
, then every neighborhood of x
contains y
.
- specializes_symmetric : Symmetric Specializes
In an R₀ space, the
Specializes
relation is symmetric.
Instances
In an R₀ space, the Specializes
relation is symmetric, dot notation version.
In an R₀ space, the Specializes
relation is symmetric, Iff
version.
In an R₀ space, Specializes
is equivalent to Inseparable
.
In an R₀ space, Specializes
implies Inseparable
.
Alias of Topology.IsInducing.r0Space
.
In an R₀ space, the closure of a singleton is a compact set.
In an R₀ space, relatively compact sets form a bornology.
Its cobounded filter is Filter.coclosedCompact
.
See also Bornology.inCompact
the bornology of sets contained in a compact set.
Equations
- Bornology.relativelyCompact X = { cobounded' := Filter.coclosedCompact X, le_cofinite' := ⋯ }
Instances For
In an R₀ space, the closure of a finite set is a compact set.
A T₁ space, also known as a Fréchet space, is a topological space
where every singleton set is closed. Equivalently, for every pair
x ≠ y
, there is an open set containing x
and not y
.
- t1 (x : X) : IsClosed {x}
A singleton in a T₁ space is a closed set.
Instances
Alias of Topology.IsEmbedding.t1Space
.
If t
is a subset of s
, except for one point,
then insert x s
is a neighborhood of x
within t
.
Removing a non-isolated point from a dense set, one still obtains a dense set.
Removing a finset from a dense set in a space without isolated points, one still obtains a dense set.
Removing a finite set from a dense set in a space without isolated points, one still obtains a dense set.
If a function to a T1Space
tends to some limit y
at some point x
, then necessarily
y = f x
.
Alias of the forward direction of continuousWithinAt_insert
.
Alias of the reverse direction of continuousWithinAt_insert
.
See also continuousWithinAt_diff_self
for the case y = x
but not requiring T1Space
.
If two sets coincide locally around x
, except maybe at y
, then it is equivalent to be
continuous at x
within one set or the other.
To prove a function to a T1Space
is continuous at some point x
, it suffices to prove that
f
admits some limit at x
.
A point with a finite neighborhood has to be isolated.
If the punctured neighborhoods of a point form a nontrivial filter, then any neighborhood is infinite.
A non-trivial connected T1 space has no isolated points.
The neighbourhoods filter of x
within s
, under the discrete topology, is equal to
the pure x
filter (which is the principal filter at the singleton {x}
.)
A point x
in a discrete subset s
of a topological space admits a neighbourhood
that only meets s
at x
.
Let x
be a point in a discrete subset s
of a topological space, then there exists an open
set that only meets s
at x
.
For point x
in a discrete subset s
of a topological space, there is a set U
such that
U
is a punctured neighborhood ofx
(ie.U ∪ {x}
is a neighbourhood ofx
),U
is disjoint froms
.
Alias of isClosedEmbedding_update
.
R₁ (preregular) spaces #
Limits are unique up to separability.
A weaker version of tendsto_nhds_unique
for R1Space
.
In an R₁ space, a point belongs to the closure of a compact set K
if and only if it is topologically inseparable from some point of K
.
In an R₁ space, the closure of a compact set is the union of the closures of its points.
The closure of a compact set in an R₁ space is a compact set.
If K
and L
are disjoint compact sets in an R₁ topological space
and L
is also closed, then K
and L
have disjoint neighborhoods.
If a compact set is covered by two open sets, then we can cover it by two compact subsets.
For every finite open cover Uᵢ
of a compact set, there exists a compact cover Kᵢ ⊆ Uᵢ
.
Alias of Topology.IsInducing.r1Space
.
If a point in an R₁ space has a compact neighborhood, then it has a basis of compact closed neighborhoods.
In an R₁ space, the filters coclosedCompact
and cocompact
are equal.
In an R₁ space, the bornologies relativelyCompact
and inCompact
are equal.
Lemmas about a weakly locally compact R₁ space #
In fact, a space with these properties is locally compact and regular. Some lemmas are formulated using the latter assumptions below.
In a (weakly) locally compact R₁ space, compact closed neighborhoods of a point x
form a basis of neighborhoods of x
.
In a (weakly) locally compact R₁ space, each point admits a compact closed neighborhood.
A weakly locally compact R₁ space is locally compact.
In a weakly locally compact R₁ space, every compact set has an open neighborhood with compact closure.
In a weakly locally compact R₁ space, every point has an open neighborhood with compact closure.