Documentation

Mathlib.Topology.Separation.Hausdorff

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 #

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 #

If the space is also compact:

References #

class T2Space (X : Type u) [TopologicalSpace X] :

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.

Instances
    theorem t2Space_iff (X : Type u) [TopologicalSpace X] :
    T2Space X Pairwise fun (x y : X) => ∃ (u : Set X) (v : Set X), IsOpen u IsOpen v x u y v Disjoint u v
    theorem t2_separation {X : Type u_1} [TopologicalSpace X] [T2Space X] {x y : X} (h : x y) :
    ∃ (u : Set X) (v : Set X), IsOpen u IsOpen v x u y v Disjoint u v

    Two different points can be separated by open sets.

    theorem t2Space_iff_disjoint_nhds {X : Type u_1} [TopologicalSpace X] :
    T2Space X Pairwise fun (x y : X) => Disjoint (nhds x) (nhds y)
    @[simp]
    theorem disjoint_nhds_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] {x y : X} :
    Disjoint (nhds x) (nhds y) x y
    theorem pairwise_disjoint_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] :
    Pairwise (Disjoint on nhds)
    theorem Set.pairwiseDisjoint_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] (s : Set X) :
    s.PairwiseDisjoint nhds
    theorem Set.Finite.t2_separation {X : Type u_1} [TopologicalSpace X] [T2Space X] {s : Set X} (hs : s.Finite) :
    ∃ (U : XSet X), (∀ (x : X), x U x IsOpen (U x)) s.PairwiseDisjoint U

    Points of a finite set can be separated by open sets from each other.

    @[instance 100]
    instance T2Space.t1Space {X : Type u_1} [TopologicalSpace X] [T2Space X] :
    @[instance 100]
    instance T2Space.r1Space {X : Type u_1} [TopologicalSpace X] [T2Space X] :
    @[instance 80]
    theorem t2_iff_nhds {X : Type u_1} [TopologicalSpace X] :
    T2Space X ∀ {x y : X}, (nhds x nhds y).NeBotx = y

    A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter.

    theorem eq_of_nhds_neBot {X : Type u_1} [TopologicalSpace X] [T2Space X] {x y : X} (h : (nhds x nhds y).NeBot) :
    x = y
    theorem t2Space_iff_nhds {X : Type u_1} [TopologicalSpace X] :
    T2Space X Pairwise fun (x y : X) => Unhds x, Vnhds y, Disjoint U V
    theorem t2_separation_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] {x y : X} (h : x y) :
    ∃ (u : Set X) (v : Set X), u nhds x v nhds y Disjoint u v
    theorem t2_separation_compact_nhds {X : Type u_1} [TopologicalSpace X] [LocallyCompactSpace X] [T2Space X] {x y : X} (h : x y) :
    ∃ (u : Set X) (v : Set X), u nhds x v nhds y IsCompact u IsCompact v Disjoint u v
    theorem t2_iff_ultrafilter {X : Type u_1} [TopologicalSpace X] :
    T2Space X ∀ {x y : X} (f : Ultrafilter X), f nhds xf nhds yx = y
    theorem tendsto_nhds_unique {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {f : YX} {l : Filter Y} {a b : X} [l.NeBot] (ha : Filter.Tendsto f l (nhds a)) (hb : Filter.Tendsto f l (nhds b)) :
    a = b
    theorem tendsto_nhds_unique' {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {f : YX} {l : Filter Y} {a b : X} :
    l.NeBot∀ (ha : Filter.Tendsto f l (nhds a)) (hb : Filter.Tendsto f l (nhds b)), a = b
    theorem tendsto_nhds_unique_of_eventuallyEq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {f g : YX} {l : Filter Y} {a b : X} [l.NeBot] (ha : Filter.Tendsto f l (nhds a)) (hb : Filter.Tendsto g l (nhds b)) (hfg : f =ᶠ[l] g) :
    a = b
    theorem tendsto_nhds_unique_of_frequently_eq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {f g : YX} {l : Filter Y} {a b : X} (ha : Filter.Tendsto f l (nhds a)) (hb : Filter.Tendsto g l (nhds b)) (hfg : ∃ᶠ (x : Y) in l, f x = g x) :
    a = b
    theorem IsCompact.nhdsSet_inter_eq {X : Type u_1} [TopologicalSpace X] [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) :

    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.

    theorem IsCompact.separation_of_not_mem {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} {t : Set X} (H1 : IsCompact t) (H2 : xt) :
    ∃ (U : Set X) (V : Set X), IsOpen U IsOpen V t U x V Disjoint U V

    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.

    theorem IsCompact.disjoint_nhdsSet_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} {t : Set X} (H1 : IsCompact t) (H2 : xt) :

    In a T2Space X, for a compact set t and a point x outside t, 𝓝ˢ t and 𝓝 x are disjoint.

    theorem Set.InjOn.exists_mem_nhdsSet {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} {s : Set X} (inj : Set.InjOn f s) (sc : IsCompact s) (fc : xs, ContinuousAt f x) (loc : xs, unhds x, Set.InjOn f u) :
    tnhdsSet s, Set.InjOn f t

    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.

    theorem Set.InjOn.exists_isOpen_superset {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} {s : Set X} (inj : Set.InjOn f s) (sc : IsCompact s) (fc : xs, ContinuousAt f x) (loc : xs, unhds x, Set.InjOn f u) :
    ∃ (t : Set X), IsOpen t s t Set.InjOn f t

    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.

    theorem lim_eq {X : Type u_1} [TopologicalSpace X] [T2Space X] {f : Filter X} {x : X} [f.NeBot] (h : f nhds x) :
    lim f = x
    theorem lim_eq_iff {X : Type u_1} [TopologicalSpace X] [T2Space X] {f : Filter X} [f.NeBot] (h : ∃ (x : X), f nhds x) {x : X} :
    lim f = x f nhds x
    theorem Ultrafilter.lim_eq_iff_le_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] [CompactSpace X] {x : X} {F : Ultrafilter X} :
    F.lim = x F nhds x
    theorem isOpen_iff_ultrafilter' {X : Type u_1} [TopologicalSpace X] [T2Space X] [CompactSpace X] (U : Set X) :
    IsOpen U ∀ (F : Ultrafilter X), F.lim UU F
    theorem Filter.Tendsto.limUnder_eq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {x : X} {f : Filter Y} [f.NeBot] {g : YX} (h : Filter.Tendsto g f (nhds x)) :
    limUnder f g = x
    theorem Filter.limUnder_eq_iff {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] {f : Filter Y} [f.NeBot] {g : YX} (h : ∃ (x : X), Filter.Tendsto g f (nhds x)) {x : X} :
    theorem Continuous.limUnder_eq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] [TopologicalSpace Y] {f : YX} (h : Continuous f) (y : Y) :
    limUnder (nhds y) f = f y
    @[simp]
    theorem lim_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] (x : X) :
    lim (nhds x) = x
    @[simp]
    theorem limUnder_nhds_id {X : Type u_1} [TopologicalSpace X] [T2Space X] (x : X) :
    limUnder (nhds x) id = x
    @[simp]
    theorem lim_nhdsWithin {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} {s : Set X} (h : x closure s) :
    lim (nhdsWithin x s) = x
    @[simp]
    theorem limUnder_nhdsWithin_id {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} {s : Set X} (h : x closure s) :
    limUnder (nhdsWithin x s) id = x

    T2Space constructions #

    We use two lemmas to prove that various standard constructions generate Hausdorff spaces from Hausdorff spaces:

    @[instance 100]
    theorem separated_by_continuous {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) {x y : X} (h : f x f y) :
    ∃ (u : Set X) (v : Set X), IsOpen u IsOpen v x u y v Disjoint u v
    theorem separated_by_isOpenEmbedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} (hf : Topology.IsOpenEmbedding f) {x y : X} (h : x y) :
    ∃ (u : Set Y) (v : Set Y), IsOpen u IsOpen v f x u f y v Disjoint u v
    @[deprecated separated_by_isOpenEmbedding]
    theorem separated_by_openEmbedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} (hf : Topology.IsOpenEmbedding f) {x y : X} (h : x y) :
    ∃ (u : Set Y) (v : Set Y), IsOpen u IsOpen v f x u f y v Disjoint u v

    Alias of separated_by_isOpenEmbedding.

    instance instT2SpaceSubtype {X : Type u_1} [TopologicalSpace X] {p : XProp} [T2Space X] :
    instance Prod.t2Space {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] [TopologicalSpace Y] [T2Space Y] :
    T2Space (X × Y)
    theorem T2Space.of_injective_continuous {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hinj : Function.Injective f) (hc : Continuous f) :

    If the codomain of an injective continuous function is a Hausdorff space, then so is its domain.

    theorem Topology.IsEmbedding.t2Space {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Topology.IsEmbedding f) :

    If the codomain of a topological embedding is a Hausdorff space, then so is its domain. See also T2Space.of_continuous_injective.

    @[deprecated Topology.IsEmbedding.t2Space]
    theorem Embedding.t2Space {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Topology.IsEmbedding f) :

    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.

    instance instT2SpaceSum {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [T2Space X] [TopologicalSpace Y] [T2Space Y] :
    T2Space (X Y)
    instance Pi.t2Space {X : Type u_1} {Y : XType v} [(a : X) → TopologicalSpace (Y a)] [∀ (a : X), T2Space (Y a)] :
    T2Space ((a : X) → Y a)
    instance Sigma.t2Space {ι : Type u_4} {X : ιType u_3} [(i : ι) → TopologicalSpace (X i)] [∀ (a : ι), T2Space (X a)] :
    T2Space ((i : ι) × X i)
    def t2Setoid (X : Type u_1) [TopologicalSpace X] :

    The smallest equivalence relation on a topological space giving a T2 quotient.

    Equations
    Instances For
      def t2Quotient (X : Type u_1) [TopologicalSpace X] :
      Type u_1

      The largest T2 quotient of a topological space. This construction is left-adjoint to the inclusion of T2 spaces into all topological spaces.

      Equations
      Instances For
        def t2Quotient.mk {X : Type u_1} [TopologicalSpace X] :
        Xt2Quotient X

        The map from a topological space to its largest T2 quotient.

        Equations
        Instances For
          theorem t2Quotient.mk_eq {X : Type u_1} [TopologicalSpace X] {x y : X} :
          t2Quotient.mk x = t2Quotient.mk y ∀ (s : Setoid X), T2Space (Quotient s)s x y
          theorem t2Quotient.continuous_mk (X : Type u_1) [TopologicalSpace X] :
          Continuous t2Quotient.mk
          theorem t2Quotient.inductionOn {X : Type u_1} [TopologicalSpace X] {motive : t2Quotient XProp} (q : t2Quotient X) (h : ∀ (x : X), motive (t2Quotient.mk x)) :
          motive q
          theorem t2Quotient.inductionOn₂ {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {motive : t2Quotient Xt2Quotient YProp} (q : t2Quotient X) (q' : t2Quotient Y) (h : ∀ (x : X) (y : Y), motive (t2Quotient.mk x) (t2Quotient.mk y)) :
          motive q q'

          The largest T2 quotient of a topological space is indeed T2.

          theorem t2Quotient.compatible {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) (a b : X) :
          a bf a = f b
          def t2Quotient.lift {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) :
          t2Quotient XY

          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
          Instances For
            theorem t2Quotient.continuous_lift {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) :
            @[simp]
            theorem t2Quotient.lift_mk {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) (x : X) :
            theorem t2Quotient.unique_lift {X : Type u_3} {Y : Type u_4} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f : XY} (hf : Continuous f) {g : t2Quotient XY} (hfg : g t2Quotient.mk = f) :
            theorem isClosed_eq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f g : YX} (hf : Continuous f) (hg : Continuous g) :
            IsClosed {y : Y | f y = g y}
            theorem IsClosed.isClosed_eq {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {f g : XY} {s : Set X} (hs : IsClosed s) (hf : ContinuousOn f s) (hg : ContinuousOn g s) :
            IsClosed {x : X | x s f x = g x}

            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.

            theorem isOpen_ne_fun {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f g : YX} (hf : Continuous f) (hg : Continuous g) :
            IsOpen {y : Y | f y g y}
            theorem Set.EqOn.closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {s : Set Y} {f g : YX} (h : Set.EqOn f g s) (hf : Continuous f) (hg : Continuous g) :

            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.

            theorem Continuous.ext_on {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {s : Set Y} (hs : Dense s) {f g : YX} (hf : Continuous f) (hg : Continuous g) (h : Set.EqOn f g s) :
            f = g

            If two continuous functions are equal on a dense set, then they are equal.

            theorem eqOn_closure₂' {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] {Z : Type u_3} [TopologicalSpace Y] [TopologicalSpace Z] [T2Space Z] {s : Set X} {t : Set Y} {f g : XYZ} (h : xs, yt, f x y = g x y) (hf₁ : ∀ (x : X), Continuous (f x)) (hf₂ : ∀ (y : Y), Continuous fun (x : X) => f x y) (hg₁ : ∀ (x : X), Continuous (g x)) (hg₂ : ∀ (y : Y), Continuous fun (x : X) => g x y) (x : X) :
            x closure syclosure t, f x y = g x y
            theorem eqOn_closure₂ {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] {Z : Type u_3} [TopologicalSpace Y] [TopologicalSpace Z] [T2Space Z] {s : Set X} {t : Set Y} {f g : XYZ} (h : xs, yt, f x y = g x y) (hf : Continuous (Function.uncurry f)) (hg : Continuous (Function.uncurry g)) (x : X) :
            x closure syclosure t, f x y = g x y
            theorem Set.EqOn.of_subset_closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {s t : Set X} {f g : XY} (h : Set.EqOn f g s) (hf : ContinuousOn f t) (hg : ContinuousOn g t) (hst : s t) (hts : t closure s) :
            Set.EqOn f g t

            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.

            theorem Function.LeftInverse.isClosed_range {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :
            @[deprecated Function.LeftInverse.isClosed_range]
            theorem Function.LeftInverse.closed_range {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :

            Alias of Function.LeftInverse.isClosed_range.

            theorem Function.LeftInverse.isClosedEmbedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :
            @[deprecated Function.LeftInverse.isClosedEmbedding]
            theorem Function.LeftInverse.closedEmbedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space X] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :

            Alias of Function.LeftInverse.isClosedEmbedding.

            theorem SeparatedNhds.of_isCompact_isCompact {X : Type u_1} [TopologicalSpace X] [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) (hst : Disjoint s t) :

            In a T2Space X, for disjoint closed sets s t such that closure sᶜ is compact, there are neighbourhoods that separate s and t.

            theorem SeparatedNhds.of_finset_finset {X : Type u_1} [TopologicalSpace X] [T2Space X] (s t : Finset X) (h : Disjoint s t) :
            SeparatedNhds s t
            theorem SeparatedNhds.of_singleton_finset {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} {s : Finset X} (h : xs) :
            SeparatedNhds {x} s
            theorem IsCompact.isClosed {X : Type u_1} [TopologicalSpace X] [T2Space X] {s : Set X} (hs : IsCompact s) :

            In a T2Space, every compact set is closed.

            theorem IsCompact.preimage_continuous {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [CompactSpace X] [T2Space Y] {f : XY} {s : Set Y} (hs : IsCompact s) (hf : Continuous f) :
            theorem Pi.isCompact_iff {ι : Type u_4} {π : ιType u_5} [(i : ι) → TopologicalSpace (π i)] [∀ (i : ι), T2Space (π i)] {s : Set ((i : ι) → π i)} :
            theorem Pi.isCompact_closure_iff {ι : Type u_4} {π : ιType u_5} [(i : ι) → TopologicalSpace (π i)] [∀ (i : ι), T2Space (π i)] {s : Set ((i : ι) → π i)} :
            theorem exists_subset_nhds_of_isCompact {X : Type u_1} [TopologicalSpace X] [T2Space X] {ι : Type u_4} [Nonempty ι] {V : ιSet X} (hV : Directed (fun (x1 x2 : Set X) => x1 x2) V) (hV_cpct : ∀ (i : ι), IsCompact (V i)) {U : Set X} (hU : x⋂ (i : ι), V i, U nhds x) :
            ∃ (i : ι), V i U

            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.

            theorem IsCompact.inter {X : Type u_1} [TopologicalSpace X] [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) :
            theorem image_closure_of_isCompact {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] {s : Set X} (hs : IsCompact (closure s)) {f : XY} (hf : ContinuousOn f (closure s)) :
            f '' closure s = closure (f '' s)
            theorem Continuous.isClosedMap {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [CompactSpace X] [T2Space Y] {f : XY} (h : Continuous f) :

            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.

            @[deprecated Continuous.isClosedEmbedding]

            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.

            @[deprecated IsQuotientMap.of_surjective_continuous]

            Alias of IsQuotientMap.of_surjective_continuous.


            A continuous surjective map from a compact space to a Hausdorff space is a quotient map.

            theorem IsPreirreducible.subsingleton {X : Type u_1} [TopologicalSpace X] [T2Space X] {S : Set X} (h : IsPreirreducible S) :
            S.Subsingleton
            theorem isIrreducible_iff_singleton {X : Type u_1} [TopologicalSpace X] [T2Space X] {S : Set X} :
            IsIrreducible S ∃ (x : X), S = {x}

            There does not exist a nontrivial preirreducible T₂ space.