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

Commit 247a102

Browse files
feat(analysis/fourier): convergence of Fourier series (#17913)
This PR adds a straightforward but useful criterion for convergence of Fourier series: for a continuous periodic function `f`, if the sequence of Fourier coefficients of `f` is absolutely summable, then the Fourier series converges uniformly, and hence pointwise everywhere, to `f`. (Note that it is obvious in this case that the Fourier series converges uniformly to _something_, the difficult part is that the limit is actually `f`.) Co-authored-by: Junyan Xu <junyanxu.math@gmail.com>
1 parent 7c60702 commit 247a102

File tree

7 files changed

+236
-90
lines changed

7 files changed

+236
-90
lines changed

docs/100.yaml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -297,8 +297,8 @@
297297
76:
298298
title : Fourier Series
299299
decls :
300-
- fourier_series_repr
301-
- has_sum_fourier_series
300+
- fourier_coeff
301+
- has_sum_fourier_series_L2
302302
author : Heather Macbeth
303303
77:
304304
title : Sum of kth powers

docs/undergrad.yaml

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -436,7 +436,7 @@ Topology:
436436
inner product space $L^2$: 'measure_theory.L2.inner_product_space'
437437
its completeness: 'measure_theory.Lp.complete_space'
438438
Hilbert bases: 'hilbert_basis' # the document specifies "in the separable case" but we don't require this
439-
example, the Hilbert basis of trigonometric polynomials: 'fourier_series'
439+
example, the Hilbert basis of trigonometric polynomials: 'fourier_basis'
440440
example, classical Hilbert bases of orthogonal polynomials: ''
441441
Lax-Milgram theorem: 'is_coercive.continuous_linear_equiv_of_bilin'
442442
$H^1_0([0,1])$ and its application to the one-dimensional Dirichlet problem: ''
@@ -521,12 +521,12 @@ Measures and integral calculus:
521521
approximation by convolution: 'cont_diff_bump_of_inner.convolution_tendsto_right'
522522
regularization by convolution: 'has_compact_support.cont_diff_convolution_left'
523523
Fourier analysis:
524-
Fourier series of locally integrable periodic real-valued functions: 'fourier_series'
524+
Fourier series of locally integrable periodic real-valued functions: 'fourier_basis'
525525
Riemann-Lebesgue lemma: ''
526526
convolution product of periodic functions: ''
527527
Dirichlet theorem: ''
528528
Fejer theorem: ''
529-
Parseval theorem: 'tsum_sq_fourier_series_repr'
529+
Parseval theorem: 'tsum_sq_fourier_coeff'
530530
Fourier transforms on $\mathrm{L}^1(\R^d)$ and $\mathrm{L}^2(\R^d)$: ''
531531
Plancherel’s theorem: ''
532532

src/analysis/fourier/add_circle.lean

Lines changed: 125 additions & 55 deletions
Original file line numberDiff line numberDiff line change
@@ -26,36 +26,34 @@ This file contains basic results on Fourier series for functions on the additive
2626
so we do not declare it as a `measure_space` instance, to avoid confusion.)
2727
* for `n : ℤ`, `fourier n` is the monomial `λ x, exp (2 π i n x / T)`, bundled as a continuous map
2828
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`).
3433
3534
## Main statements
3635
3736
The theorem `span_fourier_closure_eq_top` states that the span of the monomials `fourier n` is
3837
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, ℂ)`.
5553
-/
5654

5755
noncomputable theory
58-
open_locale ennreal complex_conjugate classical real
56+
open_locale ennreal complex_conjugate real
5957
open topological_space continuous_map measure_theory measure_theory.measure algebra submodule set
6058

6159
variables {T : ℝ}
@@ -135,13 +133,27 @@ def fourier (n : ℤ) : C(add_circle T, ℂ) :=
135133

136134
@[simp] lemma fourier_apply {n : ℤ} {x : add_circle T} : fourier n x = to_circle (n • x) := rfl
137135

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+
138146
@[simp] lemma fourier_zero {x : add_circle T} : fourier 0 x = 1 :=
139147
begin
140148
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],
143151
end
144152

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+
145157
@[simp] lemma fourier_one {x : add_circle T} : fourier 1 x = to_circle x :=
146158
by rw [fourier_apply, one_zsmul]
147159

@@ -153,9 +165,17 @@ begin
153165
end
154166

155167
@[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 :=
157169
by simp_rw [fourier_apply, add_zsmul, to_circle_add, coe_mul_unit_sphere]
158170

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+
159179
/-- For `n ≠ 0`, a translation by `T / 2 / n` negates the function `fourier n`. -/
160180
lemma fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0) (hT : 0 < T) (x : add_circle T) :
161181
fourier n (x + ((T / 2 / n) : ℝ)) = - fourier n x :=
@@ -266,64 +286,114 @@ end
266286

267287
end monomials
268288

269-
section fourier
289+
section scope_hT -- everything from here on needs `0 < T`
270290
variables [hT : fact (0 < T)]
271291
include hT
272292

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`,
274320
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) :=
276322
hilbert_basis.mk orthonormal_fourier (span_fourier_Lp_closure_eq_top (by norm_num)).ge
277323

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
279325
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 _ _
281327

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 :=
286333
begin
287334
transitivity ∫ (t : add_circle T),
288335
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] },
290337
{ apply integral_congr_ae,
291338
filter_upwards [coe_fn_fourier_Lp 2 i] with _ ht,
292-
rw [ht, ←fourier_neg] }
339+
rw [ht, ←fourier_neg, smul_eq_mul], }
293340
end
294341

295342
/-- 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 }
299346

300347
/-- **Parseval's identity**: for an `L²` function `f` on `add_circle T`, the sum of the squared
301348
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 :=
304351
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),
307355
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,
309357
have H₃ := congr_arg is_R_or_C.re (@L2.inner_def (add_circle T) ℂ ℂ _ _ _ _ f f),
310358
rw ← integral_re at H₃,
311359
{ simp only [← norm_sq_eq_inner] at H₃,
312360
rw [← H₁, H₂, H₃], },
313361
{ exact L2.integrable_inner f f },
314362
end
315363

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 :=
319382
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,
327388
end
328389

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

src/measure_theory/function/l2_space.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -236,8 +236,8 @@ lemma bounded_continuous_function.inner_to_Lp (f g : α →ᵇ 𝕜) :
236236
= ∫ x, conj (f x) * g x ∂μ :=
237237
begin
238238
apply integral_congr_ae,
239-
have hf_ae := f.coe_fn_to_Lp μ,
240-
have hg_ae := g.coe_fn_to_Lp μ,
239+
have hf_ae := f.coe_fn_to_Lp 2 μ 𝕜,
240+
have hg_ae := g.coe_fn_to_Lp 2 μ 𝕜,
241241
filter_upwards [hf_ae, hg_ae] with _ hf hg,
242242
rw [hf, hg],
243243
simp

src/measure_theory/function/lp_space.lean

Lines changed: 41 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -2826,11 +2826,11 @@ begin
28262826
(by { rintros - ⟨f, rfl⟩, exact mem_Lp f } : _ ≤ Lp E p μ),
28272827
end
28282828

2829-
variables (𝕜 : Type*)
2829+
variables (𝕜 : Type*) [fact (1 ≤ p)]
28302830

28312831
/-- The bounded linear map of considering a bounded continuous function on a finite-measure space
28322832
as an element of `Lp`. -/
2833-
def to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] :
2833+
def to_Lp [normed_field 𝕜] [normed_space 𝕜 E] :
28342834
(α →ᵇ E) →L[𝕜] (Lp E p μ) :=
28352835
linear_map.mk_continuous
28362836
(linear_map.cod_restrict
@@ -2840,23 +2840,34 @@ linear_map.mk_continuous
28402840
_
28412841
Lp_norm_le
28422842

2843+
lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] (f : α →ᵇ E) :
2844+
to_Lp p μ 𝕜 f =ᵐ[μ] f := ae_eq_fun.coe_fn_mk f _
2845+
28432846
variables {𝕜}
28442847

2845-
lemma range_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] :
2848+
lemma range_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] :
28462849
((linear_map.range (to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] Lp E p μ)).to_add_subgroup)
28472850
= measure_theory.Lp.bounded_continuous_function E p μ :=
28482851
range_to_Lp_hom p μ
28492852

28502853
variables {p}
28512854

2852-
lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] (f : α →ᵇ E) :
2853-
to_Lp p μ 𝕜 f =ᵐ[μ] f :=
2854-
ae_eq_fun.coe_fn_mk f _
2855-
2856-
lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] :
2855+
lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E]:
28572856
‖(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ :=
28582857
linear_map.mk_continuous_norm_le _ ((measure_univ_nnreal μ) ^ (p.to_real)⁻¹).coe_nonneg _
28592858

2859+
lemma to_Lp_inj {f g : α →ᵇ E} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] :
2860+
to_Lp p μ 𝕜 f = to_Lp p μ 𝕜 g ↔ f = g :=
2861+
begin
2862+
refine ⟨λ h, _, by tauto⟩,
2863+
rw [←fun_like.coe_fn_eq, ←(map_continuous f).ae_eq_iff_eq μ (map_continuous g)],
2864+
refine (coe_fn_to_Lp p μ 𝕜 f).symm.trans (eventually_eq.trans _ $ coe_fn_to_Lp p μ 𝕜 g),
2865+
rw h,
2866+
end
2867+
2868+
lemma to_Lp_injective [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] :
2869+
function.injective ⇑(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ)) := λ f g hfg, (to_Lp_inj μ).mp hfg
2870+
28602871
end bounded_continuous_function
28612872

28622873
namespace continuous_map
@@ -2906,7 +2917,28 @@ rfl
29062917
(to_Lp p μ 𝕜 f : α →ₘ[μ] E) = f.to_ae_eq_fun μ :=
29072918
rfl
29082919

2909-
variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 E]
2920+
lemma to_Lp_injective [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] :
2921+
function.injective ⇑(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ)) :=
2922+
(bounded_continuous_function.to_Lp_injective _).comp
2923+
(linear_isometry_bounded_of_compact α E 𝕜).injective
2924+
2925+
lemma to_Lp_inj {f g : C(α, E)} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] :
2926+
to_Lp p μ 𝕜 f = to_Lp p μ 𝕜 g ↔ f = g :=
2927+
(to_Lp_injective μ).eq_iff
2928+
2929+
variables {μ}
2930+
2931+
/-- If a sum of continuous functions `g n` is convergent, and the same sum converges in `Lᵖ` to `h`,
2932+
then in fact `g n` converges uniformly to `h`. -/
2933+
lemma has_sum_of_has_sum_Lp {β : Type*} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E]
2934+
{g : β → C(α, E)} {f : C(α, E)} (hg : summable g)
2935+
(hg2 : has_sum (to_Lp p μ 𝕜 ∘ g) (to_Lp p μ 𝕜 f)) : has_sum g f :=
2936+
begin
2937+
convert summable.has_sum hg,
2938+
exact to_Lp_injective μ (hg2.unique ((to_Lp p μ 𝕜).has_sum $ summable.has_sum hg)),
2939+
end
2940+
2941+
variables (μ) [nontrivially_normed_field 𝕜] [normed_space 𝕜 E]
29102942

29112943
lemma to_Lp_norm_eq_to_Lp_norm_coe :
29122944
‖(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))‖

0 commit comments

Comments
 (0)