Integrals of periodic functions #
In this file we prove that the half-open interval Ioc t (t + T) in ℝ is a fundamental domain of
the action of the subgroup ℤ ∙ T on ℝ.
A consequence is AddCircle.measurePreserving_mk: the covering map from ℝ to the "additive
circle" ℝ ⧸ (ℤ ∙ T) is measure-preserving, with respect to the restriction of Lebesgue measure to
Ioc t (t + T) (upstairs) and with respect to Haar measure (downstairs).
Another consequence (Function.Periodic.intervalIntegral_add_eq and related declarations) is that
∫ x in t..t + T, f x = ∫ x in s..s + T, f x for any (not necessarily measurable) function with
period T.
Measures and integrability on ℝ and on the circle #
Equip the "additive circle" ℝ ⧸ (ℤ ∙ T) with, as a standard measure, the Haar measure of total
mass T
Equations
- AddCircle.measureSpace T = { toMeasurableSpace := QuotientAddGroup.measurableSpace (AddSubgroup.zmultiples T), volume := ENNReal.ofReal T • MeasureTheory.Measure.addHaarMeasure ⊤ }
The covering map from ℝ to the "additive circle" ℝ ⧸ (ℤ ∙ T) is measure-preserving,
considered with respect to the standard measure (defined to be the Haar measure of total mass T)
on the additive circle, and with respect to the restriction of Lebesgue measure on ℝ to an
interval (t, t + T].
The isomorphism AddCircle T ≃ Ioc a (a + T) whose inverse is the natural quotient map,
as an equivalence of measurable spaces.
Equations
- AddCircle.measurableEquivIoc T a = { toEquiv := AddCircle.equivIoc T a, measurable_toFun := ⋯, measurable_invFun := ⋯ }
Instances For
The isomorphism AddCircle T ≃ Ico a (a + T) whose inverse is the natural quotient map,
as an equivalence of measurable spaces.
Equations
- AddCircle.measurableEquivIco T a = { toEquiv := AddCircle.equivIco T a, measurable_toFun := ⋯, measurable_invFun := ⋯ }
Instances For
The equivalence equivIoc is measure preserving with respect to the natural volume measures.
The integral of an almost-everywhere strongly measurable function over AddCircle T is equal
to the integral over an interval (t, t + T] in ℝ of its lift to ℝ.
The integral of an almost-everywhere strongly measurable function over AddCircle T is equal
to the integral over an interval (t, t + T] in ℝ of its lift to ℝ.
The integral of a function lifted to AddCircle from an interval (t, t + T] to AddCircle T
is equal the the intervalIntegral over the interval.
If a function satisfies MemLp on the interval (t, t + T], then its lift to the AddCircle
also satisfies MemLp with respect to the Haar measure.
The covering map from ℝ to the "unit additive circle" ℝ ⧸ ℤ is measure-preserving,
considered with respect to the standard measure (defined to be the Haar measure of total mass 1)
on the additive circle, and with respect to the restriction of Lebesgue measure on ℝ to an
interval (t, t + 1].
The integral of a measurable function over UnitAddCircle is equal to the integral over an
interval (t, t + 1] in ℝ of its lift to ℝ.
The integral of an almost-everywhere strongly measurable function over UnitAddCircle is
equal to the integral over an interval (t, t + 1] in ℝ of its lift to ℝ.
The integral of an almost-everywhere strongly measurable function over UnitAddCircle is
equal to the integral over an interval (t, t + 1] in ℝ of its lift to ℝ.
Interval integrability of periodic functions #
A periodic function is interval integrable over every interval if it is interval integrable over one period.
Special case of Function.Periodic.intervalIntegrable: A periodic function is interval integrable over every interval if it is interval integrable over the period starting from zero.
Interval integrals of periodic functions #
An auxiliary lemma for a more general Function.Periodic.intervalIntegral_add_eq.
If f is a periodic function with period T, then its integral over [t, t + T] does not
depend on t.
If f is an integrable periodic function with period T, then its integral over [t, s + T]
is the sum of its integrals over the intervals [t, s] and [t, t + T].
If f is an integrable periodic function with period T, and n is an integer, then its
integral over [t, t + n • T] is n times its integral over [t, t + T].
If g : ℝ → ℝ is periodic with period T > 0, then for any t : ℝ, the function
t ↦ ∫ x in 0..t, g x is bounded below by t ↦ X + ⌊t/T⌋ • Y for appropriate constants X and
Y.
If g : ℝ → ℝ is periodic with period T > 0, then for any t : ℝ, the function
t ↦ ∫ x in 0..t, g x is bounded above by t ↦ X + ⌊t/T⌋ • Y for appropriate constants X and
Y.
If g : ℝ → ℝ is periodic with period T > 0 and 0 < ∫ x in 0..T, g x, then
t ↦ ∫ x in 0..t, g x tends to ∞ as t tends to ∞.
If g : ℝ → ℝ is periodic with period T > 0 and 0 < ∫ x in 0..T, g x, then
t ↦ ∫ x in 0..t, g x tends to -∞ as t tends to -∞.
If g : ℝ → ℝ is periodic with period T > 0 and ∀ x, 0 < g x, then t ↦ ∫ x in 0..t, g x
tends to ∞ as t tends to ∞.
If g : ℝ → ℝ is periodic with period T > 0 and ∀ x, 0 < g x, then t ↦ ∫ x in 0..t, g x
tends to -∞ as t tends to -∞.