mathlib documentation

topology.urysohns_lemma

Urysohn's lemma #

In this file we prove Urysohn's lemma exists_continuous_zero_one_of_closed: for any two disjoint closed sets s and t in a normal topological space X there exists a continuous function f : X → ℝ such that

Implementation notes #

Most paper sources prove Urysohn's lemma using a family of open sets indexed by dyadic rational numbers on [0, 1]. There are many technical difficulties with formalizing this proof (e.g., one needs to formalize the "dyadic induction", then prove that the resulting family of open sets is monotone). So, we formalize a slightly different proof.

Let urysohns.CU be the type of pairs (C, U) of a closed set Cand an open set U such that C ⊆ U. Since X is a normal topological space, for each c : CU X there exists an open set u such that c.C ⊆ u ∧ closure u ⊆ c.U. We define c.left and c.right to be (c.C, u) and (closure u, c.U), respectively. Then we define a family of functions urysohns.CU.approx (c : urysohns.CU X) (n : ℕ) : X → ℝ by recursion on n:

For each x this is a monotone family of functions that are equal to zero on c.C and are equal to one outside of c.U. We also have c.approx n x ∈ [0, 1] for all c, n, and x.

Let urysohns.CU.lim c be the supremum (or equivalently, the limit) of c.approx n. Then properties of urysohns.CU.approx immediately imply that

In order to prove that c.lim is continuous at x, we prove by induction on n : ℕ that for y in a small neighborhood of x we have |c.lim y - c.lim x| ≤ (3 / 4) ^ n. Induction base follows from c.lim x ∈ [0, 1], c.lim y ∈ [0, 1]. For the induction step, consider two cases:

The actual formalization uses midpoint ℝ x y instead of (x + y) / 2 because we have more API lemmas about midpoint.

Tags #

Urysohn's lemma, normal topological space

structure urysohns.CU (X : Type u_2) [topological_space X] :
Type u_2

An auxiliary type for the proof of Urysohn's lemma: a pair of a closed set C and its open neighborhood U.

Instances for urysohns.CU
@[protected, instance]
Equations
@[simp]
theorem urysohns.CU.left_C {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :
c.left.C = c.C
def urysohns.CU.left {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :

Due to normal_exists_closure_subset, for each c : CU X there exists an open set u such chat c.C ⊆ u and closure u ⊆ c.U. c.left is the pair (c.C, u).

Equations
@[simp]
theorem urysohns.CU.right_U {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :
c.right.U = c.U
def urysohns.CU.right {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :

Due to normal_exists_closure_subset, for each c : CU X there exists an open set u such chat c.C ⊆ u and closure u ⊆ c.U. c.right is the pair (closure u, c.U).

Equations
theorem urysohns.CU.left_U_subset {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :
c.left.U c.U
theorem urysohns.CU.subset_right_C {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) :
c.C c.right.C
noncomputable def urysohns.CU.approx {X : Type u_1} [topological_space X] [normal_space X] :
urysohns.CU XX →

n-th approximation to a continuous function f : X → ℝ such that f = 0 on c.C and f = 1 outside of c.U.

Equations
theorem urysohns.CU.approx_of_mem_C {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (n : ) {x : X} (hx : x c.C) :
theorem urysohns.CU.approx_of_nmem_U {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (n : ) {x : X} (hx : x c.U) :
theorem urysohns.CU.approx_nonneg {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (n : ) (x : X) :
theorem urysohns.CU.approx_le_one {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (n : ) (x : X) :
theorem urysohns.CU.bdd_above_range_approx {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
theorem urysohns.CU.approx_le_approx_of_U_sub_C {X : Type u_1} [topological_space X] [normal_space X] {c₁ c₂ : urysohns.CU X} (h : c₁.U c₂.C) (n₁ n₂ : ) (x : X) :
urysohns.CU.approx n₂ c₂ x urysohns.CU.approx n₁ c₁ x
theorem urysohns.CU.approx_le_succ {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (n : ) (x : X) :
theorem urysohns.CU.approx_mono {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
monotone (λ (n : ), urysohns.CU.approx n c x)
@[protected]
noncomputable def urysohns.CU.lim {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :

A continuous function f : X → ℝ such that

  • 0 ≤ f x ≤ 1 for all x;
  • f equals zero on c.C and equals one outside of c.U;
Equations
theorem urysohns.CU.lim_of_mem_C {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) (h : x c.C) :
c.lim x = 0
theorem urysohns.CU.lim_of_nmem_U {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) (h : x c.U) :
c.lim x = 1
theorem urysohns.CU.lim_eq_midpoint {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
c.lim x = midpoint (c.left.lim x) (c.right.lim x)
theorem urysohns.CU.approx_le_lim {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) (n : ) :
theorem urysohns.CU.lim_nonneg {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
0 c.lim x
theorem urysohns.CU.lim_le_one {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
c.lim x 1
theorem urysohns.CU.lim_mem_Icc {X : Type u_1} [topological_space X] [normal_space X] (c : urysohns.CU X) (x : X) :
c.lim x set.Icc 0 1

Continuity of urysohns.CU.lim. See module docstring for a sketch of the proofs.

theorem exists_continuous_zero_one_of_closed {X : Type u_1} [topological_space X] [normal_space X] {s t : set X} (hs : is_closed s) (ht : is_closed t) (hd : disjoint s t) :
∃ (f : C(X, )), set.eq_on f 0 s set.eq_on f 1 t ∀ (x : X), f x set.Icc 0 1

Urysohns lemma: if s and t are two disjoint closed sets in a normal topological space X, then there exists a continuous function f : X → ℝ such that

  • f equals zero on s;
  • f equals one on t;
  • 0 ≤ f x ≤ 1 for all x.