Documentation

Mathlib.Analysis.Complex.Periodic

Periodic holomorphic functions #

We show that if f : ℂ → ℂ satisfies f (z + h) = f z, for some nonzero real h, then there is a function F such that f z = F (exp (2 * π * I * z / h)) for all z; and if f is holomorphic at some z, then F is holomorphic at exp (2 * π * I * z / h).

We also show (using Riemann's removable singularity theorem) that if f is holomorphic and bounded for all sufficiently large im z, then F extends to a holomorphic function on a neighbourhood of 0. As a consequence, if f tends to zero as im z → ∞, then in fact it decays exponentially to zero. These results are important in the theory of modular forms.

Parameter for q-expansions, qParam h z = exp (2 * π * I * z / h)

Equations
Instances For

    One-sided inverse of qParam h.

    Equations
    Instances For
      theorem Function.Periodic.abs_qParam (h : ) (z : ) :
      Complex.abs (Function.Periodic.qParam h z) = Real.exp (-2 * Real.pi * z.im / h)
      theorem Function.Periodic.im_invQParam (h : ) (q : ) :
      (Function.Periodic.invQParam h q).im = -h / (2 * Real.pi) * Real.log (Complex.abs q)
      theorem Function.Periodic.abs_qParam_lt_iff {h : } (hh : 0 < h) (A : ) (z : ) :
      Complex.abs (Function.Periodic.qParam h z) < Real.exp (-2 * Real.pi * A / h) A < z.im
      def Function.Periodic.cuspFunction (h : ) (f : ) :

      The function q ↦ f (invQParam h q), extended by a non-canonical choice of limit at 0.

      Equations
      Instances For

        Key technical lemma: the function cuspFunction h f is differentiable at the images of differentiability points of f (even if invQParam is not differentiable there).

        theorem Function.Periodic.boundedAtFilter_cuspFunction {h : } {f : } (hh : 0 < h) (h_bd : (Filter.comap Complex.im Filter.atTop).BoundedAtFilter f) :
        (nhdsWithin 0 {0}).BoundedAtFilter (Function.Periodic.cuspFunction h f)
        theorem Function.Periodic.cuspFunction_zero_of_zero_at_inf {h : } {f : } (hh : 0 < h) (h_zer : (Filter.comap Complex.im Filter.atTop).ZeroAtFilter f) :
        theorem Function.Periodic.differentiableAt_cuspFunction_zero {h : } {f : } (hh : 0 < h) (hf : Function.Periodic f h) (h_hol : ∀ᶠ (z : ) in Filter.comap Complex.im Filter.atTop, DifferentiableAt f z) (h_bd : (Filter.comap Complex.im Filter.atTop).BoundedAtFilter f) :
        theorem Function.Periodic.tendsto_at_I_inf {h : } {f : } (hh : 0 < h) (hf : Function.Periodic f h) (h_hol : ∀ᶠ (z : ) in Filter.comap Complex.im Filter.atTop, DifferentiableAt f z) (h_bd : (Filter.comap Complex.im Filter.atTop).BoundedAtFilter f) :

        If f is periodic, and holomorphic and bounded near I∞, then it tends to a limit at I∞, and this limit is the value of its cusp function at 0.

        theorem Function.Periodic.exp_decay_of_zero_at_inf {h : } {f : } (hh : 0 < h) (hf : Function.Periodic f h) (h_hol : ∀ᶠ (z : ) in Filter.comap Complex.im Filter.atTop, DifferentiableAt f z) (h_zer : (Filter.comap Complex.im Filter.atTop).ZeroAtFilter f) :
        f =O[Filter.comap Complex.im Filter.atTop] fun (z : ) => Real.exp (-2 * Real.pi * z.im / h)

        If f is periodic, holomorphic near I∞, and tends to zero at I∞, then in fact it tends to zero exponentially fast.