Integral over an interval #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
In this file we define ∫ x in a..b, f x ∂μ to be ∫ x in Ioc a b, f x ∂μ if a ≤ b and
-∫ x in Ioc b a, f x ∂μ if b ≤ a.
Implementation notes #
Avoiding if, min, and max #
In order to avoid ifs in the definition, we define interval_integrable f μ a b as
integrable_on f (Ioc a b) μ ∧ integrable_on f (Ioc b a) μ. For any a, b one of these
intervals is empty and the other coincides with set.uIoc a b = set.Ioc (min a b) (max a b).
Similarly, we define ∫ x in a..b, f x ∂μ to be ∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ.
Again, for any a, b one of these integrals is zero, and the other gives the expected result.
This way some properties can be translated from integrals over sets without dealing with
the cases a ≤ b and b ≤ a separately.
Choice of the interval #
We use integral over set.uIoc a b = set.Ioc (min a b) (max a b) instead of one of the other
three possible intervals with the same endpoints for two reasons:
- this way ∫ x in a..b, f x ∂μ + ∫ x in b..c, f x ∂μ = ∫ x in a..c, f x ∂μholds wheneverfis integrable on each interval; in particular, it works even if the measureμhas an atom atb; this rules outset.Iooandset.Iccintervals;
- with this definition for a probability measure μ, the integral∫ x in a..b, 1 ∂μequals the difference $F_μ(b)-F_μ(a)$, where $F_μ(a)=μ(-∞, a]$ is the cumulative distribution function ofμ.
Tags #
integral
Integrability on an interval #
A function f is called interval integrable with respect to a measure μ on an unordered
interval a..b if it is integrable on both intervals (a, b] and (b, a]. One of these
intervals is always empty, so this property is equivalent to f being integrable on
(min a b, max a b].
Equations
- interval_integrable f μ a b = (measure_theory.integrable_on f (set.Ioc a b) μ ∧ measure_theory.integrable_on f (set.Ioc b a) μ)
A function is interval integrable with respect to a given measure μ on a..b if and
only if it is integrable on uIoc a b with respect to μ. This is an equivalent
definition of interval_integrable.
If a function is interval integrable with respect to a given measure μ on a..b then
it is integrable on uIoc a b with respect to μ.
If a function is integrable with respect to a given measure μ then it is interval integrable
with respect to μ on uIcc a b.
A continuous function on ℝ is interval_integrable with respect to any locally finite measure
ν on ℝ.
Let l' be a measurably generated filter; let l be a of filter such that each s ∈ l'
eventually includes Ioc u v as both u and v tend to l. Let μ be a measure finite at l'.
Suppose that f : ℝ → E has a finite limit at l' ⊓ μ.ae. Then f is interval integrable on
u..v provided that both u and v tend to l.
Typeclass instances allow Lean to find l' based on l but not vice versa, so
apply tendsto.eventually_interval_integrable_ae will generate goals filter ℝ and
tendsto_Ixx_class Ioc ?m_1 l'.
Let l' be a measurably generated filter; let l be a of filter such that each s ∈ l'
eventually includes Ioc u v as both u and v tend to l. Let μ be a measure finite at l'.
Suppose that f : ℝ → E has a finite limit at l. Then f is interval integrable on u..v
provided that both u and v tend to l.
Typeclass instances allow Lean to find l' based on l but not vice versa, so
apply tendsto.eventually_interval_integrable_ae will generate goals filter ℝ and
tendsto_Ixx_class Ioc ?m_1 l'.
Interval integral: definition and basic properties #
In this section we define ∫ x in a..b, f x ∂μ as ∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ
and prove some basic properties.
The interval integral ∫ x in a..b, f x ∂μ is defined
as ∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ. If a ≤ b, then it equals
∫ x in Ioc a b, f x ∂μ, otherwise it equals -∫ x in Ioc b a, f x ∂μ.
Integral is an additive function of the interval #
In this section we prove that ∫ x in a..b, f x ∂μ + ∫ x in b..c, f x ∂μ = ∫ x in a..c, f x ∂μ
as well as a few other identities trivially equivalent to this one. We also prove that
∫ x in a..b, f x ∂μ = ∫ x, f x ∂μ provided that support f ⊆ Ioc a b.
If two functions are equal in the relevant interval, their interval integrals are also equal.
If μ is a finite measure then ∫ x in a..b, c ∂μ = (μ (Iic b) - μ (Iic a)) • c.
Lebesgue dominated convergence theorem for filters with a countable basis
Lebesgue dominated convergence theorem for series.
Interval integrals commute with countable sums, when the supremum norms are summable (a special case of the dominated convergence theorem).
Continuity of interval integral with respect to a parameter, at a point within a set.
Given F : X → ℝ → E, assume F x is ae-measurable on [a, b] for x in a
neighborhood of x₀ within s and at x₀, and assume it is bounded by a function integrable
on [a, b] independent of x in a neighborhood of x₀ within s. If (λ x, F x t)
is continuous at x₀ within s for almost every t in [a, b]
then the same holds for (λ x, ∫ t in a..b, F x t ∂μ) s x₀.
Continuity of interval integral with respect to a parameter at a point.
Given F : X → ℝ → E, assume F x is ae-measurable on [a, b] for x in a
neighborhood of x₀, and assume it is bounded by a function integrable on
[a, b] independent of x in a neighborhood of x₀. If (λ x, F x t)
is continuous at x₀ for almost every t in [a, b]
then the same holds for (λ x, ∫ t in a..b, F x t ∂μ) s x₀.
Continuity of interval integral with respect to a parameter.
Given F : X → ℝ → E, assume each F x is ae-measurable on [a, b],
and assume it is bounded by a function integrable on [a, b] independent of x.
If (λ x, F x t) is continuous for almost every t in [a, b]
then the same holds for (λ x, ∫ t in a..b, F x t ∂μ) s x₀.
Note: this assumes that f is interval_integrable, in contrast to some other lemmas here.
If f is nonnegative and integrable on the unordered interval set.uIoc a b, then its
integral over a..b is positive if and only if a < b and the measure of
function.support f ∩ set.Ioc a b is positive.
If f is nonnegative a.e.-everywhere and it is integrable on the unordered interval
set.uIoc a b, then its integral over a..b is positive if and only if a < b and the
measure of function.support f ∩ set.Ioc a b is positive.
If f : ℝ → ℝ is integrable on (a, b] for real numbers a < b, and positive on the interior
of the interval, then its integral over a..b is strictly positive.
If f : ℝ → ℝ is strictly positive everywhere, and integrable on (a, b] for real numbers
a < b, then its integral over a..b is strictly positive. (See interval_integral_pos_of_pos_on
for a version only assuming positivity of f on (a, b) rather than everywhere.)
If f and g are two functions that are interval integrable on a..b, a ≤ b,
f x ≤ g x for a.e. x ∈ set.Ioc a b, and f x < g x on a subset of set.Ioc a b
of nonzero measure, then ∫ x in a..b, f x ∂μ < ∫ x in a..b, g x ∂μ.
If f and g are continuous on [a, b], a < b, f x ≤ g x on this interval, and
f c < g c at some point c ∈ [a, b], then ∫ x in a..b, f x < ∫ x in a..b, g x.