Urysohn's lemma #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
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
f
equals zero ons
;f
equals one ont
;0 ≤ f x ≤ 1
for allx
.
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 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
:
c.approx 0
is the indicator ofc.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 allx
;c.lim
equals zero onc.C
and equals one outside ofc.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 fory
in a small neighborhood ofx
we havey ∈ c.left.U ⊆ c.right.C
(hencec.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 fory
in a small neighborhood ofx
we 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
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
- urysohns.CU.has_sizeof_inst
- urysohns.CU.inhabited
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)
.
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)
.
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 (n + 1) c x = midpoint ℝ (urysohns.CU.approx n c.left x) (urysohns.CU.approx n c.right x)
- urysohns.CU.approx 0 c x = c.Uᶜ.indicator 1 x
A continuous function f : X → ℝ
such that
0 ≤ f x ≤ 1
for allx
;f
equals zero onc.C
and equals one outside ofc.U
;
Continuity of urysohns.CU.lim
. See module docstring for a sketch of the proofs.
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 ons
;f
equals one ont
;0 ≤ f x ≤ 1
for allx
.