Documentation

Mathlib.Analysis.Fourier.AddCircleMulti

Multivariate Fourier series #

In this file we define the Fourier series of an L² function on the d-dimensional unit circle, and show that it converges to the function in the L² norm. We also prove uniform convergence of the Fourier series if f is continuous and the sequence of its Fourier coefficients is summable.

In this file we normalise the measure on ℝ / ℤ to have total volume 1.

Equations
Instances For
    theorem instIsAddHaarMeasureUnitAddCircleVolume :
    MeasureTheory.volume.IsAddHaarMeasure

    The measure on ℝ / ℤ is a Haar measure.

    The measure on ℝ / ℤ is a probability measure.

    @[reducible, inline]
    abbrev UnitAddTorus (d : Type u_1) :
    Type u_1

    The product of finitely many copies of the unit circle, indexed by d.

    Equations
    Instances For
      def UnitAddTorus.mFourier {d : Type u_1} [Fintype d] (n : d) :

      Exponential monomials in d variables.

      Equations
      Instances For
        theorem UnitAddTorus.mFourier_add {d : Type u_1} [Fintype d] {n : d} {x : UnitAddTorus d} {m : d} :
        theorem UnitAddTorus.mFourier_single {d : Type u_1} [Fintype d] [DecidableEq d] (z : dAddCircle 1) (i : d) :

        The star subalgebra of C(UnitAddTorus d, ℂ) generated by mFourier n for n ∈ ℤᵈ.

        Equations
        Instances For
          theorem UnitAddTorus.mFourierSubalgebra_coe {d : Type u_1} [Fintype d] :
          Subalgebra.toSubmodule (UnitAddTorus.mFourierSubalgebra d).toSubalgebra = Submodule.span (Set.range UnitAddTorus.mFourier)

          The star subalgebra of C(UnitAddTorus d, ℂ) generated by mFourier n for n ∈ ℤᵈ is in fact the linear span of these functions.

          The subalgebra of C(UnitAddTorus d, ℂ) generated by mFourier n for n ∈ ℤᵈ separates points.

          The subalgebra of C(UnitAddTorus d, ℂ) generated by mFourier n for n : d → ℤ is dense.

          theorem UnitAddTorus.span_mFourier_closure_eq_top {d : Type u_1} [Fintype d] :
          (Submodule.span (Set.range UnitAddTorus.mFourier)).topologicalClosure =

          The linear span of the monomials mFourier n is dense in C(UnitAddTorus d, ℂ).

          @[reducible, inline]
          abbrev UnitAddTorus.mFourierLp {d : Type u_1} [Fintype d] (p : ENNReal) [Fact (1 p)] (n : d) :
          (MeasureTheory.Lp p MeasureTheory.volume)

          The family of monomials mFourier n, parametrized by n : ℤᵈ and considered as elements of the Lp space of functions UnitAddTorus d → ℂ.

          Equations
          Instances For
            theorem UnitAddTorus.coeFn_mFourierLp {d : Type u_1} [Fintype d] (p : ENNReal) [Fact (1 p)] (n : d) :
            (UnitAddTorus.mFourierLp p n) =ᵐ[MeasureTheory.volume] (UnitAddTorus.mFourier n)
            theorem UnitAddTorus.span_mFourierLp_closure_eq_top {d : Type u_1} [Fintype d] {p : ENNReal} [Fact (1 p)] (hp : p ) :

            For each 1 ≤ p < ∞, the linear span of the monomials mFourier n is dense in the Lᵖ space of functions on UnitAddTorus d.

            The monomials mFourierLp 2 n are an orthonormal set in .

            def UnitAddTorus.mFourierCoeff {d : Type u_1} [Fintype d] {E : Type} [NormedAddCommGroup E] [NormedSpace E] (f : UnitAddTorus dE) (n : d) :
            E

            The n-th Fourier coefficient of a function UnitAddTorus d → E, for E a complete normed -vector space, defined as the integral over UnitAddTorus d of mFourier (-n) t • f t.

            Equations
            Instances For
              def UnitAddTorus.mFourierBasis {d : Type u_1} [Fintype d] :
              HilbertBasis (d) (MeasureTheory.Lp 2 MeasureTheory.volume)

              We define mFourierBasis to be a ℤᵈ-indexed Hilbert basis for the space of functions on UnitAddTorus d, which by definition is an isometric isomorphism from L²(UnitAddTorus d) to ℓ²(ℤᵈ, ℂ).

              Equations
              Instances For
                @[simp]
                theorem UnitAddTorus.coe_mFourierBasis {d : Type u_1} [Fintype d] :
                (fun (i : d) => UnitAddTorus.mFourierBasis.repr.symm (lp.single 2 i 1)) = UnitAddTorus.mFourierLp 2

                The elements of the Hilbert basis mFourierBasis are the functions mFourierLp 2, i.e. the monomials mFourier n on UnitAddTorus d considered as elements of .

                theorem UnitAddTorus.mFourierBasis_repr {d : Type u_1} [Fintype d] (f : (MeasureTheory.Lp 2 MeasureTheory.volume)) (i : d) :
                (UnitAddTorus.mFourierBasis.repr f) i = UnitAddTorus.mFourierCoeff (↑f) i

                Under the isometric isomorphism mFourierBasis from L²(UnitAddTorus d) to ℓ²(ℤᵈ, ℂ), the i-th coefficient is mFourierCoeff f i.

                theorem UnitAddTorus.hasSum_mFourier_series_L2 {d : Type u_1} [Fintype d] (f : (MeasureTheory.Lp 2 MeasureTheory.volume)) :
                HasSum (fun (i : d) => UnitAddTorus.mFourierCoeff (↑f) i UnitAddTorus.mFourierLp 2 i) f

                The Fourier series of an L2 function f sums to f in the norm.

                theorem UnitAddTorus.hasSum_prod_mFourierCoeff {d : Type u_1} [Fintype d] (f g : (MeasureTheory.Lp 2 MeasureTheory.volume)) :
                HasSum (fun (i : d) => (starRingEnd ) (UnitAddTorus.mFourierCoeff (↑f) i) * UnitAddTorus.mFourierCoeff (↑g) i) (∫ (t : UnitAddTorus d), (starRingEnd ) (f t) * g t)

                Parseval's identity for inner products: for functions f, g on UnitAddTorus d, the inner product of the Fourier coefficients of f and g is the inner product of f and g.

                theorem UnitAddTorus.hasSum_sq_mFourierCoeff {d : Type u_1} [Fintype d] (f : (MeasureTheory.Lp 2 MeasureTheory.volume)) :
                HasSum (fun (i : d) => UnitAddTorus.mFourierCoeff (↑f) i ^ 2) (∫ (t : UnitAddTorus d), f t ^ 2)

                Parseval's identity for norms: for an function f on UnitAddTorus d, the sum of the squared norms of the Fourier coefficients equals the norm of f.

                theorem UnitAddTorus.mFourierCoeff_toLp {d : Type u_1} [Fintype d] (f : C(UnitAddTorus d, )) (n : d) :
                UnitAddTorus.mFourierCoeff (↑((ContinuousMap.toLp 2 MeasureTheory.volume ) f)) n = UnitAddTorus.mFourierCoeff (⇑f) n

                If the sequence of Fourier coefficients of f is summable, then the Fourier series converges uniformly to f.

                If the sequence of Fourier coefficients of f is summable, then the Fourier series of f converges everywhere pointwise to f.