@@ -26,36 +26,34 @@ This file contains basic results on Fourier series for functions on the additive
26
26
so we do not declare it as a `measure_space` instance, to avoid confusion.)
27
27
* for `n : ℤ`, `fourier n` is the monomial `λ x, exp (2 π i n x / T)`, bundled as a continuous map
28
28
from `add_circle T` to `ℂ`.
29
- * for `n : ℤ` and `p : ℝ≥0∞`, `fourier_Lp p n` is an abbreviation for the monomial `fourier n`
30
- considered as an element of the Lᵖ-space `Lp ℂ p haar_add_circle`, via the embedding
31
- `continuous_map.to_Lp`.
32
- * `fourier_series` is the canonical isometric isomorphism from `Lp ℂ 2 haar_add_circle` to
33
- `ℓ²(ℤ, ℂ)` induced by taking Fourier coefficients.
29
+ * `fourier_basis` is the Hilbert basis of `Lp ℂ 2 haar_add_circle` given by the images of the
30
+ monomials `fourier n`.
31
+ * `fourier_coeff f n`, for `f : add_circle T → ℂ`, is the `n`-th Fourier coefficient of `f`
32
+ (defined as an integral over `add_circle T`).
34
33
35
34
## Main statements
36
35
37
36
The theorem `span_fourier_closure_eq_top` states that the span of the monomials `fourier n` is
38
37
dense in `C(add_circle T, ℂ)`, i.e. that its `submodule.topological_closure` is `⊤`. This follows
39
- from the Stone-Weierstrass theorem after checking that it is a subalgebra, closed under conjugation,
40
- and separates points.
41
-
42
- The theorem `span_fourier_Lp_closure_eq_top` states that for `1 ≤ p < ∞` the span of the monomials
43
- `fourier_Lp` is dense in the Lᵖ space of `add_circle T`, i.e. that its
44
- `submodule.topological_closure` is `⊤`. This follows from the previous theorem using general theory
45
- on approximation of Lᵖ functions by continuous functions.
46
-
47
- The theorem `orthonormal_fourier` states that the monomials `fourier_Lp 2 n` form an orthonormal set
48
- (in the L² space of `add_circle T` with respect to `haar_add_circle`).
49
-
50
- The last two results together provide that the functions `fourier_Lp 2 n` form a Hilbert basis for
51
- L²; this is named as `fourier_series`.
52
-
53
- Parseval's identity, `tsum_sq_fourier_series_repr`, is a direct consequence of the construction of
54
- this Hilbert basis.
38
+ from the Stone-Weierstrass theorem after checking that the span is a subalgebra, is closed under
39
+ conjugation, and separates points.
40
+
41
+ Using this and general theory on approximation of Lᵖ functions by continuous functions, we deduce
42
+ (`span_fourier_Lp_closure_eq_top`) that for any `1 ≤ p < ∞`, the span of the Fourier monomials is
43
+ dense in the Lᵖ space of `add_circle T`. For `p = 2` we show (`orthonormal_fourier`) that the
44
+ monomials are also orthonormal, so they form a Hilbert basis for L², which is named as
45
+ `fourier_basis`; in particular, for `L²` functions `f`, the Fourier series of `f` converges to `f`
46
+ in the `L²` topology (`has_sum_fourier_series_L2`). Parseval's identity, `tsum_sq_fourier_coeff`, is
47
+ a direct consequence.
48
+
49
+ For continuous maps `f : add_circle T → ℂ`, the theorem
50
+ `continuous_map.has_sum_fourier_series_of_summable` states that if the sequence of Fourier
51
+ coefficients of `f` is summable, then the Fourier series `∑ (i:ℤ), f.fourier_coeff i * fourier i`
52
+ converges to `f` in the uniform-convergence topology of `C(add_circle T, ℂ)`.
55
53
-/
56
54
57
55
noncomputable theory
58
- open_locale ennreal complex_conjugate classical real
56
+ open_locale ennreal complex_conjugate real
59
57
open topological_space continuous_map measure_theory measure_theory.measure algebra submodule set
60
58
61
59
variables {T : ℝ}
@@ -135,13 +133,27 @@ def fourier (n : ℤ) : C(add_circle T, ℂ) :=
135
133
136
134
@[simp] lemma fourier_apply {n : ℤ} {x : add_circle T} : fourier n x = to_circle (n • x) := rfl
137
135
136
+ @[simp] lemma fourier_coe_apply {n : ℤ} {x : ℝ} :
137
+ fourier n (x : add_circle T) = complex.exp (2 * π * complex.I * n * x / T) :=
138
+ begin
139
+ rw [fourier_apply, ←quotient_add_group.coe_zsmul, to_circle, function.periodic.lift_coe,
140
+ exp_map_circle_apply, complex.of_real_mul, complex.of_real_div, complex.of_real_mul,
141
+ zsmul_eq_mul, complex.of_real_mul, complex.of_real_int_cast, complex.of_real_bit0,
142
+ complex.of_real_one],
143
+ congr' 1 , ring,
144
+ end
145
+
138
146
@[simp] lemma fourier_zero {x : add_circle T} : fourier 0 x = 1 :=
139
147
begin
140
148
induction x using quotient_add_group.induction_on',
141
- rw [fourier_apply, to_circle, zero_zsmul, ←quotient_add_group.coe_zero ,
142
- function.periodic.lift_coe, mul_zero, exp_map_circle_zero, coe_one_unit_sphere ],
149
+ simp only [fourier_coe_apply, algebra_map.coe_zero, mul_zero, zero_mul ,
150
+ zero_div, complex.exp_zero ],
143
151
end
144
152
153
+ @[simp] lemma fourier_eval_zero (n : ℤ) : fourier n (0 : add_circle T) = 1 :=
154
+ by rw [←quotient_add_group.coe_zero, fourier_coe_apply, complex.of_real_zero,
155
+ mul_zero, zero_div, complex.exp_zero]
156
+
145
157
@[simp] lemma fourier_one {x : add_circle T} : fourier 1 x = to_circle x :=
146
158
by rw [fourier_apply, one_zsmul]
147
159
@@ -153,9 +165,17 @@ begin
153
165
end
154
166
155
167
@[simp] lemma fourier_add {m n : ℤ} {x : add_circle T} :
156
- fourier (m + n) x = ( fourier m x) * ( fourier n x) :=
168
+ fourier (m + n) x = fourier m x * fourier n x :=
157
169
by simp_rw [fourier_apply, add_zsmul, to_circle_add, coe_mul_unit_sphere]
158
170
171
+ lemma fourier_norm [fact (0 < T)] (n : ℤ) : ‖@fourier T n‖ = 1 :=
172
+ begin
173
+ rw continuous_map.norm_eq_supr_norm,
174
+ have : ∀ (x : add_circle T), ‖fourier n x‖ = 1 := λ x, abs_coe_circle _,
175
+ simp_rw this ,
176
+ exact @csupr_const _ _ _ has_zero.nonempty _,
177
+ end
178
+
159
179
/-- For `n ≠ 0`, a translation by `T / 2 / n` negates the function `fourier n`. -/
160
180
lemma fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0 ) (hT : 0 < T) (x : add_circle T) :
161
181
fourier n (x + ((T / 2 / n) : ℝ)) = - fourier n x :=
@@ -266,64 +286,114 @@ end
266
286
267
287
end monomials
268
288
269
- section fourier
289
+ section scope_hT -- everything from here on needs `0 < T`
270
290
variables [hT : fact (0 < T)]
271
291
include hT
272
292
273
- /-- We define `fourier_series` to be a `ℤ`-indexed Hilbert basis for `Lp ℂ 2 haar_add_circle`,
293
+
294
+ section fourier_coeff
295
+ variables {E : Type } [normed_add_comm_group E] [normed_space ℂ E] [complete_space E]
296
+
297
+ /-- The `n`-th Fourier coefficient of a function `add_circle T → E`, for `E` a complete normed
298
+ `ℂ`-vector space, defined as the integral over `add_circle T` of `fourier (-n) t • f t`. -/
299
+ def fourier_coeff (f : add_circle T → E) (n : ℤ) : E :=
300
+ ∫ (t : add_circle T), fourier (-n) t • f t ∂ haar_add_circle
301
+
302
+ /-- The Fourier coefficients of a function can be computed as an integral
303
+ over `[a, a + T]` for any real `a`. -/
304
+ lemma fourier_coeff_eq_interval_integral (f : add_circle T → E) (n : ℤ) (a : ℝ) :
305
+ fourier_coeff f n = (1 / T) • ∫ x in a .. a + T, @fourier T (-n) x • f x :=
306
+ begin
307
+ have : ∀ (x : ℝ), @fourier T (-n) x • f x = (λ (z : add_circle T), @fourier T (-n) z • f z) x,
308
+ { intro x, refl, },
309
+ simp_rw this ,
310
+ rw [fourier_coeff, add_circle.interval_integral_preimage T a,
311
+ volume_eq_smul_haar_add_circle, integral_smul_measure, ennreal.to_real_of_real hT.out.le,
312
+ ←smul_assoc, smul_eq_mul, one_div_mul_cancel hT.out.ne', one_smul],
313
+ end
314
+
315
+ end fourier_coeff
316
+
317
+ section fourier_L2
318
+
319
+ /-- We define `fourier_basis` to be a `ℤ`-indexed Hilbert basis for `Lp ℂ 2 haar_add_circle`,
274
320
which by definition is an isometric isomorphism from `Lp ℂ 2 haar_add_circle` to `ℓ²(ℤ, ℂ)`. -/
275
- def fourier_series : hilbert_basis ℤ ℂ (Lp ℂ 2 $ @haar_add_circle T hT) :=
321
+ def fourier_basis : hilbert_basis ℤ ℂ (Lp ℂ 2 $ @haar_add_circle T hT) :=
276
322
hilbert_basis.mk orthonormal_fourier (span_fourier_Lp_closure_eq_top (by norm_num)).ge
277
323
278
- /-- The elements of the Hilbert basis `fourier_series ` are the functions `fourier_Lp 2`, i.e. the
324
+ /-- The elements of the Hilbert basis `fourier_basis ` are the functions `fourier_Lp 2`, i.e. the
279
325
monomials `fourier n` on the circle considered as elements of `L²`. -/
280
- @[simp] lemma coe_fourier_series : ⇑(@fourier_series _ hT) = fourier_Lp 2 := hilbert_basis.coe_mk _ _
326
+ @[simp] lemma coe_fourier_basis : ⇑(@fourier_basis _ hT) = fourier_Lp 2 := hilbert_basis.coe_mk _ _
281
327
282
- /-- Under the isometric isomorphism `fourier_series` from `Lp ℂ 2 haar_circle` to `ℓ²(ℤ, ℂ)`, the
283
- `i`-th coefficient is the integral over `add_circle T` of `λ t, fourier (-i) t * f t`. -/
284
- lemma fourier_series_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) (i : ℤ) :
285
- fourier_series.repr f i = ∫ (t : add_circle T), fourier (-i) t * f t ∂ haar_add_circle :=
328
+ /-- Under the isometric isomorphism `fourier_basis` from `Lp ℂ 2 haar_circle` to `ℓ²(ℤ, ℂ)`, the
329
+ `i`-th coefficient is `fourier_coeff f i`, i.e., the integral over `add_circle T` of
330
+ `λ t, fourier (-i) t * f t` with respect to the Haar measure of total mass 1. -/
331
+ lemma fourier_basis_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) (i : ℤ) :
332
+ fourier_basis.repr f i = fourier_coeff f i :=
286
333
begin
287
334
transitivity ∫ (t : add_circle T),
288
335
conj (((@fourier_Lp T hT 2 _ i) : add_circle T → ℂ) t) * f t ∂ haar_add_circle,
289
- { simp [fourier_series .repr_apply_apply f i, measure_theory.L2.inner_def] },
336
+ { simp [fourier_basis .repr_apply_apply f i, measure_theory.L2.inner_def] },
290
337
{ apply integral_congr_ae,
291
338
filter_upwards [coe_fn_fourier_Lp 2 i] with _ ht,
292
- rw [ht, ←fourier_neg] }
339
+ rw [ht, ←fourier_neg, smul_eq_mul], }
293
340
end
294
341
295
342
/-- The Fourier series of an `L2` function `f` sums to `f`, in the `L²` space of `add_circle T`. -/
296
- lemma has_sum_fourier_series (f : Lp ℂ 2 $ @haar_add_circle T hT) :
297
- has_sum (λ i, fourier_series.repr f i • fourier_Lp 2 i) f :=
298
- by simpa using hilbert_basis.has_sum_repr fourier_series f
343
+ lemma has_sum_fourier_series_L2 (f : Lp ℂ 2 $ @haar_add_circle T hT) :
344
+ has_sum (λ i, fourier_coeff f i • fourier_Lp 2 i) f :=
345
+ by { simp_rw ←fourier_basis_repr, simpa using hilbert_basis.has_sum_repr fourier_basis f }
299
346
300
347
/-- **Parseval's identity** : for an `L²` function `f` on `add_circle T`, the sum of the squared
301
348
norms of the Fourier coefficients equals the `L²` norm of `f`. -/
302
- lemma tsum_sq_fourier_series_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) :
303
- ∑' i : ℤ, ‖fourier_series.repr f i‖ ^ 2 = ∫ (t : add_circle T), ‖f t‖ ^ 2 ∂ haar_add_circle :=
349
+ lemma tsum_sq_fourier_coeff (f : Lp ℂ 2 $ @haar_add_circle T hT) :
350
+ ∑' i : ℤ, ‖fourier_coeff f i‖ ^ 2 = ∫ (t : add_circle T), ‖f t‖ ^ 2 ∂ haar_add_circle :=
304
351
begin
305
- have H₁ : ‖fourier_series.repr f‖ ^ 2 = ∑' i, ‖fourier_series.repr f i‖ ^ 2 ,
306
- { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_series.repr f),
352
+ simp_rw ←fourier_basis_repr,
353
+ have H₁ : ‖fourier_basis.repr f‖ ^ 2 = ∑' i, ‖fourier_basis.repr f i‖ ^ 2 ,
354
+ { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_basis.repr f),
307
355
norm_num },
308
- have H₂ : ‖fourier_series .repr f‖ ^ 2 = ‖f‖ ^ 2 := by simp,
356
+ have H₂ : ‖fourier_basis .repr f‖ ^ 2 = ‖f‖ ^ 2 := by simp,
309
357
have H₃ := congr_arg is_R_or_C.re (@L2.inner_def (add_circle T) ℂ ℂ _ _ _ _ f f),
310
358
rw ← integral_re at H₃,
311
359
{ simp only [← norm_sq_eq_inner] at H₃,
312
360
rw [← H₁, H₂, H₃], },
313
361
{ exact L2.integrable_inner f f },
314
362
end
315
363
316
- /-- The Fourier coefficients are given by integrating over the interval `[a, a + T] ⊂ ℝ`. -/
317
- lemma fourier_series_repr' (f : Lp ℂ 2 $ @haar_add_circle T hT) (n : ℤ) (a : ℝ):
318
- fourier_series.repr f n = 1 / T * ∫ x in a .. a + T, @fourier T (-n) x * f x :=
364
+ end fourier_L2
365
+
366
+ section convergence
367
+
368
+ variables (f : C(add_circle T, ℂ))
369
+
370
+ lemma fourier_coeff_to_Lp (n : ℤ) :
371
+ fourier_coeff (to_Lp 2 haar_add_circle ℂ f) n = fourier_coeff f n :=
372
+ integral_congr_ae (filter.eventually_eq.mul
373
+ (filter.eventually_of_forall (by tauto))
374
+ (continuous_map.coe_fn_to_ae_eq_fun haar_add_circle f))
375
+
376
+ variables {f}
377
+
378
+ /-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series converges
379
+ uniformly to `f`. -/
380
+ lemma has_sum_fourier_series_of_summable (h : summable (fourier_coeff f)) :
381
+ has_sum (λ i, fourier_coeff f i • fourier i) f :=
319
382
begin
320
- have ha : ae_strongly_measurable (λ (t : add_circle T), fourier (-n) t * f t) haar_add_circle :=
321
- (map_continuous _).ae_strongly_measurable.mul (Lp.ae_strongly_measurable _),
322
- rw [fourier_series_repr, add_circle.interval_integral_preimage T a (ha.smul_measure _),
323
- volume_eq_smul_haar_add_circle, integral_smul_measure],
324
- have : (T : ℂ) ≠ 0 := by exact_mod_cast hT.out.ne',
325
- field_simp [ennreal.to_real_of_real hT.out.le, complex.real_smul],
326
- ring,
383
+ have sum_L2 := has_sum_fourier_series_L2 (to_Lp 2 haar_add_circle ℂ f),
384
+ simp_rw fourier_coeff_to_Lp at sum_L2,
385
+ refine continuous_map.has_sum_of_has_sum_Lp (summable_of_summable_norm _) sum_L2,
386
+ simp_rw [norm_smul, fourier_norm, mul_one, summable_norm_iff],
387
+ exact h,
327
388
end
328
389
329
- end fourier
390
+ /-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series of `f`
391
+ converges everywhere pointwise to `f`. -/
392
+ lemma has_pointwise_sum_fourier_series_of_summable
393
+ (h : summable (fourier_coeff f)) (x : add_circle T) :
394
+ has_sum (λ i, fourier_coeff f i • fourier i x) (f x) :=
395
+ (continuous_map.eval_clm ℂ x).has_sum (has_sum_fourier_series_of_summable h)
396
+
397
+ end convergence
398
+
399
+ end scope_hT
0 commit comments