mathlib documentation

topology.separation

@[class]
structure t0_space (α : Type u) [topological_space α] :
Prop

A T₀ space, also known as a Kolmogorov space, is a topological space where for every pair x ≠ y, there is an open set containing one but not the other.

Instances
theorem exists_open_singleton_of_open_finset {α : Type u} [topological_space α] [t0_space α] (s : finset α) (sne : s.nonempty) (hso : is_open s) :
∃ (x : α) (H : x s), is_open {x}

theorem exists_open_singleton_of_fintype {α : Type u} [topological_space α] [t0_space α] [f : fintype α] [ha : nonempty α] :
∃ (x : α), is_open {x}

@[instance]
def subtype.t0_space {α : Type u} [topological_space α] [t0_space α] {p : α → Prop} :

@[class]
structure t1_space (α : Type u) [topological_space α] :
Prop

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.

Instances
theorem is_closed_singleton {α : Type u} [topological_space α] [t1_space α] {x : α} :

theorem is_open_ne {α : Type u} [topological_space α] [t1_space α] {x : α} :
is_open {y : α | y x}

@[instance]
def subtype.t1_space {α : Type u} [topological_space α] [t1_space α] {p : α → Prop} :

@[instance]
def t1_space.t0_space {α : Type u} [topological_space α] [t1_space α] :

theorem compl_singleton_mem_nhds {α : Type u} [topological_space α] [t1_space α] {x y : α} (h : y x) :
{x} 𝓝 y

@[simp]
theorem closure_singleton {α : Type u} [topological_space α] [t1_space α] {a : α} :
closure {a} = {a}

theorem is_closed_map_const {α : Type u_1} {β : Type u_2} [topological_space α] [topological_space β] [t1_space β] {y : β} :

@[class]
structure t2_space (α : Type u) [topological_space α] :
Prop

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 t2_separation {α : Type u} [topological_space α] [t2_space α] {x y : α} (h : x y) :
∃ (u v : set α), is_open u is_open v x u y v u v =

@[instance]
def t2_space.t1_space {α : Type u} [topological_space α] [t2_space α] :

theorem eq_of_nhds_ne_bot {α : Type u} [topological_space α] [ht : t2_space α] {x y : α} (h : (𝓝 x 𝓝 y).ne_bot) :
x = y

theorem t2_iff_nhds {α : Type u} [topological_space α] :
t2_space α ∀ {x y : α}, (𝓝 x 𝓝 y).ne_botx = y

theorem t2_iff_ultrafilter {α : Type u} [topological_space α] :
t2_space α ∀ (f : filter α) {x y : α}, f.is_ultrafilterf 𝓝 xf 𝓝 yx = y

theorem is_closed_diagonal {α : Type u} [topological_space α] [t2_space α] :

@[simp]
theorem nhds_eq_nhds_iff {α : Type u} [topological_space α] {a b : α} [t2_space α] :
𝓝 a = 𝓝 b a = b

@[simp]
theorem nhds_le_nhds_iff {α : Type u} [topological_space α] {a b : α} [t2_space α] :
𝓝 a 𝓝 b a = b

theorem tendsto_nhds_unique {α : Type u} {β : Type v} [topological_space α] [t2_space α] {f : β → α} {l : filter β} {a b : α} [l.ne_bot] (ha : filter.tendsto f l (𝓝 a)) (hb : filter.tendsto f l (𝓝 b)) :
a = b

theorem tendsto_nhds_unique' {α : Type u} {β : Type v} [topological_space α] [t2_space α] {f : β → α} {l : filter β} {a b : α} (hl : l.ne_bot) (ha : filter.tendsto f l (𝓝 a)) (hb : filter.tendsto f l (𝓝 b)) :
a = b

Properties of Lim and lim

In this section we use explicit nonempty α instances for Lim and lim. This way the lemmas are useful without a nonempty α instance.

theorem Lim_eq {α : Type u} [topological_space α] [t2_space α] {f : filter α} {a : α} [f.ne_bot] (h : f 𝓝 a) :
Lim f = a

theorem filter.tendsto.lim_eq {α : Type u} {β : Type v} [topological_space α] [t2_space α] {a : α} {f : filter β} {g : β → α} (h : filter.tendsto g f (𝓝 a)) [f.ne_bot] :
lim f g = a

theorem continuous.lim_eq {α : Type u} {β : Type v} [topological_space α] [t2_space α] [topological_space β] {f : β → α} (h : continuous f) (a : β) :
lim (𝓝 a) f = f a

@[simp]
theorem Lim_nhds {α : Type u} [topological_space α] [t2_space α] (a : α) :
Lim (𝓝 a) = a

@[simp]
theorem lim_nhds_id {α : Type u} [topological_space α] [t2_space α] (a : α) :
lim (𝓝 a) id = a

@[simp]
theorem Lim_nhds_within {α : Type u} [topological_space α] [t2_space α] {a : α} {s : set α} (h : a closure s) :
Lim (𝓝[s] a) = a

@[simp]
theorem lim_nhds_within_id {α : Type u} [topological_space α] [t2_space α] {a : α} {s : set α} (h : a closure s) :
lim (𝓝[s] a) id = a

@[instance]
def t2_space_discrete {α : Type u_1} [topological_space α] [discrete_topology α] :

@[instance]
def subtype.t2_space {α : Type u_1} {p : α → Prop} [t : topological_space α] [t2_space α] :

@[instance]
def prod.t2_space {α : Type u_1} {β : Type u_2} [t₁ : topological_space α] [t2_space α] [t₂ : topological_space β] [t2_space β] :
t2_space × β)

@[instance]
def Pi.t2_space {α : Type u_1} {β : α → Type v} [t₂ : Π (a : α), topological_space (β a)] [∀ (a : α), t2_space (β a)] :
t2_space (Π (a : α), β a)

theorem is_closed_eq {α : Type u} {β : Type v} [topological_space α] [topological_space β] [t2_space α] {f g : β → α} (hf : continuous f) (hg : continuous g) :
is_closed {x : β | f x = g x}

theorem diagonal_eq_range_diagonal_map {α : Type u_1} :
{p : α × α | p.fst = p.snd} = set.range (λ (x : α), (x, x))

theorem prod_subset_compl_diagonal_iff_disjoint {α : Type u_1} {s t : set α} :
s.prod t {p : α × α | p.fst = p.snd} s t =

theorem compact_compact_separated {α : Type u} [topological_space α] [t2_space α] {s t : set α} (hs : is_compact s) (ht : is_compact t) (hst : s t = ) :
∃ (u v : set α), is_open u is_open v s u t v u v =

theorem is_compact.is_closed {α : Type u} [topological_space α] [t2_space α] {s : set α} (hs : is_compact s) :

theorem is_compact.inter {α : Type u} [topological_space α] [t2_space α] {s t : set α} (hs : is_compact s) (ht : is_compact t) :

theorem is_compact.binary_compact_cover {α : Type u} [topological_space α] [t2_space α] {K U V : set α} (hK : is_compact K) (hU : is_open U) (hV : is_open V) (h2K : K U V) :
∃ (K₁ K₂ : set α), is_compact K₁ is_compact K₂ K₁ U K₂ V K = K₁ K₂

If a compact set is covered by two open sets, then we can cover it by two compact subsets.

theorem is_compact.finite_compact_cover {α : Type u} [topological_space α] [t2_space α] {s : set α} (hs : is_compact s) {ι : Type u_1} (t : finset ι) (U : ι → set α) (hU : ∀ (i : ι), i tis_open (U i)) (hsC : s ⋃ (i : ι) (H : i t), U i) :
∃ (K : ι → set α), (∀ (i : ι), is_compact (K i)) (∀ (i : ι), K i U i) s = ⋃ (i : ι) (H : i t), K i

For every finite open cover Uᵢ of a compact set, there exists a compact cover Kᵢ ⊆ Uᵢ.

theorem locally_compact_of_compact_nhds {α : Type u} [topological_space α] [t2_space α] (h : ∀ (x : α), ∃ (s : set α), s 𝓝 x is_compact s) :

theorem exists_open_with_compact_closure {α : Type u} [topological_space α] [locally_compact_space α] [t2_space α] (x : α) :
∃ (U : set α), is_open U x U is_compact (closure U)

In a locally compact T₂ space, every point has an open neighborhood with compact closure

theorem exists_compact_superset {α : Type u} [topological_space α] [locally_compact_space α] [t2_space α] {K : set α} (hK : is_compact K) :
∃ (K' : set α), is_compact K' K interior K'

In a locally compact T₂ space, every compact set is contained in the interior of a compact set.

@[class]
structure regular_space (α : Type u) [topological_space α] :
Prop

A T₃ space, also known as a regular space (although this condition sometimes omits T₂), is one in which for every closed C and x ∉ C, there exist disjoint open sets containing x and C respectively.

Instances
theorem nhds_is_closed {α : Type u} [topological_space α] [regular_space α] {a : α} {s : set α} (h : s 𝓝 a) :
∃ (t : set α) (H : t 𝓝 a), t s is_closed t

theorem closed_nhds_basis {α : Type u} [topological_space α] [regular_space α] (a : α) :
(𝓝 a).has_basis (λ (s : set α), s 𝓝 a is_closed s) id

@[instance]
def subtype.regular_space {α : Type u} [topological_space α] [regular_space α] {p : α → Prop} :

@[instance]

theorem disjoint_nested_nhds {α : Type u} [topological_space α] [regular_space α] {x y : α} (h : x y) :
∃ (U₁ V₁ : set α) (H : U₁ 𝓝 x) (H : V₁ 𝓝 x) (U₂ V₂ : set α) (H : U₂ 𝓝 y) (H : V₂ 𝓝 y), is_closed V₁ is_closed V₂ is_open U₁ is_open U₂ V₁ U₁ V₂ U₂ U₁ U₂ =

@[class]
structure normal_space (α : Type u) [topological_space α] :
Prop

A T₄ space, also known as a normal space (although this condition sometimes omits T₂), is one in which for every pair of disjoint closed sets C and D, there exist disjoint open sets containing C and D respectively.

theorem normal_separation {α : Type u} [topological_space α] [normal_space α] (s t : set α) (H1 : is_closed s) (H2 : is_closed t) (H3 : disjoint s t) :
∃ (u v : set α), is_open u is_open v s u t v disjoint u v

@[instance]