T₂ and T₂.₅ spaces. #
This file defines the T₂ (Hausdorff) condition, which is the most commonly-used among the various separation axioms, and the related T₂.₅ condition.
Main definitions #
T2Space
: A T₂/Hausdorff space is a space where, for every two pointsx ≠ y
, there is two disjoint open sets, one containingx
, and the othery
. T₂ implies T₁ and R₁.T25Space
: A T₂.₅/Urysohn space is a space where, for every two pointsx ≠ y
, there is two open sets, one containingx
, and the othery
, whose closures are disjoint. T₂.₅ implies T₂.
See Mathlib.Topology.Separation.Regular
for regular, T₃, etc spaces; and
Mathlib.Topology.Separation.GDelta
for the definitions of PerfectlyNormalSpace
and T6Space
.
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 #
t2_iff_nhds
: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter.t2_iff_isClosed_diagonal
: A space is T₂ iff thediagonal
ofX
(that is, the set of all points of the form(a, a) : X × X
) is closed under the product topology.separatedNhds_of_finset_finset
: Any two disjoint finsets areSeparatedNhds
.- Most topological constructions preserve Hausdorffness;
these results are part of the typeclass inference system (e.g.
Topology.IsEmbedding.t2Space
) Set.EqOn.closure
: If two functions are equal on some sets
, they are equal on its closure.IsCompact.isClosed
: All compact sets are closed.WeaklyLocallyCompactSpace.locallyCompactSpace
: If a topological space is both weakly locally compact (i.e., each point has a compact neighbourhood) and is T₂, then it is locally compact.totallySeparatedSpace_of_t1_of_basis_clopen
: IfX
has a clopen basis, then it is aTotallySeparatedSpace
.loc_compact_t2_tot_disc_iff_tot_sep
: A locally compact T₂ space is totally disconnected iff it is totally separated.t2Quotient
: the largest T2 quotient of a given topological space.
If the space is also compact:
normalOfCompactT2
: A compact T₂ space is aNormalSpace
.connectedComponent_eq_iInter_isClopen
: The connected component of a point is the intersection of all its clopen neighbourhoods.compact_t2_tot_disc_iff_tot_sep
: Being aTotallyDisconnectedSpace
is equivalent to being aTotallySeparatedSpace
.ConnectedComponents.t2
:ConnectedComponents X
is T₂ forX
T₂ and compact.
References #
A T₂ space, also known as a Hausdorff space, is one in which for every
x ≠ y
there exists disjoint open sets around x
and y
. This is
the most widely used of the separation axioms.
- t2 : Pairwise fun (x y : X) => ∃ (u : Set X) (v : Set X), IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v
Every two points in a Hausdorff space admit disjoint open neighbourhoods.
Instances
A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter.
If s
and t
are compact sets in a T₂ space, then the set neighborhoods filter of s ∩ t
is the infimum of set neighborhoods filters for s
and t
.
For general sets, only the ≤
inequality holds, see nhdsSet_inter_le
.
In a T2Space X
, for a compact set t
and a point x
outside t
, there are open sets U
,
V
that separate t
and x
.
If a function f
is
- injective on a compact set
s
; - continuous at every point of this set;
- injective on a neighborhood of each point of this set,
then it is injective on a neighborhood of this set.
If a function f
is
- injective on a compact set
s
; - continuous at every point of this set;
- injective on a neighborhood of each point of this set,
then it is injective on an open neighborhood of this set.
Properties of lim
and limUnder
#
In this section we use explicit Nonempty X
instances for lim
and limUnder
. This way the lemmas
are useful without a Nonempty X
instance.
T2Space
constructions #
We use two lemmas to prove that various standard constructions generate Hausdorff spaces from Hausdorff spaces:
separated_by_continuous
says that two pointsx y : X
can be separated by open neighborhoods provided that there exists a continuous mapf : X → Y
with a Hausdorff codomain such thatf x ≠ f y
. We use this lemma to prove that topological spaces defined usinginduced
are Hausdorff spaces.separated_by_isOpenEmbedding
says that for an open embeddingf : X → Y
of a Hausdorff spaceX
, the images of two distinct pointsx y : X
,x ≠ y
can be separated by open neighborhoods. We use this lemma to prove that topological spaces defined usingcoinduced
are Hausdorff spaces.
Alias of separated_by_isOpenEmbedding
.
If the codomain of an injective continuous function is a Hausdorff space, then so is its domain.
If the codomain of a topological embedding is a Hausdorff space, then so is its domain.
See also T2Space.of_continuous_injective
.
Alias of Topology.IsEmbedding.t2Space
.
If the codomain of a topological embedding is a Hausdorff space, then so is its domain.
See also T2Space.of_continuous_injective
.
The largest T2 quotient of a topological space. This construction is left-adjoint to the inclusion of T2 spaces into all topological spaces.
Equations
- t2Quotient X = Quotient (t2Setoid X)
Instances For
Equations
- t2Quotient.instTopologicalSpace = inferInstanceAs (TopologicalSpace (Quotient (t2Setoid X)))
The map from a topological space to its largest T2 quotient.
Equations
- t2Quotient.mk = Quotient.mk (t2Setoid X)
Instances For
The largest T2 quotient of a topological space is indeed T2.
The universal property of the largest T2 quotient of a topological space X
: any continuous
map from X
to a T2 space Y
uniquely factors through t2Quotient X
. This declaration builds the
factored map. Its continuity is t2Quotient.continuous_lift
, the fact that it indeed factors the
original map is t2Quotient.lift_mk
and uniquenes is t2Quotient.unique_lift
.
Equations
- t2Quotient.lift hf = Quotient.lift f ⋯
Instances For
If functions f
and g
are continuous on a closed set s
,
then the set of points x ∈ s
such that f x = g x
is a closed set.
If two continuous maps are equal on s
, then they are equal on the closure of s
. See also
Set.EqOn.of_subset_closure
for a more general version.
If two continuous functions are equal on a dense set, then they are equal.
If f x = g x
for all x ∈ s
and f
, g
are continuous on t
, s ⊆ t ⊆ closure s
, then
f x = g x
for all x ∈ t
. See also Set.EqOn.closure
.
Alias of Function.LeftInverse.isClosed_range
.
Alias of Function.LeftInverse.isClosedEmbedding
.
In a T2Space X
, for disjoint closed sets s t
such that closure sᶜ
is compact,
there are neighbourhoods that separate s
and t
.
In a T2Space
, every compact set is closed.
If V : ι → Set X
is a decreasing family of compact sets then any neighborhood of
⋂ i, V i
contains some V i
. This is a version of exists_subset_nhds_of_isCompact'
where we
don't need to assume each V i
closed because it follows from compactness since X
is
assumed to be Hausdorff.
A continuous map from a compact space to a Hausdorff space is a closed map.
A continuous injective map from a compact space to a Hausdorff space is a closed embedding.
Alias of Continuous.isClosedEmbedding
.
A continuous injective map from a compact space to a Hausdorff space is a closed embedding.
A continuous surjective map from a compact space to a Hausdorff space is a quotient map.
Alias of IsQuotientMap.of_surjective_continuous
.
A continuous surjective map from a compact space to a Hausdorff space is a quotient map.
There does not exist a nontrivial preirreducible T₂ space.