Fourier analysis on the additive circle #
This file contains basic results on Fourier series for functions on the additive circle
AddCircle T = ℝ / ℤ • T
.
Main definitions #
haarAddCircle
, Haar measure onAddCircle T
, normalized to have total measure1
. (Note that this is not the same normalisation as the standard measure defined inIntegral.Periodic
, so we do not declare it as aMeasureSpace
instance, to avoid confusion.)- for
n : ℤ
,fourier n
is the monomialfun x => exp (2 π i n x / T)
, bundled as a continuous map fromAddCircle T
toℂ
. fourierBasis
is the Hilbert basis ofLp ℂ 2 haarAddCircle
given by the images of the monomialsfourier n
.fourierCoeff f n
, forf : AddCircle T → E
(withE
a complete normedℂ
-vector space), is then
-th Fourier coefficient off
, defined as an integral overAddCircle T
. The lemmafourierCoeff_eq_intervalIntegral
expresses this as an integral over[a, a + T]
for any reala
.fourierCoeffOn
, forf : ℝ → E
anda < b
reals, is then
-th Fourier coefficient of the unique periodic function of periodb - a
which agrees withf
on(a, b]
. The lemmafourierCoeffOn_eq_integral
expresses this as an integral over[a, b]
.
Main statements #
The theorem span_fourier_closure_eq_top
states that the span of the monomials fourier n
is
dense in C(AddCircle T, ℂ)
, i.e. that its Submodule.topologicalClosure
is ⊤
. This follows
from the Stone-Weierstrass theorem after checking that the span is a subalgebra, is closed under
conjugation, and separates points.
Using this and general theory on approximation of Lᵖ functions by continuous functions, we deduce
(span_fourierLp_closure_eq_top
) that for any 1 ≤ p < ∞
, the span of the Fourier monomials is
dense in the Lᵖ space of AddCircle T
. For p = 2
we show (orthonormal_fourier
) that the
monomials are also orthonormal, so they form a Hilbert basis for L², which is named as
fourierBasis
; in particular, for L²
functions f
, the Fourier series of f
converges to f
in the L²
topology (hasSum_fourier_series_L2
). Parseval's identity, tsum_sq_fourierCoeff
, is
a direct consequence.
For continuous maps f : AddCircle T → ℂ
, the theorem
hasSum_fourier_series_of_summable
states that if the sequence of Fourier
coefficients of f
is summable, then the Fourier series ∑ (i : ℤ), fourierCoeff f i * fourier i
converges to f
in the uniform-convergence topology of C(AddCircle T, ℂ)
.
Measure on AddCircle T
#
In this file we use the Haar measure on AddCircle T
normalised to have total measure 1 (which is
not the same as the standard measure defined in Topology.Instances.AddCircle
).
Haar measure on the additive circle, normalised to have total measure 1.
Instances For
The family of exponential monomials fun x => exp (2 π i n x / T)
, parametrized by n : ℤ
and
considered as bundled continuous maps from ℝ / ℤ • T
to ℂ
.
Instances For
The star subalgebra of C(AddCircle T, ℂ)
generated by fourier n
for n ∈ ℤ
.
Equations
- fourierSubalgebra = { toSubalgebra := Algebra.adjoin ℂ (Set.range fourier), star_mem' := ⋯ }
Instances For
The star subalgebra of C(AddCircle T, ℂ)
generated by fourier n
for n ∈ ℤ
is in fact the
linear span of these functions.
The subalgebra of C(AddCircle T, ℂ)
generated by fourier n
for n ∈ ℤ
separates points.
The subalgebra of C(AddCircle T, ℂ)
generated by fourier n
for n ∈ ℤ
is dense.
The family of monomials fourier n
, parametrized by n : ℤ
and considered as
elements of the Lp
space of functions AddCircle T → ℂ
.
Equations
- fourierLp p n = (ContinuousMap.toLp p AddCircle.haarAddCircle ℂ) (fourier n)
Instances For
The monomials fourier n
are an orthonormal set with respect to normalised Haar measure.
The n
-th Fourier coefficient of a function AddCircle T → E
, for E
a complete normed
ℂ
-vector space, defined as the integral over AddCircle T
of fourier (-n) t • f t
.
Equations
- fourierCoeff f n = ∫ (t : AddCircle T), (fourier (-n)) t • f t ∂AddCircle.haarAddCircle
Instances For
The Fourier coefficients of a function on AddCircle T
can be computed as an integral
over [a, a + T]
, for any real a
.
For a function on ℝ
, the Fourier coefficients of f
on [a, b]
are defined as the
Fourier coefficients of the unique periodic function agreeing with f
on Ioc a b
.
Equations
- fourierCoeffOn hab f n = fourierCoeff (AddCircle.liftIoc (b - a) a f) n
Instances For
We define fourierBasis
to be a ℤ
-indexed Hilbert basis for Lp ℂ 2 haarAddCircle
,
which by definition is an isometric isomorphism from Lp ℂ 2 haarAddCircle
to ℓ²(ℤ, ℂ)
.
Equations
Instances For
The elements of the Hilbert basis fourierBasis
are the functions fourierLp 2
, i.e. the
monomials fourier n
on the circle considered as elements of L²
.
Under the isometric isomorphism fourierBasis
from Lp ℂ 2 haarAddCircle
to ℓ²(ℤ, ℂ)
, the
i
-th coefficient is fourierCoeff f i
, i.e., the integral over AddCircle T
of
fun t => fourier (-i) t * f t
with respect to the Haar measure of total mass 1.
The Fourier series of an L2
function f
sums to f
, in the L²
space of AddCircle T
.
Parseval's identity: for an L²
function f
on AddCircle T
, the sum of the squared
norms of the Fourier coefficients equals the L²
norm of f
.
If the sequence of Fourier coefficients of f
is summable, then the Fourier series of f
converges everywhere pointwise to f
.
Express Fourier coefficients of f
on an interval in terms of those of its derivative.
Express Fourier coefficients of f
on an interval in terms of those of its derivative.
Express Fourier coefficients of f
on an interval in terms of those of its derivative.