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
- Function.Periodic.qParam h z = Complex.exp (2 * ↑Real.pi * Complex.I * z / ↑h)
Instances For
One-sided inverse of qParam h
.
Equations
- Function.Periodic.invQParam h q = ↑h / (2 * ↑Real.pi * Complex.I) * Complex.log q
Instances For
The function q ↦ f (invQParam h q)
, extended by a non-canonical choice of limit at 0.
Equations
- Function.Periodic.cuspFunction h f = Function.update (f ∘ Function.Periodic.invQParam h) 0 (limUnder (nhdsWithin 0 {0}ᶜ) (f ∘ Function.Periodic.invQParam h))
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).
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.
If f
is periodic, holomorphic near I∞
, and tends to zero at I∞
, then in fact it tends to zero
exponentially fast.