@@ -3,6 +3,7 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Heather Macbeth
5
5
-/
6
+ import measure_theory.continuous_map_dense
6
7
import measure_theory.l2_space
7
8
import measure_theory.haar_measure
8
9
import analysis.complex.circle
@@ -12,14 +13,17 @@ import topology.continuous_function.stone_weierstrass
12
13
13
14
# Fourier analysis on the circle
14
15
15
- This file contains some first steps towards Fourier series.
16
+ This file contains basic technical results for a development of Fourier series.
16
17
17
18
## Main definitions
18
19
19
20
* `haar_circle`, Haar measure on the circle, normalized to have total measure `1`
20
21
* instances `measure_space`, `probability_measure` for the circle with respect to this measure
21
22
* for `n : ℤ`, `fourier n` is the monomial `λ z, z ^ n`, bundled as a continuous map from `circle`
22
23
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`
23
27
24
28
## Main statements
25
29
@@ -28,25 +32,34 @@ dense in `C(circle, ℂ)`, i.e. that its `submodule.topological_closure` is `⊤
28
32
the Stone-Weierstrass theorem after checking that it is a subalgebra, closed under conjugation, and
29
33
separates points.
30
34
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.
33
39
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².
35
46
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
39
48
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.
42
52
43
53
-/
44
54
45
55
noncomputable theory
56
+ open_locale ennreal
46
57
open topological_space continuous_map measure_theory measure_theory.measure algebra submodule set
47
58
48
59
local attribute [instance] fact_one_le_two_ennreal
49
60
61
+ /-! ### Choice of measure on the circle -/
62
+
50
63
section haar_circle
51
64
/-! We make the circle into a measure space, using the Haar measure normalized to have total
52
65
measure 1. -/
@@ -65,6 +78,8 @@ instance : measure_space circle :=
65
78
66
79
end haar_circle
67
80
81
+ /-! ### Monomials on the circle -/
82
+
68
83
section fourier
69
84
70
85
/-- 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
142
157
fourier_subalgebra_conj_invariant
143
158
144
159
/-- 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 = ⊤ :=
146
161
begin
147
162
rw ← fourier_subalgebra_coe,
148
163
exact congr_arg subalgebra.to_submodule fourier_subalgebra_closure_eq_top,
149
164
end
150
165
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
+
151
182
/-- For `n ≠ 0`, a rotation by `n⁻¹ * real.pi` negates the monomial `z ^ n`. -/
152
183
lemma fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0 ) (z : circle) :
153
184
fourier n ((exp_map_circle (n⁻¹ * real.pi) * z)) = - fourier n z :=
@@ -160,7 +191,7 @@ begin
160
191
end
161
192
162
193
/-- 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 ) :=
164
195
begin
165
196
rw orthonormal_iff_ite,
166
197
intros i j,
0 commit comments