Urysohn's lemma #
In this file we prove Urysohn's lemma exists_continuous_zero_one_of_isClosed: for any two disjoint
closed sets s and t in a normal topological space X there exists a continuous function
f : X → ℝ such that
- fequals zero on- s;
- fequals one on- t;
- 0 ≤ f x ≤ 1for all- x.
We also give versions in a regular locally compact space where one assumes that s is compact
and t is closed, in exists_continuous_zero_one_of_isCompact
and exists_continuous_one_zero_of_isCompact (the latter providing additionally a function with
compact support).
We write a generic proof so that it applies both to normal spaces and to regular locally compact spaces.
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 C and an open set U such that
C ⊆ U. Since X is a normal topological space, for each c : CU 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) (n : ℕ) : X → ℝ by recursion on n:
- c.approx 0is the indicator of- c.Uᶜ;
- c.approx (n + 1) x = (c.left.approx n x + c.right.approx n x) / 2.
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
- c.lim x ∈ [0, 1]for all- x;
- c.limequals zero on- c.Cand equals one outside of- c.U;
- c.lim x = (c.left.lim x + c.right.lim x) / 2.
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:
- x ∈ c.left.U; then for- yin a small neighborhood of- xwe have- y ∈ c.left.U ⊆ c.right.C(hence- c.right.lim x = c.right.lim y = 0) and- |c.left.lim y - c.left.lim x| ≤ (3 / 4) ^ n. Then- |c.lim y - c.lim x| = |c.left.lim y - c.left.lim x| / 2 ≤ (3 / 4) ^ n / 2 < (3 / 4) ^ (n + 1).
- otherwise, x ∉ c.left.right.C; then foryin a small neighborhood ofxwe havey ∉ c.left.right.C ⊇ c.left.left.U(hencec.left.left.lim x = c.left.left.lim y = 1),|c.left.right.lim y - c.left.right.lim x| ≤ (3 / 4) ^ n, and|c.right.lim y - c.right.lim x| ≤ (3 / 4) ^ n. Combining these inequalities, the triangle inequality, and the recurrence formula forc.lim, we get|c.lim x - c.lim y| ≤ (3 / 4) ^ (n + 1).
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, locally compact topological space
An auxiliary type for the proof of Urysohn's lemma: a pair of a closed set C and its open
neighborhood U, together with the assumption that C and U satisfy the property P C U.
The latter assumption will make it possible to prove simultaneously both versions of Urysohn's
lemma, in normal spaces (with P always true) and in locally compact spaces
(with P C U = IsCompact C). We put also in the structure the assumption that, for any such pair,
one may find an intermediate pair in between satisfying P,
to avoid carrying it around in the argument.
- C : Set XThe inner set in the inductive construction towards Urysohn's lemma 
- U : Set XThe outer set in the inductive construction towards Urysohn's lemma 
- hP {c u : Set X} : IsClosed c → P c u → IsOpen u → c ⊆ u → ∃ (v : Set X), IsOpen v ∧ c ⊆ v ∧ closure v ⊆ u ∧ P c v ∧ P (closure v) uThe proof that we can divide CUpairs in half
Instances For
n-th approximation to a continuous function f : X → ℝ such that f = 0 on c.C and f = 1
outside of c.U.
Equations
- Urysohns.CU.approx 0 x✝¹ x✝ = x✝¹.Uᶜ.indicator 1 x✝
- Urysohns.CU.approx n.succ x✝¹ x✝ = midpoint ℝ (Urysohns.CU.approx n x✝¹.left x✝) (Urysohns.CU.approx n x✝¹.right x✝)
Instances For
Alias of Urysohns.CU.approx_of_notMem_U.
A continuous function f : X → ℝ such that
Equations
- c.lim x = ⨆ (n : ℕ), Urysohns.CU.approx n c x
Instances For
Alias of Urysohns.CU.lim_of_notMem_U.
Continuity of Urysohns.CU.lim. See module docstring for a sketch of the proofs.
Urysohn's 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
- fequals zero on- s;
- fequals one on- t;
- 0 ≤ f x ≤ 1for all- x.
Urysohn's lemma: if s and t are two disjoint sets in a regular locally compact topological
space X, with s compact and t closed, then there exists a continuous
function f : X → ℝ such that
- fequals zero on- s;
- fequals one on- t;
- 0 ≤ f x ≤ 1for all- x.
Urysohn's lemma: if s and t are two disjoint sets in a regular locally compact topological
space X, with s compact and t closed, then there exists a continuous
function f : X → ℝ such that
- fequals zero on- t;
- fequals one on- s;
- 0 ≤ f x ≤ 1for all- x.
Urysohn's lemma: if s and t are two disjoint sets in a regular locally compact topological
space X, with s compact and t closed, then there exists a continuous compactly supported
function f : X → ℝ such that
- fequals one on- s;
- fequals zero on- t;
- 0 ≤ f x ≤ 1for all- x.
Urysohn's lemma: if s and t are two disjoint sets in a regular locally compact topological
space X, with s compact and t closed, then there exists a continuous compactly supported
function f : X → ℝ such that
- fequals one on- s;
- fequals zero on- t;
- 0 ≤ f x ≤ 1for all- x.
Moreover, if s is Gδ, one can ensure that f ⁻¹ {1} is exactly s.
A variation of Urysohn's lemma. In a T2Space X, for a closed set t and a relatively
compact open set s such that t ⊆ s, there is a continuous function f supported in s,
f x = 1 on t and 0 ≤ f x ≤ 1.