The product measure #
In this file we define and prove properties about the binary product measure. If α and β have
s-finite measures μ resp. ν then α × β can be equipped with a s-finite measure μ.prod ν that
satisfies (μ.prod ν) s = ∫⁻ x, ν {y | (x, y) ∈ s} ∂μ.
We also have (μ.prod ν) (s ×ˢ t) = μ s * ν t, i.e. the measure of a rectangle is the product of
the measures of the sides.
We also prove Tonelli's theorem.
Main definition #
- MeasureTheory.Measure.prod: The product of two measures.
Main results #
- MeasureTheory.Measure.prod_applystates- μ.prod ν s = ∫⁻ x, ν {y | (x, y) ∈ s} ∂μfor measurable- s.- MeasureTheory.Measure.prod_apply_symmis the reversed version.
- MeasureTheory.Measure.prod_prodstates- μ.prod ν (s ×ˢ t) = μ s * ν tfor measurable sets- sand- t.
- MeasureTheory.lintegral_prod: Tonelli's theorem. It states that for a measurable function- α × β → ℝ≥0∞we have- ∫⁻ z, f z ∂(μ.prod ν) = ∫⁻ x, ∫⁻ y, f (x, y) ∂ν ∂μ. The version for functions- α → β → ℝ≥0∞is reversed, and called- lintegral_lintegral. Both versions have a variant with- _symmappended, where the order of integration is reversed. The lemma- Measurable.lintegral_prod_right'states that the inner integral of the right-hand side is measurable.
Implementation Notes #
Many results are proven twice, once for functions in curried form (α → β → γ) and one for
functions in uncurried form (α × β → γ). The former often has an assumption
Measurable (uncurry f), which could be inconvenient to discharge, but for the latter it is more
common that the function has to be given explicitly, since Lean cannot synthesize the function by
itself. We name the lemmas about the uncurried form with a prime.
Tonelli's theorem has a different naming scheme, since the version for the uncurried version is
reversed.
Tags #
product measure, Tonelli's theorem, Fubini-Tonelli theorem
Rectangles formed by π-systems form a π-system.
Rectangles of countably spanning sets are countably spanning.
Measurability #
Before we define the product measure, we can talk about the measurability of operations on binary
functions. We show that if f is a binary measurable function, then the function that integrates
along one of the variables (using either the Lebesgue or Bochner integral) is measurable.
The product of generated σ-algebras is the one generated by rectangles, if both generating sets are countably spanning.
If C and D generate the σ-algebras on α resp. β, then rectangles formed by C and D
generate the σ-algebra on α × β.
The product σ-algebra is generated from boxes, i.e. s ×ˢ t for sets s : Set α and
t : Set β.
Rectangles form a π-system.
If ν is a finite measure, and s ⊆ α × β is measurable, then x ↦ ν { y | (x, y) ∈ s } is
a measurable function. measurable_measure_prod_mk_left is strictly more general.
If ν is an s-finite measure, and s ⊆ α × β is measurable, then x ↦ ν { y | (x, y) ∈ s }
is a measurable function.
If μ is a σ-finite measure, and s ⊆ α × β is measurable, then y ↦ μ { x | (x, y) ∈ s } is
a measurable function.
The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable.
The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of)
Tonelli's theorem is measurable.
This version has the argument f in curried form.
The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) the symmetric version of Tonelli's theorem is measurable.
The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of)
the symmetric version of Tonelli's theorem is measurable.
This version has the argument f in curried form.
The product measure #
The binary product of measures. They are defined for arbitrary measures, but we basically prove all properties under the assumption that at least one of them is s-finite.
Instances For
Equations
- MeasureTheory.Measure.prod.measureSpace = MeasureTheory.MeasureSpace.mk (MeasureTheory.volume.prod MeasureTheory.volume)
The product measure of the product of two sets is the product of their measures. Note that we do not need the sets to be measurable.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Note: the assumption hs cannot be dropped. For a counterexample, see
Walter Rudin Real and Complex Analysis, example (c) in section 8.9.
Note: the converse is not true without assuming that s is measurable. For a counterexample,
see Walter Rudin Real and Complex Analysis, example (c) in section 8.9.
Note: the converse is not true. For a counterexample, see
Walter Rudin Real and Complex Analysis, example (c) in section 8.9. It is true if the set is
measurable, see ae_prod_mem_iff_ae_ae_mem.
If s ×ˢ t is a null measurable set and μ s ≠ 0, then t is a null measurable set.
If Prod.snd ⁻¹' t is a null measurable set and μ ≠ 0, then t is a null measurable set.
Prod.snd ⁻¹' t is null measurable w.r.t. μ.prod ν iff t is null measurable w.r.t. ν
provided that μ ≠ 0.
μ.prod ν has finite spanning sets in rectangles of finite spanning sets.
Equations
- hμ.prod hν = { set := fun (n : ℕ) => hμ.set (Nat.unpair n).1 ×ˢ hν.set (Nat.unpair n).2, set_mem := ⋯, finite := ⋯, spanning := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A measure on a product space equals the product measure if they are equal on rectangles with as sides sets that generate the corresponding σ-algebras.
A measure on a product space equals the product measure of sigma-finite measures if they are equal on rectangles.
If s ×ˢ t is a null measurable set and ν t ≠ 0, then s is a null measurable set.
If Prod.fst ⁻¹' s is a null measurable set and ν ≠ 0, then s is a null measurable set.
Prod.fst ⁻¹' s is null measurable w.r.t. μ.prod ν iff s is null measurable w.r.t. μ
provided that ν ≠ 0.
The product of two non-null sets is null measurable if and only if both of them are null measurable.
The product of two sets is null measurable if and only if both of them are null measurable or one of them has measure zero.
The product of specific measures #
If f : α → β sends the measure μa to μb and g : γ → δ sends the measure μc to μd,
then Prod.map f g sends μa.prod μc to μb.prod μd.
The Lebesgue integral on a product #
Tonelli's Theorem: For ℝ≥0∞-valued measurable functions on α × β,
the integral of f is equal to the iterated integral.
Tonelli's Theorem: For ℝ≥0∞-valued almost everywhere measurable functions on α × β,
the integral of f is equal to the iterated integral.
The symmetric version of Tonelli's Theorem: For ℝ≥0∞-valued almost everywhere measurable
functions on α × β, the integral of f is equal to the iterated integral, in reverse order.
The symmetric version of Tonelli's Theorem: For ℝ≥0∞-valued measurable
functions on α × β, the integral of f is equal to the iterated integral, in reverse order.
The reversed version of Tonelli's Theorem. In this version f is in curried form, which
makes it easier for the elaborator to figure out f automatically.
The reversed version of Tonelli's Theorem (symmetric version). In this version f is in
curried form, which makes it easier for the elaborator to figure out f automatically.
Change the order of Lebesgue integration.
Marginals of a measure defined on a product #
Marginal measure on α obtained from a measure ρ on α × β, defined by ρ.map Prod.fst.
Equations
- ρ.fst = MeasureTheory.Measure.map Prod.fst ρ
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Marginal measure on β obtained from a measure on ρ α × β, defined by ρ.map Prod.snd.
Equations
- ρ.snd = MeasureTheory.Measure.map Prod.snd ρ
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯