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

Commit 02ecb62

Browse files
committed
feat(analysis/fourier): span of monomials is dense in L^p (#8328)
1 parent 5ccf2bf commit 02ecb62

File tree

5 files changed

+86
-25
lines changed

5 files changed

+86
-25
lines changed

docs/undergrad.yaml

Lines changed: 11 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -352,7 +352,7 @@ Single Variable Real Analysis:
352352
normal convergence:
353353
continuity of the limit: 'continuous_of_uniform_approx_of_continuous'
354354
differentiability of the limit:
355-
Weierstrass polynomial approximation theorem:
355+
Weierstrass polynomial approximation theorem: 'polynomial_functions_closure_eq_top'
356356
Weierstrass trigonometric approximation theorem:
357357
Convexity:
358358
convex functions of a real variable: 'convex_on'
@@ -421,27 +421,29 @@ Topology:
421421
continuous linear maps: 'continuous_linear_map'
422422
norm of a continuous linear map: 'linear_map.mk_continuous'
423423
uniform convergence norm (sup-norm): 'emetric.tendsto_uniformly_on_iff'
424-
normed space of bounded continuous normed-space-valued functions: 'bounded_continuous_function.normed_space'
424+
normed space of bounded continuous functions: 'bounded_continuous_function.normed_space'
425425
its completeness: 'bounded_continuous_function.complete_space'
426426
Heine-Borel theorem (closed bounded subsets are compact in finite dimension): 'finite_dimensional.proper'
427427
Riesz' lemma (unit-ball characterization of finite dimension):
428428
Arzela-Ascoli theorem: 'bounded_continuous_function.arzela_ascoli'
429-
Hilbert Spaces:
429+
Hilbert spaces:
430430
Hilbert projection theorem: 'exists_norm_eq_infi_of_complete_convex'
431431
orthogonal projection onto closed vector subspaces: 'orthogonal_projection_fn'
432432
dual space: 'normed_space.dual.normed_space'
433433
Riesz representation theorem: 'inner_product_space.to_dual'
434434
inner product spaces $l^2$ and $L^2$: 'measure_theory.L2.inner_product_space'
435435
their completeness: 'measure_theory.Lp.complete_space'
436-
Hilbert (orthonormal) bases (in the separable case): 'exists_is_orthonormal_dense_span'
437-
Hilbert basis of trigonometric polynomials:
438-
Hilbert bases of orthogonal polynomials:
436+
Hilbert bases, i.e. complete orthonormal sets (in the separable case): 'exists_is_orthonormal_dense_span'
437+
example, the Hilbert basis of trigonometric polynomials: 'fourier_Lp'
438+
its orthonormality: 'orthonormal_fourier'
439+
its completeness: 'span_fourier_Lp_closure_eq_top'
440+
example, classical Hilbert bases of orthogonal polynomials, their orthonormality and completeness:
439441
Lax-Milgram theorem:
440442
$H^1_0([0,1])$ and its application to the one-dimensional Dirichlet problem:
441443

442444
# 9.
443445
Multivariable calculus:
444-
Differential Calculus:
446+
Differential calculus:
445447
differentiable functions on an open subset of $\R^n$: 'differentiable_on'
446448
differentials (linear tangent functions): 'fderiv'
447449
directional derivative:
@@ -506,7 +508,7 @@ Measures and integral Calculus:
506508
integrable functions: 'measure_theory.integrable'
507509
dominated convergence theorem: 'measure_theory.tendsto_integral_of_dominated_convergence'
508510
finite dimensional vector-valued integrable functions: 'measure_theory.integrable'
509-
continuity of integrals with respect to parameters:
511+
continuity of integrals with respect to parameters: 'interval_integral.continuous_of_dominated_interval'
510512
differentiability of integrals with respect to parameters:
511513
$\mathrm{L}^p$ spaces where $1 ≤ p ≤ ∞$: 'measure_theory.Lp'
512514
Completeness of $\mathrm{L}^p$ spaces: 'measure_theory.Lp.complete_space'
@@ -524,7 +526,7 @@ Measures and integral Calculus:
524526
Dirichlet theorem:
525527
Fejer theorem:
526528
Parseval theorem:
527-
Fourier transforms on $\mathrm{L}^1(\R^d)$ and $\mathrm{L}^2(R^d)$:
529+
Fourier transforms on $\mathrm{L}^1(\R^d)$ and $\mathrm{L}^2(\R^d)$:
528530
Plancherel’s theorem:
529531

530532
# 11.

src/analysis/fourier.lean

Lines changed: 42 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Heather Macbeth
55
-/
6+
import measure_theory.continuous_map_dense
67
import measure_theory.l2_space
78
import measure_theory.haar_measure
89
import analysis.complex.circle
@@ -12,14 +13,17 @@ import topology.continuous_function.stone_weierstrass
1213
1314
# Fourier analysis on the circle
1415
15-
This file contains some first steps towards Fourier series.
16+
This file contains basic technical results for a development of Fourier series.
1617
1718
## Main definitions
1819
1920
* `haar_circle`, Haar measure on the circle, normalized to have total measure `1`
2021
* instances `measure_space`, `probability_measure` for the circle with respect to this measure
2122
* for `n : ℤ`, `fourier n` is the monomial `λ z, z ^ n`, bundled as a continuous map from `circle`
2223
to `ℂ`
24+
* for `n : ℤ` and `p : ℝ≥0∞`, `fourier_Lp p n` is an abbreviation for the monomial `fourier n`
25+
considered as an element of the Lᵖ-space `Lp ℂ p haar_circle`, via the embedding
26+
`continuous_map.to_Lp`
2327
2428
## Main statements
2529
@@ -28,25 +32,34 @@ dense in `C(circle, ℂ)`, i.e. that its `submodule.topological_closure` is `⊤
2832
the Stone-Weierstrass theorem after checking that it is a subalgebra, closed under conjugation, and
2933
separates points.
3034
31-
The theorem `orthonormal_fourier` states that the functions `fourier n`, when sent via
32-
`continuous_map.to_Lp` to the L^2 space on the circle, form an orthonormal set.
35+
The theorem `span_fourier_Lp_closure_eq_top` states that for `1 ≤ p < ∞` the span of the monomials
36+
`fourier_Lp` is dense in `Lp ℂ p haar_circle`, i.e. that its `submodule.topological_closure` is
37+
`⊤`. This follows from the previous theorem using general theory on approximation of Lᵖ functions
38+
by continuous functions.
3339
34-
## TODO
40+
The theorem `orthonormal_fourier` states that the monomials `fourier_Lp 2 n` form an orthonormal
41+
set (in the L² space of the circle).
42+
43+
By definition, a Hilbert basis for an inner product space is an orthonormal set whose span is
44+
dense. Thus, the last two results together establish that the functions `fourier_Lp 2 n` form a
45+
Hilbert basis for L².
3546
36-
Show that the image of `submodule.span fourier` under `continuous_map.to_Lp` is dense in the `L^2`
37-
space on the circle. This follows from `span_fourier_closure_eq_top` using general theory (not yet
38-
in Lean) on approximation by continuous functions.
47+
## TODO
3948
40-
Paired with `orthonormal_fourier`, this establishes that the functions `fourier` form a Hilbert
41-
basis for `L^2`.
49+
Once mathlib has general theory showing that a Hilbert basis of an inner product space induces a
50+
unitary equivalence with L², the results in this file will give Fourier series applications such
51+
as Parseval's formula.
4252
4353
-/
4454

4555
noncomputable theory
56+
open_locale ennreal
4657
open topological_space continuous_map measure_theory measure_theory.measure algebra submodule set
4758

4859
local attribute [instance] fact_one_le_two_ennreal
4960

61+
/-! ### Choice of measure on the circle -/
62+
5063
section haar_circle
5164
/-! We make the circle into a measure space, using the Haar measure normalized to have total
5265
measure 1. -/
@@ -65,6 +78,8 @@ instance : measure_space circle :=
6578

6679
end haar_circle
6780

81+
/-! ### Monomials on the circle -/
82+
6883
section fourier
6984

7085
/-- The family of monomials `λ z, z ^ n`, parametrized by `n : ℤ` and considered as bundled
@@ -142,12 +157,28 @@ continuous_map.subalgebra_complex_topological_closure_eq_top_of_separates_points
142157
fourier_subalgebra_conj_invariant
143158

144159
/-- The linear span of the monomials `z ^ n` is dense in `C(circle, ℂ)`. -/
145-
lemma span_fourier_closure_eq_top : (span ℂ (set.range fourier)).topological_closure = ⊤ :=
160+
lemma span_fourier_closure_eq_top : (span ℂ (range fourier)).topological_closure = ⊤ :=
146161
begin
147162
rw ← fourier_subalgebra_coe,
148163
exact congr_arg subalgebra.to_submodule fourier_subalgebra_closure_eq_top,
149164
end
150165

166+
/-- The family of monomials `λ z, z ^ n`, parametrized by `n : ℤ` and considered as elements of
167+
the `Lp` space of functions on `circle` taking values in `ℂ`. -/
168+
abbreviation fourier_Lp (p : ℝ≥0∞) [fact (1 ≤ p)] (n : ℤ) : Lp ℂ p haar_circle :=
169+
to_Lp p haar_circle ℂ (fourier n)
170+
171+
/-- For each `1 ≤ p < ∞`, the linear span of the monomials `z ^ n` is dense in
172+
`Lp ℂ p haar_circle`. -/
173+
lemma span_fourier_Lp_closure_eq_top {p : ℝ≥0∞} [fact (1 ≤ p)] (hp : p ≠ ∞) :
174+
(span ℂ (range (fourier_Lp p))).topological_closure = ⊤ :=
175+
begin
176+
convert (continuous_map.to_Lp_dense_range ℂ hp haar_circle ℂ).topological_closure_map_submodule
177+
span_fourier_closure_eq_top,
178+
rw [map_span, range_comp],
179+
simp
180+
end
181+
151182
/-- For `n ≠ 0`, a rotation by `n⁻¹ * real.pi` negates the monomial `z ^ n`. -/
152183
lemma fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0) (z : circle) :
153184
fourier n ((exp_map_circle (n⁻¹ * real.pi) * z)) = - fourier n z :=
@@ -160,7 +191,7 @@ begin
160191
end
161192

162193
/-- The monomials `z ^ n` are an orthonormal set with respect to Haar measure on the circle. -/
163-
lemma orthonormal_fourier : orthonormal ℂ (λ n, to_Lp 2 haar_circle ℂ (fourier n)) :=
194+
lemma orthonormal_fourier : orthonormal ℂ (fourier_Lp 2) :=
164195
begin
165196
rw orthonormal_iff_ite,
166197
intros i j,

src/measure_theory/continuous_map_dense.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -19,10 +19,10 @@ The result is presented in several versions:
1919
`measure_theory.Lp.bounded_continuous_function` of `Lp E p μ`, the additive subgroup of
2020
`Lp E p μ` consisting of equivalence classes containing a continuous representative, is dense in
2121
`Lp E p μ`.
22-
* `bounded_continuous_function.dense_range`: For finite-measure `μ`, the continuous linear map
23-
`bounded_continuous_function.to_Lp p μ 𝕜` from `α →ᵇ E` to `Lp E p μ` has dense range.
24-
* `continuous_map.dense_range`: For compact `α` and finite-measure `μ`, the continuous linear map
25-
`continuous_map.to_Lp p μ 𝕜` from `C(α, E)` to `Lp E p μ` has dense range.
22+
* `bounded_continuous_function.to_Lp_dense_range`: For finite-measure `μ`, the continuous linear
23+
map `bounded_continuous_function.to_Lp p μ 𝕜` from `α →ᵇ E` to `Lp E p μ` has dense range.
24+
* `continuous_map.to_Lp_dense_range`: For compact `α` and finite-measure `μ`, the continuous linear
25+
map `continuous_map.to_Lp p μ 𝕜` from `C(α, E)` to `Lp E p μ` has dense range.
2626
2727
Note that for `p = ∞` this result is not true: the characteristic function of the set `[0, ∞)` in
2828
`ℝ` cannot be continuously approximated in `L∞`.

src/topology/algebra/group.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -230,6 +230,10 @@ def subgroup.topological_closure (s : subgroup G) : subgroup G :=
230230
inv_mem' := λ g m, by simpa [←mem_inv, inv_closure] using m,
231231
..s.to_submonoid.topological_closure }
232232

233+
@[simp, to_additive] lemma subgroup.topological_closure_coe {s : subgroup G} :
234+
(s.topological_closure : set G) = closure s :=
235+
rfl
236+
233237
@[to_additive]
234238
instance subgroup.topological_closure_topological_group (s : subgroup G) :
235239
topological_group (s.topological_closure) :=
@@ -254,6 +258,16 @@ by convert is_closed_closure
254258
s.topological_closure ≤ t :=
255259
closure_minimal h ht
256260

261+
@[to_additive] lemma dense_range.topological_closure_map_subgroup [group H] [topological_space H]
262+
[topological_group H] {f : G →* H} (hf : continuous f) (hf' : dense_range f) {s : subgroup G}
263+
(hs : s.topological_closure = ⊤) :
264+
(s.map f).topological_closure = ⊤ :=
265+
begin
266+
rw set_like.ext'_iff at hs ⊢,
267+
simp only [subgroup.topological_closure_coe, subgroup.coe_top, ← dense_iff_closure_eq] at hs ⊢,
268+
exact hf'.dense_image hf hs
269+
end
270+
257271
@[to_additive exists_nhds_half_neg]
258272
lemma exists_nhds_split_inv {s : set G} (hs : s ∈ 𝓝 (1 : G)) :
259273
∃ V ∈ 𝓝 (1 : G), ∀ (v ∈ V) (w ∈ V), v / w ∈ s :=

src/topology/algebra/module.lean

Lines changed: 15 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -247,9 +247,23 @@ contained in the `topological_closure` of its image. -/
247247
lemma _root_.submodule.topological_closure_map [topological_space R] [has_continuous_smul R M]
248248
[has_continuous_add M] [has_continuous_smul R M₂] [has_continuous_add M₂] (f : M →L[R] M₂)
249249
(s : submodule R M) :
250-
(s.topological_closure.map f.to_linear_map) ≤ (s.map f.to_linear_map).topological_closure :=
250+
(s.topological_closure.map ↑f) ≤ (s.map (f : M →ₗ[R] M₂)).topological_closure :=
251251
image_closure_subset_closure_image f.continuous
252252

253+
/-- Under a dense continuous linear map, a submodule whose `topological_closure` is `⊤` is sent to
254+
another such submodule. That is, the image of a dense set under a map with dense range is dense.
255+
-/
256+
lemma _root_.dense_range.topological_closure_map_submodule [topological_space R]
257+
[has_continuous_smul R M] [has_continuous_add M] [has_continuous_smul R M₂]
258+
[has_continuous_add M₂] {f : M →L[R] M₂} (hf' : dense_range f) {s : submodule R M}
259+
(hs : s.topological_closure = ⊤) :
260+
(s.map (f : M →ₗ[R] M₂)).topological_closure = ⊤ :=
261+
begin
262+
rw set_like.ext'_iff at hs ⊢,
263+
simp only [submodule.topological_closure_coe, submodule.top_coe, ← dense_iff_closure_eq] at hs ⊢,
264+
exact hf'.dense_image f.continuous hs
265+
end
266+
253267
/-- The continuous map that is constantly zero. -/
254268
instance: has_zero (M →L[R] M₂) := ⟨⟨0, continuous_zero⟩⟩
255269
instance : inhabited (M →L[R] M₂) := ⟨0

0 commit comments

Comments
 (0)