f' is interval integrable for certain classes of functions f #
This file proves that:
MonotoneOn.intervalIntegrable_deriv: Iffis monotone ona..b, thenf'is interval integrable ona..b.MonotoneOn.intervalIntegral_deriv_mem_uIcc: Iffis monotone ona..b, then the integral off'ona..bis inuIcc 0 (f b - f a).BoundedVariationOn.intervalIntegrable_deriv: Iffhas bounded variation ona..b, thenf'is interval integrable ona..b.AbsolutelyContinuousOnInterval.intervalIntegrable_deriv: Iffis absolutely continuous ona..b, thenf'is interval integrable ona..b.
Tags #
interval integrable, monotone, bounded variation, absolutely continuous
theorem
MonotoneOn.exists_tendsto_deriv_liminf_lintegral_enorm_le
{f : ℝ → ℝ}
{a b : ℝ}
(hab : a ≤ b)
(hf : MonotoneOn f (Set.Icc a b))
:
∃ (G : ℕ → ℝ → ℝ),
(∀ᵐ (x : ℝ) ∂MeasureTheory.volume.restrict (Set.Icc a b), Filter.Tendsto (fun (n : ℕ) => G n x) Filter.atTop (nhds (deriv f x))) ∧ (∀ (n : ℕ), MeasureTheory.AEStronglyMeasurable (G n) (MeasureTheory.volume.restrict (Set.Icc a b))) ∧ Filter.liminf (fun (n : ℕ) => ∫⁻ (x : ℝ) in Set.Icc a b, ‖G n x‖ₑ) Filter.atTop ≤ ENNReal.ofReal (f b - f a)
If f is monotone on [a, b], then f' is the limit of G n a.e. on [a, b], where each
G n is AEStronglyMeasurable and the liminf of the lower Lebesgue integral of ‖G n ·‖ₑ is at
most f b - f a.
theorem
MonotoneOn.intervalIntegrable_deriv
{f : ℝ → ℝ}
{a b : ℝ}
(hf : MonotoneOn f (Set.uIcc a b))
:
If f is monotone on a..b, then f' is interval integrable on a..b.
theorem
BoundedVariationOn.intervalIntegrable_deriv
{f : ℝ → ℝ}
{a b : ℝ}
(hf : BoundedVariationOn f (Set.uIcc a b))
:
If f has bounded variation on uIcc a b, then f' is interval integrable on a..b.
theorem
AbsolutelyContinuousOnInterval.intervalIntegrable_deriv
{f : ℝ → ℝ}
{a b : ℝ}
(hf : AbsolutelyContinuousOnInterval f a b)
:
If f is absolutely continuous on uIcc a b, then f' is interval integrable on a..b.