Documentation

Mathlib.Topology.Separation.SeparatedNhds

Separated neighbourhoods #

This file defines the predicates SeparatedNhds and HasSeparatingCover, which are used in formulating separation axioms for topological spaces.

Main definitions #

References #

def SeparatedNhds {X : Type u_1} [TopologicalSpace X] :
Set XSet XProp

SeparatedNhds is a predicate on pairs of subSets of a topological space. It holds if the two subSets are contained in disjoint open sets.

Equations
Instances For

    Alias of the forward direction of separatedNhds_iff_disjoint.

    def HasSeparatingCover {X : Type u_1} [TopologicalSpace X] :
    Set XSet XProp

    HasSeparatingCovers can be useful witnesses for SeparatedNhds.

    Equations
    Instances For

      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₂) :
      theorem SeparatedNhds.symm {X : Type u_1} [TopologicalSpace X] {s t : Set X} :
      theorem SeparatedNhds.preimage {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s t : Set Y} (h : SeparatedNhds s t) (hf : Continuous f) :
      theorem SeparatedNhds.disjoint {X : Type u_1} [TopologicalSpace X] {s t : Set X} (h : SeparatedNhds s t) :
      @[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} :
      theorem SeparatedNhds.union_right {X : Type u_1} [TopologicalSpace X] {s t u : Set X} (ht : SeparatedNhds s t) (hu : SeparatedNhds s u) :