Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 41e1ab7

Browse files
committed
feat(topology/uniform_space/equicontinuity): definition and basic properties of [uniform] equicontinuity (#16467)
1 parent e753057 commit 41e1ab7

File tree

5 files changed

+595
-50
lines changed

5 files changed

+595
-50
lines changed

src/topology/continuous_function/bounded.lean

Lines changed: 25 additions & 48 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import analysis.normed_space.operator_norm
88
import analysis.normed_space.star.basic
99
import data.real.sqrt
1010
import topology.continuous_function.algebra
11+
import topology.metric_space.equicontinuity
1112

1213
/-!
1314
# Bounded continuous functions
@@ -18,7 +19,7 @@ the uniform distance.
1819
-/
1920

2021
noncomputable theory
21-
open_locale topological_space classical nnreal
22+
open_locale topological_space classical nnreal uniformity uniform_convergence
2223

2324
open set filter metric function
2425

@@ -226,6 +227,20 @@ iff.intro
226227
λ n hn, lt_of_le_of_lt ((dist_le (half_pos ε_pos).le).mpr $
227228
λ x, dist_comm (f x) (F n x) ▸ le_of_lt (hn x)) (half_lt_self ε_pos)))
228229

230+
/-- The topology on `α →ᵇ β` is exactly the topology induced by the natural map to `α →ᵤ β`. -/
231+
lemma inducing_coe_fn : inducing (uniform_fun.of_fun ∘ coe_fn : (α →ᵇ β) → (α →ᵤ β)) :=
232+
begin
233+
rw inducing_iff_nhds,
234+
refine λ f, eq_of_forall_le_iff (λ l, _),
235+
rw [← tendsto_iff_comap, ← tendsto_id', tendsto_iff_tendsto_uniformly,
236+
uniform_fun.tendsto_iff_tendsto_uniformly],
237+
refl
238+
end
239+
240+
-- TODO: upgrade to a `uniform_embedding`
241+
lemma embedding_coe_fn : embedding (uniform_fun.of_fun ∘ coe_fn : (α →ᵇ β) → (α →ᵤ β)) :=
242+
⟨inducing_coe_fn, λ f g h, ext $ λ x, congr_fun h x⟩
243+
229244
variables (α) {β}
230245

231246
/-- Constant as a continuous bounded function. -/
@@ -417,10 +432,10 @@ and several useful variations around it. -/
417432
theorem arzela_ascoli₁ [compact_space β]
418433
(A : set (α →ᵇ β))
419434
(closed : is_closed A)
420-
(H : ∀ (x:α) (ε > 0), ∃U ∈ 𝓝 x, ∀ (y z ∈ U) (f : α →ᵇ β),
421-
f ∈ A → dist (f y) (f z) < ε) :
435+
(H : equicontinuous (coe_fn : A → α → β)) :
422436
is_compact A :=
423437
begin
438+
simp_rw [equicontinuous, metric.equicontinuous_at_iff_pair] at H,
424439
refine is_compact_of_totally_bounded_is_closed _ closed,
425440
refine totally_bounded_of_finite_discretization (λ ε ε0, _),
426441
rcases exists_between ε0 with ⟨ε₁, ε₁0, εε₁⟩,
@@ -437,7 +452,7 @@ begin
437452
f ∈ A → dist (f y) (f z) < ε₂ := λ x,
438453
let ⟨U, nhdsU, hU⟩ := H x _ ε₂0,
439454
⟨V, VU, openV, xV⟩ := _root_.mem_nhds_iff.1 nhdsU in
440-
⟨V, xV, openV, λy hy z hz f hf, hU y (VU hy) z (VU hz) f hf⟩,
455+
⟨V, xV, openV, λy hy z hz f hf, hU y (VU hy) z (VU hz) ⟨f, hf⟩,
441456
choose U hU using this,
442457
/- For all x, the set hU x is an open set containing x on which the elements of A
443458
fluctuate by at most ε₂.
@@ -481,8 +496,7 @@ theorem arzela_ascoli₂
481496
(A : set (α →ᵇ β))
482497
(closed : is_closed A)
483498
(in_s : ∀(f : α →ᵇ β) (x : α), f ∈ A → f x ∈ s)
484-
(H : ∀(x:α) (ε > 0), ∃U ∈ 𝓝 x, ∀ (y z ∈ U) (f : α →ᵇ β),
485-
f ∈ A → dist (f y) (f z) < ε) :
499+
(H : equicontinuous (coe_fn : A → α → β)) :
486500
is_compact A :=
487501
/- This version is deduced from the previous one by restricting to the compact type in the target,
488502
using compactness there and then lifting everything to the original space. -/
@@ -492,10 +506,9 @@ begin
492506
refine is_compact_of_is_closed_subset
493507
((_ : is_compact (F ⁻¹' A)).image (continuous_comp M)) closed (λ f hf, _),
494508
{ haveI : compact_space s := is_compact_iff_compact_space.1 hs,
495-
refine arzela_ascoli₁ _ (continuous_iff_is_closed.1 (continuous_comp M) _ closed)
496-
(λ x ε ε0, bex.imp_right (λ U U_nhds hU y hy z hz f hf, _) (H x ε ε0)),
497-
calc dist (f y) (f z) = dist (F f y) (F f z) : rfl
498-
... < ε : hU y hy z hz (F f) hf },
509+
refine arzela_ascoli₁ _ (continuous_iff_is_closed.1 (continuous_comp M) _ closed) _,
510+
rw uniform_embedding_subtype_coe.to_uniform_inducing.equicontinuous_iff,
511+
exact H.comp (A.restrict_preimage F) },
499512
{ let g := cod_restrict s f (λx, in_s f x hf),
500513
rw [show f = F g, by ext; refl] at hf ⊢,
501514
exact ⟨g, hf, rfl⟩ }
@@ -507,51 +520,15 @@ theorem arzela_ascoli [t2_space β]
507520
(s : set β) (hs : is_compact s)
508521
(A : set (α →ᵇ β))
509522
(in_s : ∀(f : α →ᵇ β) (x : α), f ∈ A → f x ∈ s)
510-
(H : ∀(x:α) (ε > 0), ∃U ∈ 𝓝 x, ∀ (y z ∈ U) (f : α →ᵇ β),
511-
f ∈ A → dist (f y) (f z) < ε) :
523+
(H : equicontinuous (coe_fn : A → α → β)) :
512524
is_compact (closure A) :=
513525
/- This version is deduced from the previous one by checking that the closure of A, in
514526
addition to being closed, still satisfies the properties of compact range and equicontinuity -/
515527
arzela_ascoli₂ s hs (closure A) is_closed_closure
516528
(λ f x hf, (mem_of_closed' hs.is_closed).2 $ λ ε ε0,
517529
let ⟨g, gA, dist_fg⟩ := metric.mem_closure_iff.1 hf ε ε0 in
518530
⟨g x, in_s g x gA, lt_of_le_of_lt (dist_coe_le_dist _) dist_fg⟩)
519-
(λ x ε ε0, show ∃ U ∈ 𝓝 x,
520-
∀ y z ∈ U, ∀ (f : α →ᵇ β), f ∈ closure A → dist (f y) (f z) < ε,
521-
begin
522-
refine bex.imp_right (λ U U_set hU y hy z hz f hf, _) (H x (ε/2) (half_pos ε0)),
523-
rcases metric.mem_closure_iff.1 hf (ε/2/2) (half_pos (half_pos ε0)) with ⟨g, gA, dist_fg⟩,
524-
replace dist_fg := λ x, lt_of_le_of_lt (dist_coe_le_dist x) dist_fg,
525-
calc dist (f y) (f z) ≤ dist (f y) (g y) + dist (f z) (g z) + dist (g y) (g z) :
526-
dist_triangle4_right _ _ _ _
527-
... < ε/2/2 + ε/2/2 + ε/2 :
528-
add_lt_add (add_lt_add (dist_fg y) (dist_fg z)) (hU y hy z hz g gA)
529-
... = ε : by rw [add_halves, add_halves]
530-
end)
531-
532-
/- To apply the previous theorems, one needs to check the equicontinuity. An important
533-
instance is when the source space is a metric space, and there is a fixed modulus of continuity
534-
for all the functions in the set A -/
535-
536-
lemma equicontinuous_of_continuity_modulus {α : Type u} [pseudo_metric_space α]
537-
(b : ℝ → ℝ) (b_lim : tendsto b (𝓝 0) (𝓝 0))
538-
(A : set (α →ᵇ β))
539-
(H : ∀(x y:α) (f : α →ᵇ β), f ∈ A → dist (f x) (f y) ≤ b (dist x y))
540-
(x:α) (ε : ℝ) (ε0 : 0 < ε) : ∃U ∈ 𝓝 x, ∀ (y z ∈ U) (f : α →ᵇ β),
541-
f ∈ A → dist (f y) (f z) < ε :=
542-
begin
543-
rcases tendsto_nhds_nhds.1 b_lim ε ε0 with ⟨δ, δ0, hδ⟩,
544-
refine ⟨ball x (δ/2), ball_mem_nhds x (half_pos δ0), λ y hy z hz f hf, _⟩,
545-
have : dist y z < δ := calc
546-
dist y z ≤ dist y x + dist z x : dist_triangle_right _ _ _
547-
... < δ/2 + δ/2 : add_lt_add hy hz
548-
... = δ : add_halves _,
549-
calc
550-
dist (f y) (f z) ≤ b (dist y z) : H y z f hf
551-
... ≤ |b (dist y z)| : le_abs_self _
552-
... = dist (b (dist y z)) 0 : by simp [real.dist_eq]
553-
... < ε : hδ (by simpa [real.dist_eq] using this),
554-
end
531+
(H.closure' continuous_coe)
555532

556533
end arzela_ascoli
557534

Lines changed: 127 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,127 @@
1+
/-
2+
Copyright (c) 2022 Anatole Dedecker. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Anatole Dedecker
5+
-/
6+
7+
import topology.metric_space.basic
8+
import topology.uniform_space.equicontinuity
9+
/-!
10+
# Equicontinuity in metric spaces
11+
12+
This files contains various facts about (uniform) equicontinuity in metric spaces. Most
13+
importantly, we prove the usual characterization of equicontinuity of `F` at `x₀` in the case of
14+
(pseudo) metric spaces: `∀ ε > 0, ∃ δ > 0, ∀ x, dist x x₀ < δ → ∀ i, dist (F i x₀) (F i x) < ε`,
15+
and we prove that functions sharing a common (local or global) continuity modulus are
16+
(locally or uniformly) equicontinuous.
17+
18+
## Main statements
19+
20+
* `equicontinuous_at_iff`: characterization of equicontinuity for families of functions between
21+
(pseudo) metric spaces.
22+
* `equicontinuous_at_of_continuity_modulus`: convenient way to prove equicontinuity at a point of
23+
a family of functions to a (pseudo) metric space by showing that they share a common *local*
24+
continuity modulus.
25+
* `uniform_equicontinuous_of_continuity_modulus`: convenient way to prove uniform equicontinuity
26+
of a family of functions to a (pseudo) metric space by showing that they share a common *global*
27+
continuity modulus.
28+
29+
## Tags
30+
31+
equicontinuity, continuity modulus
32+
-/
33+
34+
open filter
35+
open_locale topological_space uniformity
36+
37+
variables {α β ι : Type*} [pseudo_metric_space α]
38+
39+
namespace metric
40+
41+
/-- Characterization of equicontinuity for families of functions taking values in a (pseudo) metric
42+
space. -/
43+
lemma equicontinuous_at_iff_right {ι : Type*} [topological_space β] {F : ι → β → α} {x₀ : β} :
44+
equicontinuous_at F x₀ ↔ ∀ ε > 0, ∀ᶠ x in 𝓝 x₀, ∀ i, dist (F i x₀) (F i x) < ε :=
45+
uniformity_basis_dist.equicontinuous_at_iff_right
46+
47+
/-- Characterization of equicontinuity for families of functions between (pseudo) metric spaces. -/
48+
lemma equicontinuous_at_iff {ι : Type*} [pseudo_metric_space β] {F : ι → β → α} {x₀ : β} :
49+
equicontinuous_at F x₀ ↔ ∀ ε > 0, ∃ δ > 0, ∀ x, dist x x₀ < δ → ∀ i, dist (F i x₀) (F i x) < ε :=
50+
nhds_basis_ball.equicontinuous_at_iff uniformity_basis_dist
51+
52+
/-- Reformulation of `equicontinuous_at_iff_pair` for families of functions taking values in a
53+
(pseudo) metric space. -/
54+
protected lemma equicontinuous_at_iff_pair {ι : Type*} [topological_space β] {F : ι → β → α}
55+
{x₀ : β} :
56+
equicontinuous_at F x₀ ↔ ∀ ε > 0, ∃ U ∈ 𝓝 x₀, ∀ (x x' ∈ U), ∀ i, dist (F i x) (F i x') < ε :=
57+
begin
58+
rw equicontinuous_at_iff_pair,
59+
split; intros H,
60+
{ intros ε hε,
61+
refine exists_imp_exists (λ V, exists_imp_exists $ λ hV h, _) (H _ (dist_mem_uniformity hε)),
62+
exact λ x hx x' hx', h _ hx _ hx' },
63+
{ intros U hU,
64+
rcases mem_uniformity_dist.mp hU with ⟨ε, hε, hεU⟩,
65+
refine exists_imp_exists (λ V, exists_imp_exists $ λ hV h, _) (H _ hε),
66+
exact λ x hx x' hx' i, hεU (h _ hx _ hx' i) }
67+
end
68+
69+
/-- Characterization of uniform equicontinuity for families of functions taking values in a
70+
(pseudo) metric space. -/
71+
lemma uniform_equicontinuous_iff_right {ι : Type*} [uniform_space β] {F : ι → β → α} :
72+
uniform_equicontinuous F ↔
73+
∀ ε > 0, ∀ᶠ (xy : β × β) in 𝓤 β, ∀ i, dist (F i xy.1) (F i xy.2) < ε :=
74+
uniformity_basis_dist.uniform_equicontinuous_iff_right
75+
76+
/-- Characterization of uniform equicontinuity for families of functions between
77+
(pseudo) metric spaces. -/
78+
lemma uniform_equicontinuous_iff {ι : Type*} [pseudo_metric_space β] {F : ι → β → α} :
79+
uniform_equicontinuous F ↔
80+
∀ ε > 0, ∃ δ > 0, ∀ x y, dist x y < δ → ∀ i, dist (F i x) (F i y) < ε :=
81+
uniformity_basis_dist.uniform_equicontinuous_iff uniformity_basis_dist
82+
83+
/-- For a family of functions to a (pseudo) metric spaces, a convenient way to prove
84+
equicontinuity at a point is to show that all of the functions share a common *local* continuity
85+
modulus. -/
86+
lemma equicontinuous_at_of_continuity_modulus {ι : Type*} [topological_space β] {x₀ : β}
87+
(b : β → ℝ)
88+
(b_lim : tendsto b (𝓝 x₀) (𝓝 0))
89+
(F : ι → β → α)
90+
(H : ∀ᶠ x in 𝓝 x₀, ∀ i, dist (F i x₀) (F i x) ≤ b x) :
91+
equicontinuous_at F x₀ :=
92+
begin
93+
rw metric.equicontinuous_at_iff_right,
94+
intros ε ε0,
95+
filter_upwards [b_lim (Iio_mem_nhds ε0), H] using λ x hx₁ hx₂ i, (hx₂ i).trans_lt hx₁
96+
end
97+
98+
/-- For a family of functions between (pseudo) metric spaces, a convenient way to prove
99+
uniform equicontinuity is to show that all of the functions share a common *global* continuity
100+
modulus. -/
101+
lemma uniform_equicontinuous_of_continuity_modulus {ι : Type*} [pseudo_metric_space β] (b : ℝ → ℝ)
102+
(b_lim : tendsto b (𝓝 0) (𝓝 0))
103+
(F : ι → β → α)
104+
(H : ∀ (x y : β) i, dist (F i x) (F i y) ≤ b (dist x y)) :
105+
uniform_equicontinuous F :=
106+
begin
107+
rw metric.uniform_equicontinuous_iff,
108+
intros ε ε0,
109+
rcases tendsto_nhds_nhds.1 b_lim ε ε0 with ⟨δ, δ0, hδ⟩,
110+
refine ⟨δ, δ0, λ x y hxy i, _⟩,
111+
calc
112+
dist (F i x) (F i y) ≤ b (dist x y) : H x y i
113+
... ≤ |b (dist x y)| : le_abs_self _
114+
... = dist (b (dist x y)) 0 : by simp [real.dist_eq]
115+
... < ε : hδ (by simpa only [real.dist_eq, tsub_zero, abs_dist] using hxy)
116+
end
117+
118+
/-- For a family of functions between (pseudo) metric spaces, a convenient way to prove
119+
equicontinuity is to show that all of the functions share a common *global* continuity modulus. -/
120+
lemma equicontinuous_of_continuity_modulus {ι : Type*} [pseudo_metric_space β] (b : ℝ → ℝ)
121+
(b_lim : tendsto b (𝓝 0) (𝓝 0))
122+
(F : ι → β → α)
123+
(H : ∀ (x y : β) i, dist (F i x) (F i y) ≤ b (dist x y)) :
124+
equicontinuous F :=
125+
(uniform_equicontinuous_of_continuity_modulus b b_lim F H).equicontinuous
126+
127+
end metric

src/topology/metric_space/gromov_hausdorff_realized.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -269,7 +269,7 @@ begin
269269
{ have : tendsto (λ (t : ℝ), 2 * (max_var X Y : ℝ) * t) (𝓝 0) (𝓝 (2 * max_var X Y * 0)) :=
270270
tendsto_const_nhds.mul tendsto_id,
271271
simpa using this },
272-
{ assume x y f hf,
272+
{ rintros x y ⟨f, hf,
273273
exact (candidates_lipschitz hf).dist_le_mul _ _ } }
274274
end
275275

src/topology/uniform_space/compact.lean

Lines changed: 17 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Patrick Massot, Yury Kudryashov
55
-/
66
import topology.uniform_space.uniform_convergence
7+
import topology.uniform_space.equicontinuity
78
import topology.separation
89

910
/-!
@@ -178,7 +179,7 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp
178179
### Heine-Cantor theorem
179180
-/
180181

181-
/-- Heine-Cantor: a continuous function on a compact separated uniform space is uniformly
182+
/-- Heine-Cantor: a continuous function on a compact uniform space is uniformly
182183
continuous. -/
183184
lemma compact_space.uniform_continuous_of_continuous [compact_space α]
184185
{f : α → β} (h : continuous f) : uniform_continuous f :=
@@ -220,3 +221,18 @@ locally compact and `β` is compact. -/
220221
lemma continuous.tendsto_uniformly [locally_compact_space α] [compact_space β] [uniform_space γ]
221222
(f : α → β → γ) (h : continuous ↿f) (x : α) : tendsto_uniformly f (f x) (𝓝 x) :=
222223
h.continuous_on.tendsto_uniformly univ_mem
224+
225+
section uniform_convergence
226+
227+
/-- An equicontinuous family of functions defined on a compact uniform space is automatically
228+
uniformly equicontinuous. -/
229+
lemma compact_space.uniform_equicontinuous_of_equicontinuous {ι : Type*} {F : ι → β → α}
230+
[compact_space β] (h : equicontinuous F) :
231+
uniform_equicontinuous F :=
232+
begin
233+
rw equicontinuous_iff_continuous at h,
234+
rw uniform_equicontinuous_iff_uniform_continuous,
235+
exact compact_space.uniform_continuous_of_continuous h
236+
end
237+
238+
end uniform_convergence

0 commit comments

Comments
 (0)