Documentation

Mathlib.Probability.Moments.Variance

Variance of random variables #

We define the variance of a real-valued random variable as Var[X] = 𝔼[(X - 𝔼[X])^2] (in the ProbabilityTheory locale).

Main definitions #

Main results #

def ProbabilityTheory.evariance {Ω : Type u_1} { : MeasurableSpace Ω} (X : Ω) (μ : MeasureTheory.Measure Ω) :

The ℝ≥0∞-valued variance of a real-valued random variable defined as the Lebesgue integral of ‖X - 𝔼[X]‖^2.

Equations
    Instances For
      def ProbabilityTheory.variance {Ω : Type u_1} { : MeasurableSpace Ω} (X : Ω) (μ : MeasureTheory.Measure Ω) :

      The -valued variance of a real-valued random variable defined by applying ENNReal.toReal to evariance.

      Equations
        Instances For

          The ℝ≥0∞-valued variance of the real-valued random variable X according to the measure μ.

          This is defined as the Lebesgue integral of (X - 𝔼[X])^2.

          Equations
            Instances For

              The ℝ≥0∞-valued variance of the real-valued random variable X according to the volume measure.

              This is defined as the Lebesgue integral of (X - 𝔼[X])^2.

              Equations
                Instances For

                  The -valued variance of the real-valued random variable X according to the measure μ.

                  It is set to 0 if X has infinite variance.

                  Equations
                    Instances For

                      The -valued variance of the real-valued random variable X according to the volume measure.

                      It is set to 0 if X has infinite variance.

                      Equations
                        Instances For
                          theorem ProbabilityTheory.evariance_congr {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} (h : X =ᵐ[μ] Y) :
                          theorem ProbabilityTheory.variance_congr {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} (h : X =ᵐ[μ] Y) :
                          variance X μ = variance Y μ
                          @[deprecated ProbabilityTheory.evariance_lt_top_iff_memLp (since := "2025-02-21")]

                          Alias of ProbabilityTheory.evariance_lt_top_iff_memLp.

                          @[deprecated ProbabilityTheory.evariance_lt_top (since := "2025-02-21")]
                          theorem MeasureTheory.Memℒp.evariance_lt_top {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : Measure Ω} [IsFiniteMeasure μ] (hX : MemLp X 2 μ) :

                          Alias of ProbabilityTheory.evariance_lt_top.

                          @[deprecated ProbabilityTheory.evariance_ne_top (since := "2025-02-21")]
                          theorem MeasureTheory.Memℒp.evariance_ne_top {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : Measure Ω} [IsFiniteMeasure μ] (hX : MemLp X 2 μ) :

                          Alias of ProbabilityTheory.evariance_ne_top.

                          @[deprecated ProbabilityTheory.ofReal_variance (since := "2025-02-21")]

                          Alias of ProbabilityTheory.ofReal_variance.

                          theorem ProbabilityTheory.evariance_eq_lintegral_ofReal {Ω : Type u_1} { : MeasurableSpace Ω} (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          evariance X μ = ∫⁻ (ω : Ω), ENNReal.ofReal ((X ω - (x : Ω), X x μ) ^ 2) μ
                          theorem ProbabilityTheory.variance_eq_integral {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} (hX : AEMeasurable X μ) :
                          variance X μ = (ω : Ω), (X ω - (x : Ω), X x μ) ^ 2 μ
                          theorem ProbabilityTheory.variance_of_integral_eq_zero {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} (hX : AEMeasurable X μ) (hXint : (x : Ω), X x μ = 0) :
                          variance X μ = (ω : Ω), X ω ^ 2 μ
                          @[deprecated ProbabilityTheory.variance_of_integral_eq_zero (since := "2025-01-23")]
                          theorem MeasureTheory.Memℒp.variance_eq_of_integral_eq_zero {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : Measure Ω} (hX : AEMeasurable X μ) (hXint : (x : Ω), X x μ = 0) :
                          ProbabilityTheory.variance X μ = (ω : Ω), X ω ^ 2 μ

                          Alias of ProbabilityTheory.variance_of_integral_eq_zero.

                          @[simp]
                          theorem ProbabilityTheory.evariance_eq_zero_iff {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} (hX : AEMeasurable X μ) :
                          evariance X μ = 0 X =ᵐ[μ] fun (x : Ω) => (x : Ω), X x μ
                          theorem ProbabilityTheory.evariance_mul {Ω : Type u_1} { : MeasurableSpace Ω} (c : ) (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          evariance (fun (ω : Ω) => c * X ω) μ = ENNReal.ofReal (c ^ 2) * evariance X μ
                          @[simp]
                          theorem ProbabilityTheory.covariance_self {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X : Ω} (hX : AEMeasurable X μ) :
                          covariance X X μ = variance X μ
                          @[deprecated ProbabilityTheory.covariance_self (since := "2025-06-25")]
                          theorem ProbabilityTheory.covariance_same {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X : Ω} (hX : AEMeasurable X μ) :
                          covariance X X μ = variance X μ

                          Alias of ProbabilityTheory.covariance_self.

                          theorem ProbabilityTheory.variance_nonneg {Ω : Type u_1} { : MeasurableSpace Ω} (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          0 variance X μ
                          theorem ProbabilityTheory.variance_mul {Ω : Type u_1} { : MeasurableSpace Ω} (c : ) (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          variance (fun (ω : Ω) => c * X ω) μ = c ^ 2 * variance X μ
                          theorem ProbabilityTheory.variance_smul {Ω : Type u_1} { : MeasurableSpace Ω} (c : ) (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          variance (c X) μ = c ^ 2 * variance X μ
                          theorem ProbabilityTheory.variance_smul' {Ω : Type u_1} { : MeasurableSpace Ω} {A : Type u_2} [CommSemiring A] [Algebra A ] (c : A) (X : Ω) (μ : MeasureTheory.Measure Ω) :
                          variance (c X) μ = c ^ 2 variance X μ
                          theorem ProbabilityTheory.variance_eq_sub {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] {X : Ω} (hX : MeasureTheory.MemLp X 2 μ) :
                          variance X μ = (x : Ω), (X ^ 2) x μ - ( (x : Ω), X x μ) ^ 2
                          @[deprecated ProbabilityTheory.variance_eq_sub (since := "2025-08-07")]
                          theorem ProbabilityTheory.variance_def' {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] {X : Ω} (hX : MeasureTheory.MemLp X 2 μ) :
                          variance X μ = (x : Ω), (X ^ 2) x μ - ( (x : Ω), X x μ) ^ 2

                          Alias of ProbabilityTheory.variance_eq_sub.

                          theorem ProbabilityTheory.variance_add_const {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] (hX : MeasureTheory.AEStronglyMeasurable X μ) (c : ) :
                          variance (fun (ω : Ω) => X ω + c) μ = variance X μ
                          theorem ProbabilityTheory.variance_const_add {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] (hX : MeasureTheory.AEStronglyMeasurable X μ) (c : ) :
                          variance (fun (ω : Ω) => c + X ω) μ = variance X μ
                          theorem ProbabilityTheory.variance_fun_neg {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} :
                          variance (fun (ω : Ω) => -X ω) μ = variance X μ
                          theorem ProbabilityTheory.variance_neg {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} :
                          variance (-X) μ = variance X μ
                          theorem ProbabilityTheory.variance_sub_const {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] (hX : MeasureTheory.AEStronglyMeasurable X μ) (c : ) :
                          variance (fun (ω : Ω) => X ω - c) μ = variance X μ
                          theorem ProbabilityTheory.variance_const_sub {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] (hX : MeasureTheory.AEStronglyMeasurable X μ) (c : ) :
                          variance (fun (ω : Ω) => c - X ω) μ = variance X μ
                          theorem ProbabilityTheory.variance_add {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] (hX : MeasureTheory.MemLp X 2 μ) (hY : MeasureTheory.MemLp Y 2 μ) :
                          variance (X + Y) μ = variance X μ + 2 * covariance X Y μ + variance Y μ
                          theorem ProbabilityTheory.variance_fun_add {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] (hX : MeasureTheory.MemLp X 2 μ) (hY : MeasureTheory.MemLp Y 2 μ) :
                          variance (fun (ω : Ω) => X ω + Y ω) μ = variance X μ + 2 * covariance X Y μ + variance Y μ
                          theorem ProbabilityTheory.variance_sub {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] (hX : MeasureTheory.MemLp X 2 μ) (hY : MeasureTheory.MemLp Y 2 μ) :
                          variance (X - Y) μ = variance X μ - 2 * covariance X Y μ + variance Y μ
                          theorem ProbabilityTheory.variance_fun_sub {Ω : Type u_1} { : MeasurableSpace Ω} {X Y : Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] (hX : MeasureTheory.MemLp X 2 μ) (hY : MeasureTheory.MemLp Y 2 μ) :
                          variance (fun (ω : Ω) => X ω - Y ω) μ = variance X μ - 2 * covariance X Y μ + variance Y μ
                          theorem ProbabilityTheory.variance_sum' {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {ι : Type u_2} {s : Finset ι} {X : ιΩ} [MeasureTheory.IsFiniteMeasure μ] (hX : is, MeasureTheory.MemLp (X i) 2 μ) :
                          variance (∑ is, X i) μ = is, js, covariance (X i) (X j) μ
                          theorem ProbabilityTheory.variance_sum {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {ι : Type u_2} {X : ιΩ} [MeasureTheory.IsFiniteMeasure μ] [Fintype ι] (hX : ∀ (i : ι), MeasureTheory.MemLp (X i) 2 μ) :
                          variance (∑ i : ι, X i) μ = i : ι, j : ι, covariance (X i) (X j) μ
                          theorem ProbabilityTheory.variance_fun_sum' {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {ι : Type u_2} {s : Finset ι} {X : ιΩ} [MeasureTheory.IsFiniteMeasure μ] (hX : is, MeasureTheory.MemLp (X i) 2 μ) :
                          variance (fun (ω : Ω) => is, X i ω) μ = is, js, covariance (X i) (X j) μ
                          theorem ProbabilityTheory.variance_fun_sum {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {ι : Type u_2} {X : ιΩ} [MeasureTheory.IsFiniteMeasure μ] [Fintype ι] (hX : ∀ (i : ι), MeasureTheory.MemLp (X i) 2 μ) :
                          variance (fun (ω : Ω) => i : ι, X i ω) μ = i : ι, j : ι, covariance (X i) (X j) μ
                          @[simp]
                          theorem ProbabilityTheory.variance_map {Ω : Type u_1} { : MeasurableSpace Ω} {X : Ω} {Ω' : Type u_3} {mΩ' : MeasurableSpace Ω'} {μ : MeasureTheory.Measure Ω'} {Y : Ω'Ω} (hX : AEMeasurable X (MeasureTheory.Measure.map Y μ)) (hY : AEMeasurable Y μ) :
                          theorem ProbabilityTheory.variance_map_equiv {Ω : Type u_1} { : MeasurableSpace Ω} {Ω' : Type u_3} {mΩ' : MeasurableSpace Ω'} {μ : MeasureTheory.Measure Ω'} (X : Ω) (Y : Ω' ≃ᵐ Ω) :
                          variance X (MeasureTheory.Measure.map (⇑Y) μ) = variance (X Y) μ
                          theorem ProbabilityTheory.meas_ge_le_evariance_div_sq {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X : Ω} (hX : MeasureTheory.AEStronglyMeasurable X μ) {c : NNReal} (hc : c 0) :
                          μ {ω : Ω | c |X ω - (x : Ω), X x μ|} evariance X μ / c ^ 2

                          Chebyshev's inequality for ℝ≥0∞-valued variance.

                          theorem ProbabilityTheory.meas_ge_le_variance_div_sq {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsFiniteMeasure μ] {X : Ω} (hX : MeasureTheory.MemLp X 2 μ) {c : } (hc : 0 < c) :
                          μ {ω : Ω | c |X ω - (x : Ω), X x μ|} ENNReal.ofReal (variance X μ / c ^ 2)

                          Chebyshev's inequality: one can control the deviation probability of a real random variable from its expectation in terms of the variance.

                          theorem ProbabilityTheory.IndepFun.variance_add {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {X Y : Ω} (hX : MeasureTheory.MemLp X 2 μ) (hY : MeasureTheory.MemLp Y 2 μ) (h : IndepFun X Y μ) :
                          variance (X + Y) μ = variance X μ + variance Y μ

                          The variance of the sum of two independent random variables is the sum of the variances.

                          theorem ProbabilityTheory.IndepFun.variance_sum {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} {ι : Type u_3} {X : ιΩ} {s : Finset ι} (hs : is, MeasureTheory.MemLp (X i) 2 μ) (h : (↑s).Pairwise fun (i j : ι) => IndepFun (X i) (X j) μ) :
                          variance (∑ is, X i) μ = is, variance (X i) μ

                          The variance of a finite sum of pairwise independent random variables is the sum of the variances.

                          theorem ProbabilityTheory.variance_le_sub_mul_sub {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] {a b : } {X : Ω} (h : ∀ᵐ (ω : Ω) μ, X ω Set.Icc a b) (hX : AEMeasurable X μ) :
                          variance X μ (b - (x : Ω), X x μ) * ( (x : Ω), X x μ - a)

                          The Bhatia-Davis inequality on variance

                          The variance of a random variable X satisfying a ≤ X ≤ b almost everywhere is at most (b - 𝔼 X) * (𝔼 X - a).

                          theorem ProbabilityTheory.variance_le_sq_of_bounded {Ω : Type u_1} { : MeasurableSpace Ω} {μ : MeasureTheory.Measure Ω} [MeasureTheory.IsProbabilityMeasure μ] {a b : } {X : Ω} (h : ∀ᵐ (ω : Ω) μ, X ω Set.Icc a b) (hX : AEMeasurable X μ) :
                          variance X μ ((b - a) / 2) ^ 2

                          Popoviciu's inequality on variance

                          The variance of a random variable X satisfying a ≤ X ≤ b almost everywhere is at most ((b - a) / 2) ^ 2.