Documentation

Mathlib.Probability.Integration

Integration in Probability Theory #

Integration results for independent random variables. Specifically, for two independent random variables X and Y over the extended non-negative reals, E[X * Y] = E[X] * E[Y], and similar results.

Implementation notes #

Many lemmas in this file take two arguments of the same typeclass. It is worth remembering that lean will always pick the later typeclass in this situation, and does not care whether the arguments are [], {}, or (). All of these use the MeasurableSpace M2 to define μ:

example {M1 : MeasurableSpace Ω} [M2 : MeasurableSpace Ω] {μ : Measure Ω} : sorry := sorry
example [M1 : MeasurableSpace Ω] {M2 : MeasurableSpace Ω} {μ : Measure Ω} : sorry := sorry
theorem ProbabilityTheory.lintegral_mul_indicator_eq_lintegral_mul_lintegral_indicator {Ω : Type u_1} {f : ΩENNReal} {Mf : MeasurableSpace Ω} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} (hMf : Mf ) (c : ENNReal) {T : Set Ω} (h_meas_T : MeasurableSet T) (h_ind : ProbabilityTheory.IndepSets {s : Set Ω | MeasurableSet s} {T} μ) (h_meas_f : Measurable f) :
∫⁻ (ω : Ω), f ω * T.indicator (fun (x : Ω) => c) ωμ = (∫⁻ (ω : Ω), f ωμ) * ∫⁻ (ω : Ω), T.indicator (fun (x : Ω) => c) ωμ

If a random variable f in ℝ≥0∞ is independent of an event T, then if you restrict the random variable to T, then E[f * indicator T c 0]=E[f] * E[indicator T c 0]. It is useful for lintegral_mul_eq_lintegral_mul_lintegral_of_independent_measurableSpace.

theorem ProbabilityTheory.lintegral_mul_eq_lintegral_mul_lintegral_of_independent_measurableSpace {Ω : Type u_1} {f : ΩENNReal} {g : ΩENNReal} {Mf : MeasurableSpace Ω} {Mg : MeasurableSpace Ω} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} (hMf : Mf ) (hMg : Mg ) (h_ind : ProbabilityTheory.Indep Mf Mg μ) (h_meas_f : Measurable f) (h_meas_g : Measurable g) :
∫⁻ (ω : Ω), f ω * g ωμ = (∫⁻ (ω : Ω), f ωμ) * ∫⁻ (ω : Ω), g ωμ

If f and g are independent random variables with values in ℝ≥0∞, then E[f * g] = E[f] * E[g]. However, instead of directly using the independence of the random variables, it uses the independence of measurable spaces for the domains of f and g. This is similar to the sigma-algebra approach to independence. See lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun for a more common variant of the product of independent variables.

theorem ProbabilityTheory.lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {f : ΩENNReal} {g : ΩENNReal} (h_meas_f : Measurable f) (h_meas_g : Measurable g) (h_indep_fun : ProbabilityTheory.IndepFun f g μ) :
∫⁻ (ω : Ω), (f * g) ωμ = (∫⁻ (ω : Ω), f ωμ) * ∫⁻ (ω : Ω), g ωμ

If f and g are independent random variables with values in ℝ≥0∞, then E[f * g] = E[f] * E[g].

theorem ProbabilityTheory.lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun' {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {f : ΩENNReal} {g : ΩENNReal} (h_meas_f : AEMeasurable f μ) (h_meas_g : AEMeasurable g μ) (h_indep_fun : ProbabilityTheory.IndepFun f g μ) :
∫⁻ (ω : Ω), (f * g) ωμ = (∫⁻ (ω : Ω), f ωμ) * ∫⁻ (ω : Ω), g ωμ

If f and g with values in ℝ≥0∞ are independent and almost everywhere measurable, then E[f * g] = E[f] * E[g] (slightly generalizing lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun).

theorem ProbabilityTheory.lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {f : ΩENNReal} {g : ΩENNReal} (h_meas_f : AEMeasurable f μ) (h_meas_g : AEMeasurable g μ) (h_indep_fun : ProbabilityTheory.IndepFun f g μ) :
∫⁻ (ω : Ω), f ω * g ωμ = (∫⁻ (ω : Ω), f ωμ) * ∫⁻ (ω : Ω), g ωμ
theorem ProbabilityTheory.IndepFun.integrable_mul {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {β : Type u_2} [MeasurableSpace β] {X : Ωβ} {Y : Ωβ} [NormedDivisionRing β] [BorelSpace β] (hXY : ProbabilityTheory.IndepFun X Y μ) (hX : MeasureTheory.Integrable X μ) (hY : MeasureTheory.Integrable Y μ) :

The product of two independent, integrable, real-valued random variables is integrable.

theorem ProbabilityTheory.IndepFun.integrable_left_of_integrable_mul {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {β : Type u_2} [MeasurableSpace β] {X : Ωβ} {Y : Ωβ} [NormedDivisionRing β] [BorelSpace β] (hXY : ProbabilityTheory.IndepFun X Y μ) (h'XY : MeasureTheory.Integrable (X * Y) μ) (hX : MeasureTheory.AEStronglyMeasurable X μ) (hY : MeasureTheory.AEStronglyMeasurable Y μ) (h'Y : ¬μ.ae.EventuallyEq Y 0) :

If the product of two independent real-valued random variables is integrable and the second one is not almost everywhere zero, then the first one is integrable.

theorem ProbabilityTheory.IndepFun.integrable_right_of_integrable_mul {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {β : Type u_2} [MeasurableSpace β] {X : Ωβ} {Y : Ωβ} [NormedDivisionRing β] [BorelSpace β] (hXY : ProbabilityTheory.IndepFun X Y μ) (h'XY : MeasureTheory.Integrable (X * Y) μ) (hX : MeasureTheory.AEStronglyMeasurable X μ) (hY : MeasureTheory.AEStronglyMeasurable Y μ) (h'X : ¬μ.ae.EventuallyEq X 0) :

If the product of two independent real-valued random variables is integrable and the first one is not almost everywhere zero, then the second one is integrable.

theorem ProbabilityTheory.IndepFun.integral_mul_of_nonneg {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X : Ω} {Y : Ω} (hXY : ProbabilityTheory.IndepFun X Y μ) (hXp : 0 X) (hYp : 0 Y) (hXm : AEMeasurable X μ) (hYm : AEMeasurable Y μ) :

The (Bochner) integral of the product of two independent, nonnegative random variables is the product of their integrals. The proof is just plumbing around lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'.

The (Bochner) integral of the product of two independent, integrable random variables is the product of their integrals. The proof is pedestrian decomposition into their positive and negative parts in order to apply IndepFun.integral_mul_of_nonneg four times.

The (Bochner) integral of the product of two independent random variables is the product of their integrals.

theorem ProbabilityTheory.IndepFun.integral_mul' {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X : Ω} {Y : Ω} (hXY : ProbabilityTheory.IndepFun X Y μ) (hX : MeasureTheory.AEStronglyMeasurable X μ) (hY : MeasureTheory.AEStronglyMeasurable Y μ) :
∫ (ω : Ω), X ω * Y ωμ = MeasureTheory.integral μ X * MeasureTheory.integral μ Y
theorem ProbabilityTheory.indepFun_iff_integral_comp_mul {Ω : Type u_1} {mΩ : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] {β : Type u_2} {β' : Type u_3} {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} {f : Ωβ} {g : Ωβ'} {hfm : Measurable f} {hgm : Measurable g} :
ProbabilityTheory.IndepFun f g μ ∀ {φ : β} {ψ : β'}, Measurable φMeasurable ψMeasureTheory.Integrable (φ f) μMeasureTheory.Integrable (ψ g) μMeasureTheory.integral μ (φ f * ψ g) = MeasureTheory.integral μ (φ f) * MeasureTheory.integral μ (ψ g)

Independence of functions f and g into arbitrary types is characterized by the relation E[(φ ∘ f) * (ψ ∘ g)] = E[φ ∘ f] * E[ψ ∘ g] for all measurable φ and ψ with values in satisfying appropriate integrability conditions.