Documentation

Mathlib.MeasureTheory.Integral.Bochner

Bochner integral #

The Bochner integral extends the definition of the Lebesgue integral to functions that map from a measure space into a Banach space (complete normed vector space). It is constructed here using the L1 Bochner integral constructed in the file Mathlib.MeasureTheory.Integral.BochnerL1.

Main definitions #

The Bochner integral is defined through the extension process described in the file Mathlib.MeasureTheory.Integral.SetToL1, which follows these steps:

The result of that construction is ∫ a, f a ∂μ, which is definitionally equal to setToFun (dominatedFinMeasAdditive_weightedSMul μ) f. Some basic properties of the integral (like linearity) are particular cases of the properties of setToFun (which are described in the file Mathlib.MeasureTheory.Integral.SetToL1).

Main statements #

  1. Basic properties of the Bochner integral on functions of type α → E, where α is a measure space and E is a real normed space.
  1. Basic properties of the Bochner integral on functions of type α → ℝ, where α is a measure space.
  1. Propositions connecting the Bochner integral with the integral on ℝ≥0∞-valued functions, which is called lintegral and has the notation ∫⁻.
  1. (In the file Mathlib.MeasureTheory.Integral.DominatedConvergence) tendsto_integral_of_dominated_convergence : the Lebesgue dominated convergence theorem

  2. (In the file Mathlib.MeasureTheory.Integral.SetIntegral) integration commutes with continuous linear maps.

Notes #

Some tips on how to prove a proposition if the API for the Bochner integral is not enough so that you need to unfold the definition of the Bochner integral and go back to simple functions.

One method is to use the theorem Integrable.induction in the file Mathlib.MeasureTheory.Function.SimpleFuncDenseLp (or one of the related results, like Lp.induction for functions in Lp), which allows you to prove something for an arbitrary integrable function.

Another method is using the following steps. See integral_eq_lintegral_pos_part_sub_lintegral_neg_part for a complicated example, which proves that ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻, with the first integral sign being the Bochner integral of a real-valued function f : α → ℝ, and second and third integral sign being the integral on ℝ≥0∞-valued functions (called lintegral). The proof of integral_eq_lintegral_pos_part_sub_lintegral_neg_part is scattered in sections with the name posPart.

Here are the usual steps of proving that a property p, say ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻, holds for all functions :

  1. First go to the space.

    For example, if you see ENNReal.toReal (∫⁻ a, ENNReal.ofReal <| ‖f a‖), that is the norm of f in space. Rewrite using L1.norm_of_fun_eq_lintegral_norm.

  2. Show that the set {f ∈ L¹ | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻} is closed in using isClosed_eq.

  3. Show that the property holds for all simple functions s in space.

    Typically, you need to convert various notions to their SimpleFunc counterpart, using lemmas like L1.integral_coe_eq_integral.

  4. Since simple functions are dense in ,

univ = closure {s simple}
     = closure {s simple | ∫ s = ∫⁻ s⁺ - ∫⁻ s⁻} : the property holds for all simple functions
     ⊆ closure {f | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻}
     = {f | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻} : closure of a closed set is itself

Use isClosed_property or DenseRange.induction_on for this argument.

Notations #

We also define notations for integral on a set, which are described in the file Mathlib/MeasureTheory/Integral/SetIntegral.lean.

Note : is typed using \_s. Sometimes it shows as a box if the font is missing.

Tags #

Bochner integral, simple function, function space, Lebesgue dominated convergence theorem

The Bochner integral on functions #

Define the Bochner integral on functions generally to be the L1 Bochner integral, for integrable functions, and 0 otherwise; prove its basic properties.

theorem MeasureTheory.integral_def {α : Type u_6} {G : Type u_7} [NormedAddCommGroup G] [NormedSpace G] {x✝ : MeasurableSpace α} (μ : Measure α) (f : αG) :
integral μ f = if x : CompleteSpace G then if hf : Integrable f μ then L1.integral (Integrable.toL1 f hf) else 0 else 0
@[irreducible]
def MeasureTheory.integral {α : Type u_6} {G : Type u_7} [NormedAddCommGroup G] [NormedSpace G] {x✝ : MeasurableSpace α} (μ : Measure α) (f : αG) :
G

The Bochner integral

Equations
Instances For

    In the notation for integrals, an expression like ∫ x, g ‖x‖ ∂μ will not be parsed correctly, and needs parentheses. We do not set the binding power of r to 0, because then ∫ x, f x = 0 will be parsed incorrectly.

    The Bochner integral

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      Pretty printer defined by notation3 command.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        The Bochner integral

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          Pretty printer defined by notation3 command.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            Pretty printer defined by notation3 command.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For

              The Bochner integral

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                Pretty printer defined by notation3 command.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  The Bochner integral

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    theorem MeasureTheory.integral_eq {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (f : αE) (hf : Integrable f μ) :
                    (a : α), f a μ = L1.integral (Integrable.toL1 f hf)
                    theorem MeasureTheory.integral_eq_setToFun {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (f : αE) :
                    (a : α), f a μ = setToFun μ (weightedSMul μ) f
                    theorem MeasureTheory.L1.integral_eq_integral {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (f : (Lp E 1 μ)) :
                    integral f = (a : α), f a μ
                    theorem MeasureTheory.integral_undef {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} (h : ¬Integrable f μ) :
                    (a : α), f a μ = 0
                    theorem MeasureTheory.Integrable.of_integral_ne_zero {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} (h : (a : α), f a μ 0) :
                    theorem MeasureTheory.integral_non_aestronglyMeasurable {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} (h : ¬AEStronglyMeasurable f μ) :
                    (a : α), f a μ = 0
                    @[simp]
                    theorem MeasureTheory.integral_zero (α : Type u_1) (G : Type u_5) [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} :
                    (x : α), 0 μ = 0
                    @[simp]
                    theorem MeasureTheory.integral_zero' (α : Type u_1) (G : Type u_5) [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} :
                    integral μ 0 = 0
                    theorem MeasureTheory.integrable_of_integral_eq_one {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (h : (x : α), f x μ = 1) :
                    theorem MeasureTheory.integral_add {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f g : αG} (hf : Integrable f μ) (hg : Integrable g μ) :
                    (a : α), f a + g a μ = (a : α), f a μ + (a : α), g a μ
                    theorem MeasureTheory.integral_add' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f g : αG} (hf : Integrable f μ) (hg : Integrable g μ) :
                    (a : α), (f + g) a μ = (a : α), f a μ + (a : α), g a μ
                    theorem MeasureTheory.integral_finset_sum {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} (s : Finset ι) {f : ιαG} (hf : is, Integrable (f i) μ) :
                    (a : α), is, f i a μ = is, (a : α), f i a μ
                    theorem MeasureTheory.integral_neg {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), -f a μ = - (a : α), f a μ
                    theorem MeasureTheory.integral_neg' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), (-f) a μ = - (a : α), f a μ
                    theorem MeasureTheory.integral_sub {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f g : αG} (hf : Integrable f μ) (hg : Integrable g μ) :
                    (a : α), f a - g a μ = (a : α), f a μ - (a : α), g a μ
                    theorem MeasureTheory.integral_sub' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f g : αG} (hf : Integrable f μ) (hg : Integrable g μ) :
                    (a : α), (f - g) a μ = (a : α), f a μ - (a : α), g a μ
                    theorem MeasureTheory.integral_smul {α : Type u_1} {𝕜 : Type u_4} [NontriviallyNormedField 𝕜] {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} [NormedSpace 𝕜 G] [SMulCommClass 𝕜 G] (c : 𝕜) (f : αG) :
                    (a : α), c f a μ = c (a : α), f a μ
                    theorem MeasureTheory.integral_mul_left {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {L : Type u_6} [RCLike L] (r : L) (f : αL) :
                    (a : α), r * f a μ = r * (a : α), f a μ
                    theorem MeasureTheory.integral_mul_right {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {L : Type u_6} [RCLike L] (r : L) (f : αL) :
                    (a : α), f a * r μ = ( (a : α), f a μ) * r
                    theorem MeasureTheory.integral_div {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {L : Type u_6} [RCLike L] (r : L) (f : αL) :
                    (a : α), f a / r μ = ( (a : α), f a μ) / r
                    theorem MeasureTheory.integral_congr_ae {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f g : αG} (h : f =ᶠ[ae μ] g) :
                    (a : α), f a μ = (a : α), g a μ
                    theorem MeasureTheory.integral_congr_ae₂ {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_6} {x✝ : MeasurableSpace β} {ν : Measure β} {f g : αβG} (h : ∀ᵐ (a : α) μ, f a =ᶠ[ae ν] g a) :
                    (a : α), (b : β), f a b ν μ = (a : α), (b : β), g a b ν μ
                    @[simp]
                    theorem MeasureTheory.L1.integral_of_fun_eq_integral {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} (hf : Integrable f μ) :
                    (a : α), (Integrable.toL1 f hf) a μ = (a : α), f a μ
                    theorem MeasureTheory.continuous_integral {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} :
                    Continuous fun (f : (Lp G 1 μ)) => (a : α), f a μ
                    theorem MeasureTheory.norm_integral_le_lintegral_norm {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), f a μ (∫⁻ (a : α), ENNReal.ofReal f a μ).toReal
                    theorem MeasureTheory.enorm_integral_le_lintegral_enorm {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), f a μ‖ₑ ∫⁻ (a : α), f a‖ₑ μ
                    @[deprecated MeasureTheory.enorm_integral_le_lintegral_enorm (since := "2025-01-21")]
                    theorem MeasureTheory.ennnorm_integral_le_lintegral_ennnorm {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), f a μ‖ₑ ∫⁻ (a : α), f a‖ₑ μ

                    Alias of MeasureTheory.enorm_integral_le_lintegral_enorm.

                    theorem MeasureTheory.integral_eq_zero_of_ae {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} (hf : f =ᶠ[ae μ] 0) :
                    (a : α), f a μ = 0
                    theorem MeasureTheory.HasFiniteIntegral.tendsto_setIntegral_nhds_zero {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} {f : αG} (hf : HasFiniteIntegral f μ) {l : Filter ι} {s : ιSet α} (hs : Filter.Tendsto (μ s) l (nhds 0)) :
                    Filter.Tendsto (fun (i : ι) => (x : α) in s i, f x μ) l (nhds 0)

                    If f has finite integral, then ∫ x in s, f x ∂μ is absolutely continuous in s: it tends to zero as μ s tends to zero.

                    theorem MeasureTheory.Integrable.tendsto_setIntegral_nhds_zero {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} {f : αG} (hf : Integrable f μ) {l : Filter ι} {s : ιSet α} (hs : Filter.Tendsto (μ s) l (nhds 0)) :
                    Filter.Tendsto (fun (i : ι) => (x : α) in s i, f x μ) l (nhds 0)

                    If f is integrable, then ∫ x in s, f x ∂μ is absolutely continuous in s: it tends to zero as μ s tends to zero.

                    theorem MeasureTheory.tendsto_integral_of_L1 {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} (f : αG) (hfi : Integrable f μ) {F : ιαG} {l : Filter ι} (hFi : ∀ᶠ (i : ι) in l, Integrable (F i) μ) (hF : Filter.Tendsto (fun (i : ι) => ∫⁻ (x : α), F i x - f x‖ₑ μ) l (nhds 0)) :
                    Filter.Tendsto (fun (i : ι) => (x : α), F i x μ) l (nhds ( (x : α), f x μ))

                    If F i → f in L1, then ∫ x, F i x ∂μ → ∫ x, f x ∂μ.

                    theorem MeasureTheory.tendsto_integral_of_L1' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} (f : αG) (hfi : Integrable f μ) {F : ιαG} {l : Filter ι} (hFi : ∀ᶠ (i : ι) in l, Integrable (F i) μ) (hF : Filter.Tendsto (fun (i : ι) => eLpNorm (F i - f) 1 μ) l (nhds 0)) :
                    Filter.Tendsto (fun (i : ι) => (x : α), F i x μ) l (nhds ( (x : α), f x μ))

                    If F i → f in L1, then ∫ x, F i x ∂μ → ∫ x, f x ∂μ.

                    theorem MeasureTheory.tendsto_setIntegral_of_L1 {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} (f : αG) (hfi : Integrable f μ) {F : ιαG} {l : Filter ι} (hFi : ∀ᶠ (i : ι) in l, Integrable (F i) μ) (hF : Filter.Tendsto (fun (i : ι) => ∫⁻ (x : α), F i x - f x‖ₑ μ) l (nhds 0)) (s : Set α) :
                    Filter.Tendsto (fun (i : ι) => (x : α) in s, F i x μ) l (nhds ( (x : α) in s, f x μ))

                    If F i → f in L1, then ∫ x in s, F i x ∂μ → ∫ x in s, f x ∂μ.

                    theorem MeasureTheory.tendsto_setIntegral_of_L1' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {ι : Type u_6} (f : αG) (hfi : Integrable f μ) {F : ιαG} {l : Filter ι} (hFi : ∀ᶠ (i : ι) in l, Integrable (F i) μ) (hF : Filter.Tendsto (fun (i : ι) => eLpNorm (F i - f) 1 μ) l (nhds 0)) (s : Set α) :
                    Filter.Tendsto (fun (i : ι) => (x : α) in s, F i x μ) l (nhds ( (x : α) in s, f x μ))

                    If F i → f in L1, then ∫ x in s, F i x ∂μ → ∫ x in s, f x ∂μ.

                    theorem MeasureTheory.continuousWithinAt_of_dominated {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {X : Type u_6} [TopologicalSpace X] [FirstCountableTopology X] {F : XαG} {x₀ : X} {bound : α} {s : Set X} (hF_meas : ∀ᶠ (x : X) in nhdsWithin x₀ s, AEStronglyMeasurable (F x) μ) (h_bound : ∀ᶠ (x : X) in nhdsWithin x₀ s, ∀ᵐ (a : α) μ, F x a bound a) (bound_integrable : Integrable bound μ) (h_cont : ∀ᵐ (a : α) μ, ContinuousWithinAt (fun (x : X) => F x a) s x₀) :
                    ContinuousWithinAt (fun (x : X) => (a : α), F x a μ) s x₀
                    theorem MeasureTheory.continuousAt_of_dominated {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {X : Type u_6} [TopologicalSpace X] [FirstCountableTopology X] {F : XαG} {x₀ : X} {bound : α} (hF_meas : ∀ᶠ (x : X) in nhds x₀, AEStronglyMeasurable (F x) μ) (h_bound : ∀ᶠ (x : X) in nhds x₀, ∀ᵐ (a : α) μ, F x a bound a) (bound_integrable : Integrable bound μ) (h_cont : ∀ᵐ (a : α) μ, ContinuousAt (fun (x : X) => F x a) x₀) :
                    ContinuousAt (fun (x : X) => (a : α), F x a μ) x₀
                    theorem MeasureTheory.continuousOn_of_dominated {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {X : Type u_6} [TopologicalSpace X] [FirstCountableTopology X] {F : XαG} {bound : α} {s : Set X} (hF_meas : xs, AEStronglyMeasurable (F x) μ) (h_bound : xs, ∀ᵐ (a : α) μ, F x a bound a) (bound_integrable : Integrable bound μ) (h_cont : ∀ᵐ (a : α) μ, ContinuousOn (fun (x : X) => F x a) s) :
                    ContinuousOn (fun (x : X) => (a : α), F x a μ) s
                    theorem MeasureTheory.continuous_of_dominated {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {X : Type u_6} [TopologicalSpace X] [FirstCountableTopology X] {F : XαG} {bound : α} (hF_meas : ∀ (x : X), AEStronglyMeasurable (F x) μ) (h_bound : ∀ (x : X), ∀ᵐ (a : α) μ, F x a bound a) (bound_integrable : Integrable bound μ) (h_cont : ∀ᵐ (a : α) μ, Continuous fun (x : X) => F x a) :
                    Continuous fun (x : X) => (a : α), F x a μ
                    theorem MeasureTheory.integral_eq_lintegral_pos_part_sub_lintegral_neg_part {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : Integrable f μ) :
                    (a : α), f a μ = (∫⁻ (a : α), ENNReal.ofReal (f a) μ).toReal - (∫⁻ (a : α), ENNReal.ofReal (-f a) μ).toReal

                    The Bochner integral of a real-valued function f : α → ℝ is the difference between the integral of the positive part of f and the integral of the negative part of f.

                    theorem MeasureTheory.integral_eq_lintegral_of_nonneg_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 ≤ᶠ[ae μ] f) (hfm : AEStronglyMeasurable f μ) :
                    (a : α), f a μ = (∫⁻ (a : α), ENNReal.ofReal (f a) μ).toReal
                    theorem MeasureTheory.integral_norm_eq_lintegral_enorm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {P : Type u_7} [NormedAddCommGroup P] {f : αP} (hf : AEStronglyMeasurable f μ) :
                    (x : α), f x μ = (∫⁻ (x : α), f x‖ₑ μ).toReal
                    @[deprecated MeasureTheory.integral_norm_eq_lintegral_enorm (since := "2025-01-21")]
                    theorem MeasureTheory.integral_norm_eq_lintegral_nnnorm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {P : Type u_7} [NormedAddCommGroup P] {f : αP} (hf : AEStronglyMeasurable f μ) :
                    (x : α), f x μ = (∫⁻ (x : α), f x‖ₑ μ).toReal

                    Alias of MeasureTheory.integral_norm_eq_lintegral_enorm.

                    theorem MeasureTheory.ofReal_integral_norm_eq_lintegral_enorm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {P : Type u_7} [NormedAddCommGroup P] {f : αP} (hf : Integrable f μ) :
                    ENNReal.ofReal ( (x : α), f x μ) = ∫⁻ (x : α), f x‖ₑ μ
                    @[deprecated MeasureTheory.ofReal_integral_norm_eq_lintegral_enorm (since := "2025-01-21")]
                    theorem MeasureTheory.ofReal_integral_norm_eq_lintegral_nnnorm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {P : Type u_7} [NormedAddCommGroup P] {f : αP} (hf : Integrable f μ) :
                    ENNReal.ofReal ( (x : α), f x μ) = ∫⁻ (x : α), f x‖ₑ μ

                    Alias of MeasureTheory.ofReal_integral_norm_eq_lintegral_enorm.

                    theorem MeasureTheory.integral_eq_integral_pos_part_sub_integral_neg_part {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : Integrable f μ) :
                    (a : α), f a μ = (a : α), (f a).toNNReal μ - (a : α), (-f a).toNNReal μ
                    theorem MeasureTheory.integral_nonneg_of_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 ≤ᶠ[ae μ] f) :
                    0 (a : α), f a μ
                    theorem MeasureTheory.lintegral_coe_eq_integral {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} (f : αNNReal) (hfi : Integrable (fun (x : α) => (f x)) μ) :
                    ∫⁻ (a : α), (f a) μ = ENNReal.ofReal ( (a : α), (f a) μ)
                    theorem MeasureTheory.ofReal_integral_eq_lintegral_ofReal {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hfi : Integrable f μ) (f_nn : 0 ≤ᶠ[ae μ] f) :
                    ENNReal.ofReal ( (x : α), f x μ) = ∫⁻ (x : α), ENNReal.ofReal (f x) μ
                    theorem MeasureTheory.integral_toReal {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : αENNReal} (hfm : AEMeasurable f μ) (hf : ∀ᵐ (x : α) μ, f x < ) :
                    (a : α), (f a).toReal μ = (∫⁻ (a : α), f a μ).toReal
                    theorem MeasureTheory.lintegral_coe_le_coe_iff_integral_le {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : αNNReal} (hfi : Integrable (fun (x : α) => (f x)) μ) {b : NNReal} :
                    ∫⁻ (a : α), (f a) μ b (a : α), (f a) μ b
                    theorem MeasureTheory.integral_coe_le_of_lintegral_coe_le {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : αNNReal} {b : NNReal} (h : ∫⁻ (a : α), (f a) μ b) :
                    (a : α), (f a) μ b
                    theorem MeasureTheory.integral_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 f) :
                    0 (a : α), f a μ
                    theorem MeasureTheory.integral_nonpos_of_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : f ≤ᶠ[ae μ] 0) :
                    (a : α), f a μ 0
                    theorem MeasureTheory.integral_nonpos {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : f 0) :
                    (a : α), f a μ 0
                    theorem MeasureTheory.integral_eq_zero_iff_of_nonneg_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 ≤ᶠ[ae μ] f) (hfi : Integrable f μ) :
                    (x : α), f x μ = 0 f =ᶠ[ae μ] 0
                    theorem MeasureTheory.integral_eq_zero_iff_of_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 f) (hfi : Integrable f μ) :
                    (x : α), f x μ = 0 f =ᶠ[ae μ] 0
                    theorem MeasureTheory.integral_eq_iff_of_ae_le {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f g : α} (hf : Integrable f μ) (hg : Integrable g μ) (hfg : f ≤ᶠ[ae μ] g) :
                    (a : α), f a μ = (a : α), g a μ f =ᶠ[ae μ] g
                    theorem MeasureTheory.integral_pos_iff_support_of_nonneg_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 ≤ᶠ[ae μ] f) (hfi : Integrable f μ) :
                    0 < (x : α), f x μ 0 < μ (Function.support f)
                    theorem MeasureTheory.integral_pos_iff_support_of_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf : 0 f) (hfi : Integrable f μ) :
                    0 < (x : α), f x μ 0 < μ (Function.support f)
                    theorem MeasureTheory.integral_exp_pos {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} [hμ : NeZero μ] (hf : Integrable (fun (x : α) => Real.exp (f x)) μ) :
                    0 < (x : α), Real.exp (f x) μ
                    theorem MeasureTheory.integral_tendsto_of_tendsto_of_monotone {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} {F : α} (hf : ∀ (n : ), Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ (x : α) μ, Monotone fun (n : ) => f n x) (h_tendsto : ∀ᵐ (x : α) μ, Filter.Tendsto (fun (n : ) => f n x) Filter.atTop (nhds (F x))) :
                    Filter.Tendsto (fun (n : ) => (x : α), f n x μ) Filter.atTop (nhds ( (x : α), F x μ))

                    Monotone convergence theorem for real-valued functions and Bochner integrals

                    theorem MeasureTheory.integral_tendsto_of_tendsto_of_antitone {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} {F : α} (hf : ∀ (n : ), Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ (x : α) μ, Antitone fun (n : ) => f n x) (h_tendsto : ∀ᵐ (x : α) μ, Filter.Tendsto (fun (n : ) => f n x) Filter.atTop (nhds (F x))) :
                    Filter.Tendsto (fun (n : ) => (x : α), f n x μ) Filter.atTop (nhds ( (x : α), F x μ))

                    Monotone convergence theorem for real-valued functions and Bochner integrals

                    theorem MeasureTheory.tendsto_of_integral_tendsto_of_monotone {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} {F : α} (hf_int : ∀ (n : ), Integrable (f n) μ) (hF_int : Integrable F μ) (hf_tendsto : Filter.Tendsto (fun (i : ) => (a : α), f i a μ) Filter.atTop (nhds ( (a : α), F a μ))) (hf_mono : ∀ᵐ (a : α) μ, Monotone fun (i : ) => f i a) (hf_bound : ∀ᵐ (a : α) μ, ∀ (i : ), f i a F a) :
                    ∀ᵐ (a : α) μ, Filter.Tendsto (fun (i : ) => f i a) Filter.atTop (nhds (F a))

                    If a monotone sequence of functions has an upper bound and the sequence of integrals of these functions tends to the integral of the upper bound, then the sequence of functions converges almost everywhere to the upper bound.

                    theorem MeasureTheory.tendsto_of_integral_tendsto_of_antitone {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} {F : α} (hf_int : ∀ (n : ), Integrable (f n) μ) (hF_int : Integrable F μ) (hf_tendsto : Filter.Tendsto (fun (i : ) => (a : α), f i a μ) Filter.atTop (nhds ( (a : α), F a μ))) (hf_mono : ∀ᵐ (a : α) μ, Antitone fun (i : ) => f i a) (hf_bound : ∀ᵐ (a : α) μ, ∀ (i : ), F a f i a) :
                    ∀ᵐ (a : α) μ, Filter.Tendsto (fun (i : ) => f i a) Filter.atTop (nhds (F a))

                    If an antitone sequence of functions has a lower bound and the sequence of integrals of these functions tends to the integral of the lower bound, then the sequence of functions converges almost everywhere to the lower bound.

                    theorem MeasureTheory.L1.norm_eq_integral_norm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {H : Type u_7} [NormedAddCommGroup H] (f : (Lp H 1 μ)) :
                    f = (a : α), f a μ
                    theorem MeasureTheory.L1.dist_eq_integral_dist {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {H : Type u_7} [NormedAddCommGroup H] (f g : (Lp H 1 μ)) :
                    dist f g = (a : α), dist (f a) (g a) μ
                    theorem MeasureTheory.L1.norm_of_fun_eq_integral_norm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {H : Type u_7} [NormedAddCommGroup H] {f : αH} (hf : Integrable f μ) :
                    theorem MeasureTheory.MemLp.eLpNorm_eq_integral_rpow_norm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {H : Type u_7} [NormedAddCommGroup H] {f : αH} {p : ENNReal} (hp1 : p 0) (hp2 : p ) (hf : MemLp f p μ) :
                    eLpNorm f p μ = ENNReal.ofReal (( (a : α), f a ^ p.toReal μ) ^ p.toReal⁻¹)
                    @[deprecated MeasureTheory.MemLp.eLpNorm_eq_integral_rpow_norm (since := "2025-02-21")]
                    theorem MeasureTheory.Memℒp.eLpNorm_eq_integral_rpow_norm {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {H : Type u_7} [NormedAddCommGroup H] {f : αH} {p : ENNReal} (hp1 : p 0) (hp2 : p ) (hf : MemLp f p μ) :
                    eLpNorm f p μ = ENNReal.ofReal (( (a : α), f a ^ p.toReal μ) ^ p.toReal⁻¹)

                    Alias of MeasureTheory.MemLp.eLpNorm_eq_integral_rpow_norm.

                    theorem MeasureTheory.integral_mono_ae {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f g : α} (hf : Integrable f μ) (hg : Integrable g μ) (h : f ≤ᶠ[ae μ] g) :
                    (a : α), f a μ (a : α), g a μ
                    theorem MeasureTheory.integral_mono {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f g : α} (hf : Integrable f μ) (hg : Integrable g μ) (h : f g) :
                    (a : α), f a μ (a : α), g a μ
                    theorem MeasureTheory.integral_mono_of_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f g : α} (hf : 0 ≤ᶠ[ae μ] f) (hgi : Integrable g μ) (h : f ≤ᶠ[ae μ] g) :
                    (a : α), f a μ (a : α), g a μ
                    theorem MeasureTheory.integral_mono_measure {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} {ν : Measure α} (hle : μ ν) (hf : 0 ≤ᶠ[ae ν] f) (hfi : Integrable f ν) :
                    (a : α), f a μ (a : α), f a ν
                    theorem MeasureTheory.norm_integral_le_integral_norm {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) :
                    (a : α), f a μ (a : α), f a μ
                    theorem MeasureTheory.abs_integral_le_integral_abs {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} :
                    | (a : α), f a μ| (a : α), |f a| μ
                    theorem MeasureTheory.norm_integral_le_of_norm_le {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {f : αG} {g : α} (hg : Integrable g μ) (h : ∀ᵐ (x : α) μ, f x g x) :
                    (x : α), f x μ (x : α), g x μ
                    theorem MeasureTheory.SimpleFunc.integral_eq_integral {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (f : SimpleFunc α E) (hfi : Integrable (⇑f) μ) :
                    integral μ f = (x : α), f x μ
                    theorem MeasureTheory.SimpleFunc.integral_eq_sum {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (f : SimpleFunc α E) (hfi : Integrable (⇑f) μ) :
                    (x : α), f x μ = xf.range, (μ (f ⁻¹' {x})).toReal x
                    @[simp]
                    theorem MeasureTheory.integral_const {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} (c : E) :
                    (x : α), c μ = (μ Set.univ).toReal c
                    theorem MeasureTheory.norm_integral_le_of_norm_le_const {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} [IsFiniteMeasure μ] {f : αG} {C : } (h : ∀ᵐ (x : α) μ, f x C) :
                    (x : α), f x μ C * (μ Set.univ).toReal
                    theorem MeasureTheory.tendsto_integral_approxOn_of_measurable {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [MeasurableSpace E] [BorelSpace E] {f : αE} {s : Set E} [TopologicalSpace.SeparableSpace s] (hfi : Integrable f μ) (hfm : Measurable f) (hs : ∀ᵐ (x : α) μ, f x closure s) {y₀ : E} (h₀ : y₀ s) (h₀i : Integrable (fun (x : α) => y₀) μ) :
                    Filter.Tendsto (fun (n : ) => SimpleFunc.integral μ (SimpleFunc.approxOn f hfm s y₀ h₀ n)) Filter.atTop (nhds ( (x : α), f x μ))
                    theorem MeasureTheory.tendsto_integral_approxOn_of_measurable_of_range_subset {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [MeasurableSpace E] [BorelSpace E] {f : αE} (fmeas : Measurable f) (hf : Integrable f μ) (s : Set E) [TopologicalSpace.SeparableSpace s] (hs : Set.range f {0} s) :
                    Filter.Tendsto (fun (n : ) => SimpleFunc.integral μ (SimpleFunc.approxOn f fmeas s 0 n)) Filter.atTop (nhds ( (x : α), f x μ))
                    theorem MeasureTheory.tendsto_integral_norm_approxOn_sub {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {E : Type u_7} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] {f : αE} (fmeas : Measurable f) (hf : Integrable f μ) [TopologicalSpace.SeparableSpace ↑(Set.range f {0})] :
                    Filter.Tendsto (fun (n : ) => (x : α), (SimpleFunc.approxOn f fmeas (Set.range f {0}) 0 n) x - f x μ) Filter.atTop (nhds 0)
                    theorem MeasureTheory.integral_add_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ ν : Measure α} {f : αG} (hμ : Integrable f μ) (hν : Integrable f ν) :
                    (x : α), f x (μ + ν) = (x : α), f x μ + (x : α), f x ν
                    @[simp]
                    theorem MeasureTheory.integral_zero_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} (f : αG) :
                    (x : α), f x 0 = 0
                    @[simp]
                    theorem MeasureTheory.setIntegral_zero_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} (f : αG) {μ : Measure α} {s : Set α} (hs : μ s = 0) :
                    (x : α) in s, f x μ = 0
                    theorem MeasureTheory.integral_of_isEmpty {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} [IsEmpty α] {f : αG} :
                    (x : α), f x μ = 0
                    theorem MeasureTheory.integral_finset_sum_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {ι : Type u_7} {m : MeasurableSpace α} {f : αG} {μ : ιMeasure α} {s : Finset ι} (hf : is, Integrable f (μ i)) :
                    (a : α), f a is, μ i = is, (a : α), f a μ i
                    theorem MeasureTheory.nndist_integral_add_measure_le_lintegral {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ ν : Measure α} {f : αG} (h₁ : Integrable f μ) (h₂ : Integrable f ν) :
                    (nndist ( (x : α), f x μ) ( (x : α), f x (μ + ν))) ∫⁻ (x : α), f x‖ₑ ν
                    theorem MeasureTheory.hasSum_integral_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {ι : Type u_7} {m : MeasurableSpace α} {f : αG} {μ : ιMeasure α} (hf : Integrable f (Measure.sum μ)) :
                    HasSum (fun (i : ι) => (a : α), f a μ i) ( (a : α), f a Measure.sum μ)
                    theorem MeasureTheory.integral_sum_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {ι : Type u_7} {x✝ : MeasurableSpace α} {f : αG} {μ : ιMeasure α} (hf : Integrable f (Measure.sum μ)) :
                    (a : α), f a Measure.sum μ = ∑' (i : ι), (a : α), f a μ i
                    @[simp]
                    theorem MeasureTheory.integral_smul_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) (c : ENNReal) :
                    (x : α), f x c μ = c.toReal (x : α), f x μ
                    @[simp]
                    theorem MeasureTheory.integral_smul_nnreal_measure {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} (f : αG) (c : NNReal) :
                    (x : α), f x c μ = c (x : α), f x μ
                    theorem MeasureTheory.integral_map_of_stronglyMeasurable {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_7} [MeasurableSpace β] {φ : αβ} (hφ : Measurable φ) {f : βG} (hfm : StronglyMeasurable f) :
                    (y : β), f y Measure.map φ μ = (x : α), f (φ x) μ
                    theorem MeasureTheory.integral_map {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_7} [MeasurableSpace β] {φ : αβ} (hφ : AEMeasurable φ μ) {f : βG} (hfm : AEStronglyMeasurable f (Measure.map φ μ)) :
                    (y : β), f y Measure.map φ μ = (x : α), f (φ x) μ
                    theorem MeasurableEmbedding.integral_map {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {β : Type u_7} {x✝ : MeasurableSpace β} {f : αβ} (hf : MeasurableEmbedding f) (g : βG) :
                    (y : β), g y MeasureTheory.Measure.map f μ = (x : α), g (f x) μ
                    theorem Topology.IsClosedEmbedding.integral_map {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {β : Type u_7} [TopologicalSpace α] [BorelSpace α] [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : αβ} (hφ : IsClosedEmbedding φ) (f : βG) :
                    (y : β), f y MeasureTheory.Measure.map φ μ = (x : α), f (φ x) μ
                    @[deprecated Topology.IsClosedEmbedding.integral_map (since := "2024-10-20")]
                    theorem ClosedEmbedding.integral_map {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {β : Type u_7} [TopologicalSpace α] [BorelSpace α] [TopologicalSpace β] [MeasurableSpace β] [BorelSpace β] {φ : αβ} (hφ : Topology.IsClosedEmbedding φ) (f : βG) :
                    (y : β), f y MeasureTheory.Measure.map φ μ = (x : α), f (φ x) μ

                    Alias of Topology.IsClosedEmbedding.integral_map.

                    theorem MeasureTheory.integral_map_equiv {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_7} [MeasurableSpace β] (e : α ≃ᵐ β) (f : βG) :
                    (y : β), f y Measure.map (⇑e) μ = (x : α), f (e x) μ
                    theorem MeasureTheory.MeasurePreserving.integral_comp {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_7} {x✝ : MeasurableSpace β} {f : αβ} {ν : Measure β} (h₁ : MeasurePreserving f μ ν) (h₂ : MeasurableEmbedding f) (g : βG) :
                    (x : α), g (f x) μ = (y : β), g y ν
                    theorem MeasureTheory.MeasurePreserving.integral_comp' {α : Type u_1} {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {m : MeasurableSpace α} {μ : Measure α} {β : Type u_7} [MeasurableSpace β] {ν : Measure β} {f : α ≃ᵐ β} (h : MeasurePreserving (⇑f) μ ν) (g : βG) :
                    (x : α), g (f x) μ = (y : β), g y ν
                    theorem MeasureTheory.integral_subtype_comap {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {α : Type u_7} [MeasurableSpace α] {μ : Measure α} {s : Set α} (hs : MeasurableSet s) (f : αG) :
                    (x : s), f x Measure.comap Subtype.val μ = (x : α) in s, f x μ
                    theorem MeasureTheory.integral_subtype {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {α : Type u_7} [MeasureSpace α] {s : Set α} (hs : MeasurableSet s) (f : αG) :
                    (x : s), f x = (x : α) in s, f x
                    @[simp]
                    theorem MeasureTheory.integral_dirac' {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] [MeasurableSpace α] (f : αE) (a : α) (hfm : StronglyMeasurable f) :
                    (x : α), f x Measure.dirac a = f a
                    @[simp]
                    theorem MeasureTheory.integral_dirac {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] [MeasurableSpace α] [MeasurableSingletonClass α] (f : αE) (a : α) :
                    (x : α), f x Measure.dirac a = f a
                    theorem MeasureTheory.setIntegral_dirac' {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {mα : MeasurableSpace α} {f : αE} (hf : StronglyMeasurable f) (a : α) {s : Set α} (hs : MeasurableSet s) [Decidable (a s)] :
                    (x : α) in s, f x Measure.dirac a = if a s then f a else 0
                    theorem MeasureTheory.setIntegral_dirac {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] [MeasurableSpace α] [MeasurableSingletonClass α] (f : αE) (a : α) (s : Set α) [Decidable (a s)] :
                    (x : α) in s, f x Measure.dirac a = if a s then f a else 0
                    theorem MeasureTheory.mul_meas_ge_le_integral_of_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {f : α} (hf_nonneg : 0 ≤ᶠ[ae μ] f) (hf_int : Integrable f μ) (ε : ) :
                    ε * (μ {x : α | ε f x}).toReal (x : α), f x μ

                    Markov's inequality also known as Chebyshev's first inequality.

                    theorem MeasureTheory.integral_mul_norm_le_Lp_mul_Lq {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {E : Type u_7} [NormedAddCommGroup E] {f g : αE} {p q : } (hpq : p.IsConjExponent q) (hf : MemLp f (ENNReal.ofReal p) μ) (hg : MemLp g (ENNReal.ofReal q) μ) :
                    (a : α), f a * g a μ ( (a : α), f a ^ p μ) ^ (1 / p) * ( (a : α), g a ^ q μ) ^ (1 / q)

                    Hölder's inequality for the integral of a product of norms. The integral of the product of two norms of functions is bounded by the product of their ℒp and ℒq seminorms when p and q are conjugate exponents.

                    theorem MeasureTheory.integral_mul_le_Lp_mul_Lq_of_nonneg {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} {p q : } (hpq : p.IsConjExponent q) {f g : α} (hf_nonneg : 0 ≤ᶠ[ae μ] f) (hg_nonneg : 0 ≤ᶠ[ae μ] g) (hf : MemLp f (ENNReal.ofReal p) μ) (hg : MemLp g (ENNReal.ofReal q) μ) :
                    (a : α), f a * g a μ ( (a : α), f a ^ p μ) ^ (1 / p) * ( (a : α), g a ^ q μ) ^ (1 / q)

                    Hölder's inequality for functions α → ℝ. The integral of the product of two nonnegative functions is bounded by the product of their ℒp and ℒq seminorms when p and q are conjugate exponents.

                    theorem MeasureTheory.integral_countable' {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} [Countable α] [MeasurableSingletonClass α] {μ : Measure α} {f : αE} (hf : Integrable f μ) :
                    (a : α), f a μ = ∑' (a : α), (μ {a}).toReal f a
                    theorem MeasureTheory.integral_singleton' {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} {f : αE} (hf : StronglyMeasurable f) (a : α) :
                    (a : α) in {a}, f a μ = (μ {a}).toReal f a
                    theorem MeasureTheory.integral_singleton {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} [MeasurableSingletonClass α] {μ : Measure α} (f : αE) (a : α) :
                    (a : α) in {a}, f a μ = (μ {a}).toReal f a
                    theorem MeasureTheory.integral_countable {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [MeasurableSingletonClass α] (f : αE) {s : Set α} (hs : s.Countable) (hf : IntegrableOn f s μ) :
                    (a : α) in s, f a μ = ∑' (a : s), (μ {a}).toReal f a
                    theorem MeasureTheory.integral_finset {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [MeasurableSingletonClass α] (s : Finset α) (f : αE) (hf : IntegrableOn f (↑s) μ) :
                    (x : α) in s, f x μ = xs, (μ {x}).toReal f x
                    theorem MeasureTheory.integral_fintype {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [MeasurableSingletonClass α] [Fintype α] (f : αE) (hf : Integrable f μ) :
                    (x : α), f x μ = x : α, (μ {x}).toReal f x
                    theorem MeasureTheory.integral_unique {α : Type u_1} {E : Type u_2} [NormedAddCommGroup E] [hE : CompleteSpace E] [NormedSpace E] {m : MeasurableSpace α} {μ : Measure α} [Unique α] (f : αE) :
                    (x : α), f x μ = (μ Set.univ).toReal f default
                    theorem MeasureTheory.integral_pos_of_integrable_nonneg_nonzero {α : Type u_1} {m : MeasurableSpace α} {μ : Measure α} [TopologicalSpace α] [μ.IsOpenPosMeasure] {f : α} {x : α} (f_cont : Continuous f) (f_int : Integrable f μ) (f_nonneg : 0 f) (f_x : f x 0) :
                    0 < (x : α), f x μ
                    def MeasureTheory.SimpleFunc.toLargerSpace {β : Type u_6} {γ : Type u_7} {m m0 : MeasurableSpace β} (hm : m m0) (f : SimpleFunc β γ) :

                    Simple function seen as simple function of a larger MeasurableSpace.

                    Equations
                    Instances For
                      theorem MeasureTheory.SimpleFunc.coe_toLargerSpace_eq {β : Type u_6} {γ : Type u_7} {m m0 : MeasurableSpace β} (hm : m m0) (f : SimpleFunc β γ) :
                      (toLargerSpace hm f) = f
                      theorem MeasureTheory.integral_simpleFunc_larger_space {F : Type u_3} [NormedAddCommGroup F] [NormedSpace F] [CompleteSpace F] {β : Type u_6} {m m0 : MeasurableSpace β} {μ : Measure β} (hm : m m0) (f : SimpleFunc β F) (hf_int : Integrable (⇑f) μ) :
                      (x : β), f x μ = xf.range, (μ (f ⁻¹' {x})).toReal x
                      theorem MeasureTheory.integral_trim_simpleFunc {F : Type u_3} [NormedAddCommGroup F] [NormedSpace F] [CompleteSpace F] {β : Type u_6} {m m0 : MeasurableSpace β} {μ : Measure β} (hm : m m0) (f : SimpleFunc β F) (hf_int : Integrable (⇑f) μ) :
                      (x : β), f x μ = (x : β), f x μ.trim hm
                      theorem MeasureTheory.integral_trim {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {β : Type u_6} {m m0 : MeasurableSpace β} {μ : Measure β} (hm : m m0) {f : βG} (hf : StronglyMeasurable f) :
                      (x : β), f x μ = (x : β), f x μ.trim hm
                      theorem MeasureTheory.integral_trim_ae {G : Type u_5} [NormedAddCommGroup G] [NormedSpace G] {β : Type u_6} {m m0 : MeasurableSpace β} {μ : Measure β} (hm : m m0) {f : βG} (hf : AEStronglyMeasurable f (μ.trim hm)) :
                      (x : β), f x μ = (x : β), f x μ.trim hm
                      theorem MeasureTheory.eLpNorm_one_le_of_le {α : Type u_1} {m0 : MeasurableSpace α} {μ : Measure α} {f : α} {r : NNReal} (hfint : Integrable f μ) (hfint' : 0 (x : α), f x μ) (hf : ∀ᵐ (ω : α) μ, f ω r) :
                      eLpNorm f 1 μ 2 * μ Set.univ * r
                      theorem MeasureTheory.eLpNorm_one_le_of_le' {α : Type u_1} {m0 : MeasurableSpace α} {μ : Measure α} {f : α} {r : } (hfint : Integrable f μ) (hfint' : 0 (x : α), f x μ) (hf : ∀ᵐ (ω : α) μ, f ω r) :

                      Positivity extension for integrals.

                      This extension only proves non-negativity, strict positivity is more delicate for integration and requires more assumptions.

                      Instances For