@@ -8,6 +8,7 @@ import analysis.normed_space.operator_norm
8
8
import analysis.normed_space.star.basic
9
9
import data.real.sqrt
10
10
import topology.continuous_function.algebra
11
+ import topology.metric_space.equicontinuity
11
12
12
13
/-!
13
14
# Bounded continuous functions
@@ -18,7 +19,7 @@ the uniform distance.
18
19
-/
19
20
20
21
noncomputable theory
21
- open_locale topological_space classical nnreal
22
+ open_locale topological_space classical nnreal uniformity uniform_convergence
22
23
23
24
open set filter metric function
24
25
@@ -226,6 +227,20 @@ iff.intro
226
227
λ n hn, lt_of_le_of_lt ((dist_le (half_pos ε_pos).le).mpr $
227
228
λ x, dist_comm (f x) (F n x) ▸ le_of_lt (hn x)) (half_lt_self ε_pos)))
228
229
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
+
229
244
variables (α) {β}
230
245
231
246
/-- Constant as a continuous bounded function. -/
@@ -417,10 +432,10 @@ and several useful variations around it. -/
417
432
theorem arzela_ascoli₁ [compact_space β]
418
433
(A : set (α →ᵇ β))
419
434
(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 → α → β)) :
422
436
is_compact A :=
423
437
begin
438
+ simp_rw [equicontinuous, metric.equicontinuous_at_iff_pair] at H,
424
439
refine is_compact_of_totally_bounded_is_closed _ closed,
425
440
refine totally_bounded_of_finite_discretization (λ ε ε0 , _),
426
441
rcases exists_between ε0 with ⟨ε₁, ε₁0 , εε₁⟩,
@@ -437,7 +452,7 @@ begin
437
452
f ∈ A → dist (f y) (f z) < ε₂ := λ x,
438
453
let ⟨U, nhdsU, hU⟩ := H x _ ε₂0 ,
439
454
⟨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⟩ ⟩,
441
456
choose U hU using this ,
442
457
/- For all x, the set hU x is an open set containing x on which the elements of A
443
458
fluctuate by at most ε₂.
@@ -481,8 +496,7 @@ theorem arzela_ascoli₂
481
496
(A : set (α →ᵇ β))
482
497
(closed : is_closed A)
483
498
(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 → α → β)) :
486
500
is_compact A :=
487
501
/- This version is deduced from the previous one by restricting to the compact type in the target,
488
502
using compactness there and then lifting everything to the original space. -/
@@ -492,10 +506,9 @@ begin
492
506
refine is_compact_of_is_closed_subset
493
507
((_ : is_compact (F ⁻¹' A)).image (continuous_comp M)) closed (λ f hf, _),
494
508
{ 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) },
499
512
{ let g := cod_restrict s f (λx, in_s f x hf),
500
513
rw [show f = F g, by ext; refl] at hf ⊢,
501
514
exact ⟨g, hf, rfl⟩ }
@@ -507,51 +520,15 @@ theorem arzela_ascoli [t2_space β]
507
520
(s : set β) (hs : is_compact s)
508
521
(A : set (α →ᵇ β))
509
522
(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 → α → β)) :
512
524
is_compact (closure A) :=
513
525
/- This version is deduced from the previous one by checking that the closure of A, in
514
526
addition to being closed, still satisfies the properties of compact range and equicontinuity -/
515
527
arzela_ascoli₂ s hs (closure A) is_closed_closure
516
528
(λ f x hf, (mem_of_closed' hs.is_closed).2 $ λ ε ε0 ,
517
529
let ⟨g, gA, dist_fg⟩ := metric.mem_closure_iff.1 hf ε ε0 in
518
530
⟨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)
555
532
556
533
end arzela_ascoli
557
534
0 commit comments