Documentation

Mathlib.MeasureTheory.Integral.SetIntegral

Set integral #

In this file we prove some properties of ∫ x in s, f x ∂μ. Recall that this notation is defined as ∫ x, f x ∂(μ.restrict s). In integral_indicator we prove that for a measurable function f and a measurable set s this definition coincides with another natural definition: ∫ x, indicator s f x ∂μ = ∫ x in s, f x ∂μ, where indicator s f x is equal to f x for x ∈ s and is zero otherwise.

Since ∫ x in s, f x ∂μ is a notation, one can rewrite or apply any theorem about ∫ x, f x ∂μ directly. In this file we prove some theorems about dependence of ∫ x in s, f x ∂μ on s, e.g. integral_union, integral_empty, integral_univ.

We use the property IntegrableOn f s μ := Integrable f (μ.restrict s), defined in MeasureTheory.IntegrableOn. We also defined in that same file a predicate IntegrableAtFilter (f : X → E) (l : Filter X) (μ : Measure X) saying that f is integrable at some set s ∈ l.

Finally, we prove a version of the Fundamental theorem of calculus for set integral, see Filter.Tendsto.integral_sub_linear_isLittleO_ae and its corollaries. Namely, consider a measurably generated filter l, a measure μ finite at this filter, and a function f that has a finite limit c at l ⊓ μ.ae. Then ∫ x in s, f x ∂μ = μ s • c + o(μ s) as s tends to l.smallSets, i.e. for any ε>0 there exists t ∈ l such that ‖∫ x in s, f x ∂μ - μ s • c‖ ≤ ε * μ s whenever s ⊆ t. We also formulate a version of this theorem for a locally finite measure μ and a function f continuous at a point a.

Notation #

We provide the following notations for expressing the integral of a function on a set :

Note that the set notations are defined in the file Mathlib/MeasureTheory/Integral/Bochner.lean, but we reference them here because all theorems about set integrals are in this file.

theorem MeasureTheory.set_integral_congr_ae₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {g : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasureTheory.NullMeasurableSet s μ) (h : ∀ᵐ (x : X) ∂μ, x sf x = g x) :
∫ (x : X) in s, f xμ = ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_congr_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {g : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasurableSet s) (h : ∀ᵐ (x : X) ∂μ, x sf x = g x) :
∫ (x : X) in s, f xμ = ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_congr₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {g : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasureTheory.NullMeasurableSet s μ) (h : Set.EqOn f g s) :
∫ (x : X) in s, f xμ = ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_congr {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {g : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasurableSet s) (h : Set.EqOn f g s) :
∫ (x : X) in s, f xμ = ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_congr_set_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (hst : s =ᶠ[MeasureTheory.Measure.ae μ] t) :
∫ (x : X) in s, f xμ = ∫ (x : X) in t, f xμ
theorem MeasureTheory.integral_union_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (hst : MeasureTheory.AEDisjoint μ s t) (ht : MeasureTheory.NullMeasurableSet t μ) (hfs : MeasureTheory.IntegrableOn f s μ) (hft : MeasureTheory.IntegrableOn f t μ) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ + ∫ (x : X) in t, f xμ
theorem MeasureTheory.integral_union {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (hst : Disjoint s t) (ht : MeasurableSet t) (hfs : MeasureTheory.IntegrableOn f s μ) (hft : MeasureTheory.IntegrableOn f t μ) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ + ∫ (x : X) in t, f xμ
theorem MeasureTheory.integral_diff {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasurableSet t) (hfs : MeasureTheory.IntegrableOn f s μ) (hts : t s) :
∫ (x : X) in s \ t, f xμ = ∫ (x : X) in s, f xμ - ∫ (x : X) in t, f xμ
theorem MeasureTheory.integral_inter_add_diff₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasureTheory.NullMeasurableSet t μ) (hfs : MeasureTheory.IntegrableOn f s μ) :
∫ (x : X) in s t, f xμ + ∫ (x : X) in s \ t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.integral_inter_add_diff {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasurableSet t) (hfs : MeasureTheory.IntegrableOn f s μ) :
∫ (x : X) in s t, f xμ + ∫ (x : X) in s \ t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.integral_finset_biUnion {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} (t : Finset ι) {s : ιSet X} (hs : it, MeasurableSet (s i)) (h's : Set.Pairwise (t) (Disjoint on s)) (hf : it, MeasureTheory.IntegrableOn f (s i) μ) :
∫ (x : X) in ⋃ i ∈ t, s i, f xμ = Finset.sum t fun (i : ι) => ∫ (x : X) in s i, f xμ
theorem MeasureTheory.integral_fintype_iUnion {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Fintype ι] {s : ιSet X} (hs : ∀ (i : ι), MeasurableSet (s i)) (h's : Pairwise (Disjoint on s)) (hf : ∀ (i : ι), MeasureTheory.IntegrableOn f (s i) μ) :
∫ (x : X) in ⋃ (i : ι), s i, f xμ = Finset.sum Finset.univ fun (i : ι) => ∫ (x : X) in s i, f xμ
theorem MeasureTheory.integral_empty {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} :
∫ (x : X) in , f xμ = 0
theorem MeasureTheory.integral_univ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} :
∫ (x : X) in Set.univ, f xμ = ∫ (x : X), f xμ
theorem MeasureTheory.integral_add_compl₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasureTheory.NullMeasurableSet s μ) (hfi : MeasureTheory.Integrable f μ) :
∫ (x : X) in s, f xμ + ∫ (x : X) in s, f xμ = ∫ (x : X), f xμ
theorem MeasureTheory.integral_add_compl {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasurableSet s) (hfi : MeasureTheory.Integrable f μ) :
∫ (x : X) in s, f xμ + ∫ (x : X) in s, f xμ = ∫ (x : X), f xμ
theorem MeasureTheory.integral_indicator {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} (hs : MeasurableSet s) :
∫ (x : X), Set.indicator s f xμ = ∫ (x : X) in s, f xμ

For a function f and a measurable set s, the integral of indicator s f over the whole space is equal to ∫ x in s, f x ∂μ defined as ∫ x, f x ∂(μ.restrict s).

theorem MeasureTheory.set_integral_indicator {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasurableSet t) :
∫ (x : X) in s, Set.indicator t f xμ = ∫ (x : X) in s t, f xμ
theorem MeasureTheory.ofReal_set_integral_one_of_measure_ne_top {X : Type u_5} {m : MeasurableSpace X} {μ : MeasureTheory.Measure X} {s : Set X} (hs : μ s ) :
ENNReal.ofReal (∫ (x : X) in s, 1μ) = μ s
theorem MeasureTheory.ofReal_set_integral_one {X : Type u_5} :
∀ {x : MeasurableSpace X} (μ : MeasureTheory.Measure X) [inst : MeasureTheory.IsFiniteMeasure μ] (s : Set X), ENNReal.ofReal (∫ (x : X) in s, 1μ) = μ s
theorem MeasureTheory.integral_piecewise {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {g : XE} {s : Set X} {μ : MeasureTheory.Measure X} [DecidablePred fun (x : X) => x s] (hs : MeasurableSet s) (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) :
∫ (x : X), Set.piecewise s f g xμ = ∫ (x : X) in s, f xμ + ∫ (x : X) in s, g xμ
theorem MeasureTheory.tendsto_set_integral_of_monotone {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Countable ι] [SemilatticeSup ι] {s : ιSet X} (hsm : ∀ (i : ι), MeasurableSet (s i)) (h_mono : Monotone s) (hfi : MeasureTheory.IntegrableOn f (⋃ (n : ι), s n) μ) :
Filter.Tendsto (fun (i : ι) => ∫ (x : X) in s i, f xμ) Filter.atTop (nhds (∫ (x : X) in ⋃ (n : ι), s n, f xμ))
theorem MeasureTheory.hasSum_integral_iUnion_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Countable ι] {s : ιSet X} (hm : ∀ (i : ι), MeasureTheory.NullMeasurableSet (s i) μ) (hd : Pairwise (MeasureTheory.AEDisjoint μ on s)) (hfi : MeasureTheory.IntegrableOn f (⋃ (i : ι), s i) μ) :
HasSum (fun (n : ι) => ∫ (x : X) in s n, f xμ) (∫ (x : X) in ⋃ (n : ι), s n, f xμ)
theorem MeasureTheory.hasSum_integral_iUnion {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Countable ι] {s : ιSet X} (hm : ∀ (i : ι), MeasurableSet (s i)) (hd : Pairwise (Disjoint on s)) (hfi : MeasureTheory.IntegrableOn f (⋃ (i : ι), s i) μ) :
HasSum (fun (n : ι) => ∫ (x : X) in s n, f xμ) (∫ (x : X) in ⋃ (n : ι), s n, f xμ)
theorem MeasureTheory.integral_iUnion {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Countable ι] {s : ιSet X} (hm : ∀ (i : ι), MeasurableSet (s i)) (hd : Pairwise (Disjoint on s)) (hfi : MeasureTheory.IntegrableOn f (⋃ (i : ι), s i) μ) :
∫ (x : X) in ⋃ (n : ι), s n, f xμ = ∑' (n : ι), ∫ (x : X) in s n, f xμ
theorem MeasureTheory.integral_iUnion_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} {ι : Type u_5} [Countable ι] {s : ιSet X} (hm : ∀ (i : ι), MeasureTheory.NullMeasurableSet (s i) μ) (hd : Pairwise (MeasureTheory.AEDisjoint μ on s)) (hfi : MeasureTheory.IntegrableOn f (⋃ (i : ι), s i) μ) :
∫ (x : X) in ⋃ (n : ι), s n, f xμ = ∑' (n : ι), ∫ (x : X) in s n, f xμ
theorem MeasureTheory.set_integral_eq_zero_of_ae_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {t : Set X} {μ : MeasureTheory.Measure X} (ht_eq : ∀ᵐ (x : X) ∂μ, x tf x = 0) :
∫ (x : X) in t, f xμ = 0
theorem MeasureTheory.set_integral_eq_zero_of_forall_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {t : Set X} {μ : MeasureTheory.Measure X} (ht_eq : xt, f x = 0) :
∫ (x : X) in t, f xμ = 0
theorem MeasureTheory.integral_union_eq_left_of_ae_aux {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht_eq : ∀ᵐ (x : X) ∂μ.restrict t, f x = 0) (haux : MeasureTheory.StronglyMeasurable f) (H : MeasureTheory.IntegrableOn f (s t) μ) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.integral_union_eq_left_of_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht_eq : ∀ᵐ (x : X) ∂μ.restrict t, f x = 0) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.integral_union_eq_left_of_forall₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} {f : XE} (ht : MeasureTheory.NullMeasurableSet t μ) (ht_eq : xt, f x = 0) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.integral_union_eq_left_of_forall {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} {f : XE} (ht : MeasurableSet t) (ht_eq : xt, f x = 0) :
∫ (x : X) in s t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_eq_of_subset_of_ae_diff_eq_zero_aux {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (hts : s t) (h't : ∀ᵐ (x : X) ∂μ, x t \ sf x = 0) (haux : MeasureTheory.StronglyMeasurable f) (h'aux : MeasureTheory.IntegrableOn f t μ) :
∫ (x : X) in t, f xμ = ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_eq_of_subset_of_ae_diff_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasureTheory.NullMeasurableSet t μ) (hts : s t) (h't : ∀ᵐ (x : X) ∂μ, x t \ sf x = 0) :
∫ (x : X) in t, f xμ = ∫ (x : X) in s, f xμ

If a function vanishes almost everywhere on t \ s with s ⊆ t, then its integrals on s and t coincide if t is null-measurable.

theorem MeasureTheory.set_integral_eq_of_subset_of_forall_diff_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} (ht : MeasurableSet t) (hts : s t) (h't : xt \ s, f x = 0) :
∫ (x : X) in t, f xμ = ∫ (x : X) in s, f xμ

If a function vanishes on t \ s with s ⊆ t, then its integrals on s and t coincide if t is measurable.

theorem MeasureTheory.set_integral_eq_integral_of_ae_compl_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} (h : ∀ᵐ (x : X) ∂μ, xsf x = 0) :
∫ (x : X) in s, f xμ = ∫ (x : X), f xμ

If a function vanishes almost everywhere on sᶜ, then its integral on s coincides with its integral on the whole space.

theorem MeasureTheory.set_integral_eq_integral_of_forall_compl_eq_zero {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} (h : xs, f x = 0) :
∫ (x : X) in s, f xμ = ∫ (x : X), f xμ

If a function vanishes on sᶜ, then its integral on s coincides with its integral on the whole space.

theorem MeasureTheory.set_integral_neg_eq_set_integral_nonpos {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} [LinearOrder E] {f : XE} (hf : MeasureTheory.AEStronglyMeasurable f μ) :
∫ (x : X) in {x : X | f x < 0}, f xμ = ∫ (x : X) in {x : X | f x 0}, f xμ
theorem MeasureTheory.integral_norm_eq_pos_sub_neg {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} (hfi : MeasureTheory.Integrable f μ) :
∫ (x : X), f xμ = ∫ (x : X) in {x : X | 0 f x}, f xμ - ∫ (x : X) in {x : X | f x 0}, f xμ
theorem MeasureTheory.set_integral_const {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {s : Set X} {μ : MeasureTheory.Measure X} [CompleteSpace E] (c : E) :
∫ (x : X) in s, cμ = (μ s).toReal c
@[simp]
theorem MeasureTheory.integral_indicator_const {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} [CompleteSpace E] (e : E) ⦃s : Set X (s_meas : MeasurableSet s) :
∫ (x : X), Set.indicator s (fun (x : X) => e) xμ = (μ s).toReal e
@[simp]
theorem MeasureTheory.integral_indicator_one {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} ⦃s : Set X (hs : MeasurableSet s) :
∫ (x : X), Set.indicator s 1 xμ = (μ s).toReal
theorem MeasureTheory.set_integral_indicatorConstLp {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {s : Set X} {t : Set X} {μ : MeasureTheory.Measure X} [CompleteSpace E] {p : ENNReal} (hs : MeasurableSet s) (ht : MeasurableSet t) (hμt : μ t ) (e : E) :
∫ (x : X) in s, (MeasureTheory.indicatorConstLp p ht hμt e) xμ = (μ (t s)).toReal e
theorem MeasureTheory.integral_indicatorConstLp {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {t : Set X} {μ : MeasureTheory.Measure X} [CompleteSpace E] {p : ENNReal} (ht : MeasurableSet t) (hμt : μ t ) (e : E) :
∫ (x : X), (MeasureTheory.indicatorConstLp p ht hμt e) xμ = (μ t).toReal e
theorem MeasureTheory.set_integral_map {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} {Y : Type u_5} [MeasurableSpace Y] {g : XY} {f : YE} {s : Set Y} (hs : MeasurableSet s) (hf : MeasureTheory.AEStronglyMeasurable f (MeasureTheory.Measure.map g μ)) (hg : AEMeasurable g μ) :
∫ (y : Y) in s, f yMeasureTheory.Measure.map g μ = ∫ (x : X) in g ⁻¹' s, f (g x)μ
theorem MeasurableEmbedding.set_integral_map {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} {Y : Type u_5} :
∀ {x : MeasurableSpace Y} {f : XY}, MeasurableEmbedding f∀ (g : YE) (s : Set Y), ∫ (y : Y) in s, g yMeasureTheory.Measure.map f μ = ∫ (x : X) in f ⁻¹' s, g (f x)μ
theorem ClosedEmbedding.set_integral_map {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} [TopologicalSpace X] [BorelSpace X] {Y : Type u_5} [MeasurableSpace Y] [TopologicalSpace Y] [BorelSpace Y] {g : XY} {f : YE} (s : Set Y) (hg : ClosedEmbedding g) :
∫ (y : Y) in s, f yMeasureTheory.Measure.map g μ = ∫ (x : X) in g ⁻¹' s, f (g x)μ
theorem MeasureTheory.MeasurePreserving.set_integral_preimage_emb {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} {Y : Type u_5} :
∀ {x : MeasurableSpace Y} {f : XY} {ν : MeasureTheory.Measure Y}, MeasureTheory.MeasurePreserving f μ νMeasurableEmbedding f∀ (g : YE) (s : Set Y), ∫ (x : X) in f ⁻¹' s, g (f x)μ = ∫ (y : Y) in s, g yν
theorem MeasureTheory.MeasurePreserving.set_integral_image_emb {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} {Y : Type u_5} :
∀ {x : MeasurableSpace Y} {f : XY} {ν : MeasureTheory.Measure Y}, MeasureTheory.MeasurePreserving f μ νMeasurableEmbedding f∀ (g : YE) (s : Set X), ∫ (y : Y) in f '' s, g yν = ∫ (x : X) in s, g (f x)μ
theorem MeasureTheory.set_integral_map_equiv {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {μ : MeasureTheory.Measure X} {Y : Type u_5} [MeasurableSpace Y] (e : X ≃ᵐ Y) (f : YE) (s : Set Y) :
∫ (y : Y) in s, f yMeasureTheory.Measure.map (e) μ = ∫ (x : X) in e ⁻¹' s, f (e x)μ
theorem MeasureTheory.norm_set_integral_le_of_norm_le_const_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} {C : } (hs : μ s < ) (hC : ∀ᵐ (x : X) ∂μ.restrict s, f x C) :
∫ (x : X) in s, f xμ C * (μ s).toReal
theorem MeasureTheory.norm_set_integral_le_of_norm_le_const_ae' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} {C : } (hs : μ s < ) (hC : ∀ᵐ (x : X) ∂μ, x sf x C) (hfm : MeasureTheory.AEStronglyMeasurable f (μ.restrict s)) :
∫ (x : X) in s, f xμ C * (μ s).toReal
theorem MeasureTheory.norm_set_integral_le_of_norm_le_const_ae'' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} {C : } (hs : μ s < ) (hsm : MeasurableSet s) (hC : ∀ᵐ (x : X) ∂μ, x sf x C) :
∫ (x : X) in s, f xμ C * (μ s).toReal
theorem MeasureTheory.norm_set_integral_le_of_norm_le_const {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} {C : } (hs : μ s < ) (hC : xs, f x C) (hfm : MeasureTheory.AEStronglyMeasurable f (μ.restrict s)) :
∫ (x : X) in s, f xμ C * (μ s).toReal
theorem MeasureTheory.norm_set_integral_le_of_norm_le_const' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {s : Set X} {μ : MeasureTheory.Measure X} {C : } (hs : μ s < ) (hsm : MeasurableSet s) (hC : xs, f x C) :
∫ (x : X) in s, f xμ C * (μ s).toReal
theorem MeasureTheory.set_integral_eq_zero_iff_of_nonneg_ae {X : Type u_1} [MeasurableSpace X] {s : Set X} {μ : MeasureTheory.Measure X} {f : X} (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] f) (hfi : MeasureTheory.IntegrableOn f s μ) :
∫ (x : X) in s, f xμ = 0 f =ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] 0
theorem MeasureTheory.set_integral_pos_iff_support_of_nonneg_ae {X : Type u_1} [MeasurableSpace X] {s : Set X} {μ : MeasureTheory.Measure X} {f : X} (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] f) (hfi : MeasureTheory.IntegrableOn f s μ) :
0 < ∫ (x : X) in s, f xμ 0 < μ (Function.support f s)
theorem MeasureTheory.set_integral_gt_gt {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {R : } {f : X} (hR : 0 R) (hfm : Measurable f) (hfint : MeasureTheory.IntegrableOn f {x : X | R < f x} μ) (hμ : μ {x : X | R < f x} 0) :
(μ {x : X | R < f x}).toReal * R < ∫ (x : X) in {x : X | R < f x}, f xμ
theorem MeasureTheory.set_integral_trim {E : Type u_3} [NormedAddCommGroup E] [NormedSpace E] {X : Type u_5} {m : MeasurableSpace X} {m0 : MeasurableSpace X} {μ : MeasureTheory.Measure X} (hm : m m0) {f : XE} (hf_meas : MeasureTheory.StronglyMeasurable f) {s : Set X} (hs : MeasurableSet s) :
∫ (x : X) in s, f xμ = ∫ (x : X) in s, f xμ.trim hm

Lemmas about adding and removing interval boundaries #

The primed lemmas take explicit arguments about the endpoint having zero measure, while the unprimed ones use [NoAtoms μ].

theorem MeasureTheory.integral_Icc_eq_integral_Ioc' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} (hx : μ {x} = 0) :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ioc x y, f tμ
theorem MeasureTheory.integral_Icc_eq_integral_Ico' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} (hy : μ {y} = 0) :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ico x y, f tμ
theorem MeasureTheory.integral_Ioc_eq_integral_Ioo' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} (hy : μ {y} = 0) :
∫ (t : X) in Set.Ioc x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Ico_eq_integral_Ioo' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} (hx : μ {x} = 0) :
∫ (t : X) in Set.Ico x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Icc_eq_integral_Ioo' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} (hx : μ {x} = 0) (hy : μ {y} = 0) :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Iic_eq_integral_Iio' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} (hx : μ {x} = 0) :
∫ (t : X) in Set.Iic x, f tμ = ∫ (t : X) in Set.Iio x, f tμ
theorem MeasureTheory.integral_Ici_eq_integral_Ioi' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} (hx : μ {x} = 0) :
∫ (t : X) in Set.Ici x, f tμ = ∫ (t : X) in Set.Ioi x, f tμ
theorem MeasureTheory.integral_Icc_eq_integral_Ioc {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ioc x y, f tμ
theorem MeasureTheory.integral_Icc_eq_integral_Ico {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ico x y, f tμ
theorem MeasureTheory.integral_Ioc_eq_integral_Ioo {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Ioc x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Ico_eq_integral_Ioo {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Ico x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Icc_eq_integral_Ioo {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} {y : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Icc x y, f tμ = ∫ (t : X) in Set.Ioo x y, f tμ
theorem MeasureTheory.integral_Iic_eq_integral_Iio {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Iic x, f tμ = ∫ (t : X) in Set.Iio x, f tμ
theorem MeasureTheory.integral_Ici_eq_integral_Ioi {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] [NormedSpace E] {f : XE} {μ : MeasureTheory.Measure X} [PartialOrder X] {x : X} [MeasureTheory.NoAtoms μ] :
∫ (t : X) in Set.Ici x, f tμ = ∫ (t : X) in Set.Ioi x, f tμ
theorem MeasureTheory.set_integral_mono_ae_restrict {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {g : X} {s : Set X} (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) (h : f ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] g) :
∫ (x : X) in s, f xμ ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_mono_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {g : X} {s : Set X} (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) (h : f ≤ᶠ[MeasureTheory.Measure.ae μ] g) :
∫ (x : X) in s, f xμ ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_mono_on {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {g : X} {s : Set X} (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) (hs : MeasurableSet s) (h : xs, f x g x) :
∫ (x : X) in s, f xμ ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_mono_on_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {g : X} {s : Set X} (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) (hs : MeasurableSet s) (h : ∀ᵐ (x : X) ∂μ, x sf x g x) :
∫ (x : X) in s, f xμ ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_mono {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {g : X} {s : Set X} (hf : MeasureTheory.IntegrableOn f s μ) (hg : MeasureTheory.IntegrableOn g s μ) (h : f g) :
∫ (x : X) in s, f xμ ∫ (x : X) in s, g xμ
theorem MeasureTheory.set_integral_mono_set {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} {t : Set X} (hfi : MeasureTheory.IntegrableOn f t μ) (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict t)] f) (hst : s ≤ᶠ[MeasureTheory.Measure.ae μ] t) :
∫ (x : X) in s, f xμ ∫ (x : X) in t, f xμ
theorem MeasureTheory.set_integral_le_integral {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hfi : MeasureTheory.Integrable f μ) (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae μ] f) :
∫ (x : X) in s, f xμ ∫ (x : X), f xμ
theorem MeasureTheory.set_integral_ge_of_const_le {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} {c : } (hs : MeasurableSet s) (hμs : μ s ) (hf : xs, c f x) (hfint : MeasureTheory.IntegrableOn (fun (x : X) => f x) s μ) :
c * (μ s).toReal ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_nonneg_of_ae_restrict {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] f) :
0 ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_nonneg_of_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hf : 0 ≤ᶠ[MeasureTheory.Measure.ae μ] f) :
0 ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_nonneg {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : xs, 0 f x) :
0 ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_nonneg_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : ∀ᵐ (x : X) ∂μ, x s0 f x) :
0 ∫ (x : X) in s, f xμ
theorem MeasureTheory.set_integral_le_nonneg {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : MeasureTheory.StronglyMeasurable f) (hfi : MeasureTheory.Integrable f μ) :
∫ (x : X) in s, f xμ ∫ (x : X) in {y : X | 0 f y}, f xμ
theorem MeasureTheory.set_integral_nonpos_of_ae_restrict {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hf : f ≤ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] 0) :
∫ (x : X) in s, f xμ 0
theorem MeasureTheory.set_integral_nonpos_of_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hf : f ≤ᶠ[MeasureTheory.Measure.ae μ] 0) :
∫ (x : X) in s, f xμ 0
theorem MeasureTheory.set_integral_nonpos_ae {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : ∀ᵐ (x : X) ∂μ, x sf x 0) :
∫ (x : X) in s, f xμ 0
theorem MeasureTheory.set_integral_nonpos {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : xs, f x 0) :
∫ (x : X) in s, f xμ 0
theorem MeasureTheory.set_integral_nonpos_le {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : MeasurableSet s) (hf : MeasureTheory.StronglyMeasurable f) (hfi : MeasureTheory.Integrable f μ) :
∫ (x : X) in {y : X | f y 0}, f xμ ∫ (x : X) in s, f xμ
theorem MeasureTheory.Integrable.measure_le_integral {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} (f_int : MeasureTheory.Integrable f μ) (f_nonneg : 0 ≤ᶠ[MeasureTheory.Measure.ae μ] f) {s : Set X} (hs : xs, 1 f x) :
μ s ENNReal.ofReal (∫ (x : X), f xμ)
theorem MeasureTheory.integral_le_measure {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {f : X} {s : Set X} (hs : xs, f x 1) (h's : xs, f x 0) :
ENNReal.ofReal (∫ (x : X), f xμ) μ s
theorem MeasureTheory.integrableOn_iUnion_of_summable_integral_norm {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [Countable ι] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] {f : XE} {s : ιSet X} (hs : ∀ (i : ι), MeasurableSet (s i)) (hi : ∀ (i : ι), MeasureTheory.IntegrableOn f (s i) μ) (h : Summable fun (i : ι) => ∫ (x : X) in s i, f xμ) :
theorem MeasureTheory.integrableOn_iUnion_of_summable_norm_restrict {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [Countable ι] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [TopologicalSpace X] [BorelSpace X] [TopologicalSpace.MetrizableSpace X] [MeasureTheory.IsLocallyFiniteMeasure μ] {f : C(X, E)} {s : ιTopologicalSpace.Compacts X} (hf : Summable fun (i : ι) => ContinuousMap.restrict ((s i)) f * (μ (s i)).toReal) :
MeasureTheory.IntegrableOn (f) (⋃ (i : ι), (s i)) μ

If s is a countable family of compact sets, f is a continuous function, and the sequence ‖f.restrict (s i)‖ * μ (s i) is summable, then f is integrable on the union of the s i.

theorem MeasureTheory.integrable_of_summable_norm_restrict {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [Countable ι] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [TopologicalSpace X] [BorelSpace X] [TopologicalSpace.MetrizableSpace X] [MeasureTheory.IsLocallyFiniteMeasure μ] {f : C(X, E)} {s : ιTopologicalSpace.Compacts X} (hf : Summable fun (i : ι) => ContinuousMap.restrict ((s i)) f * (μ (s i)).toReal) (hs : ⋃ (i : ι), (s i) = Set.univ) :

If s is a countable family of compact sets covering X, f is a continuous function, and the sequence ‖f.restrict (s i)‖ * μ (s i) is summable, then f is integrable.

Continuity of the set integral #

We prove that for any set s, the function fun f : X →₁[μ] E => ∫ x in s, f x ∂μ is continuous.

theorem MeasureTheory.Lp_toLp_restrict_add {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] {p : ENNReal} {μ : MeasureTheory.Measure X} (f : (MeasureTheory.Lp E p μ)) (g : (MeasureTheory.Lp E p μ)) (s : Set X) :

For f : Lp E p μ, we can define an element of Lp E p (μ.restrict s) by (Lp.memℒp f).restrict s).toLp f. This map is additive.

theorem MeasureTheory.Lp_toLp_restrict_smul {X : Type u_1} {F : Type u_4} [MeasurableSpace X] {𝕜 : Type u_5} [NormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] {p : ENNReal} {μ : MeasureTheory.Measure X} (c : 𝕜) (f : (MeasureTheory.Lp F p μ)) (s : Set X) :

For f : Lp E p μ, we can define an element of Lp E p (μ.restrict s) by (Lp.memℒp f).restrict s).toLp f. This map commutes with scalar multiplication.

For f : Lp E p μ, we can define an element of Lp E p (μ.restrict s) by (Lp.memℒp f).restrict s).toLp f. This map is non-expansive.

def MeasureTheory.LpToLpRestrictCLM (X : Type u_1) (F : Type u_4) [MeasurableSpace X] (𝕜 : Type u_5) [NormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] (μ : MeasureTheory.Measure X) (p : ENNReal) [hp : Fact (1 p)] (s : Set X) :
(MeasureTheory.Lp F p μ) →L[𝕜] (MeasureTheory.Lp F p (μ.restrict s))

Continuous linear map sending a function of Lp F p μ to the same function in Lp F p (μ.restrict s).

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    theorem MeasureTheory.LpToLpRestrictCLM_coeFn {X : Type u_1} {F : Type u_4} [MeasurableSpace X] (𝕜 : Type u_5) [NormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] {p : ENNReal} {μ : MeasureTheory.Measure X} [Fact (1 p)] (s : Set X) (f : (MeasureTheory.Lp F p μ)) :
    ((MeasureTheory.LpToLpRestrictCLM X F 𝕜 μ p s) f) =ᶠ[MeasureTheory.Measure.ae (μ.restrict s)] f
    theorem MeasureTheory.continuous_set_integral {X : Type u_1} {E : Type u_3} [MeasurableSpace X] [NormedAddCommGroup E] {μ : MeasureTheory.Measure X} [NormedSpace E] (s : Set X) :
    Continuous fun (f : (MeasureTheory.Lp E 1 μ)) => ∫ (x : X) in s, f xμ

    Fundamental theorem of calculus for set integrals

    theorem Filter.Tendsto.integral_sub_linear_isLittleO_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {μ : MeasureTheory.Measure X} {l : Filter X} [Filter.IsMeasurablyGenerated l] {f : XE} {b : E} (h : Filter.Tendsto f (l MeasureTheory.Measure.ae μ) (nhds b)) (hfm : StronglyMeasurableAtFilter f l μ) (hμ : MeasureTheory.Measure.FiniteAtFilter μ l) {s : ιSet X} {li : Filter ι} (hs : Filter.Tendsto s li (Filter.smallSets l)) (m : optParam (ι) fun (i : ι) => (μ (s i)).toReal) (hsμ : autoParam ((fun (i : ι) => (μ (s i)).toReal) =ᶠ[li] m) _auto✝) :
    (fun (i : ι) => ∫ (x : X) in s i, f xμ - m i b) =o[li] m

    Fundamental theorem of calculus for set integrals: if μ is a measure that is finite at a filter l and f is a measurable function that has a finite limit b at l ⊓ μ.ae, then ∫ x in s i, f x ∂μ = μ (s i) • b + o(μ (s i)) at a filter li provided that s i tends to l.smallSets along li. Since μ (s i) is an ℝ≥0∞ number, we use (μ (s i)).toReal in the actual statement.

    Often there is a good formula for (μ (s i)).toReal, so the formalization can take an optional argument m with this formula and a proof of (fun i => (μ (s i)).toReal) =ᶠ[li] m. Without these arguments, m i = (μ (s i)).toReal is used in the output.

    theorem ContinuousWithinAt.integral_sub_linear_isLittleO_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] [TopologicalSpace X] [OpensMeasurableSpace X] {μ : MeasureTheory.Measure X} [MeasureTheory.IsLocallyFiniteMeasure μ] {x : X} {t : Set X} {f : XE} (hx : ContinuousWithinAt f t x) (ht : MeasurableSet t) (hfm : StronglyMeasurableAtFilter f (nhdsWithin x t) μ) {s : ιSet X} {li : Filter ι} (hs : Filter.Tendsto s li (Filter.smallSets (nhdsWithin x t))) (m : optParam (ι) fun (i : ι) => (μ (s i)).toReal) (hsμ : autoParam ((fun (i : ι) => (μ (s i)).toReal) =ᶠ[li] m) _auto✝) :
    (fun (i : ι) => ∫ (x : X) in s i, f xμ - m i f x) =o[li] m

    Fundamental theorem of calculus for set integrals, nhdsWithin version: if μ is a locally finite measure and f is an almost everywhere measurable function that is continuous at a point a within a measurable set t, then ∫ x in s i, f x ∂μ = μ (s i) • f a + o(μ (s i)) at a filter li provided that s i tends to (𝓝[t] a).smallSets along li. Since μ (s i) is an ℝ≥0∞ number, we use (μ (s i)).toReal in the actual statement.

    Often there is a good formula for (μ (s i)).toReal, so the formalization can take an optional argument m with this formula and a proof of (fun i => (μ (s i)).toReal) =ᶠ[li] m. Without these arguments, m i = (μ (s i)).toReal is used in the output.

    theorem ContinuousAt.integral_sub_linear_isLittleO_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] [TopologicalSpace X] [OpensMeasurableSpace X] {μ : MeasureTheory.Measure X} [MeasureTheory.IsLocallyFiniteMeasure μ] {x : X} {f : XE} (hx : ContinuousAt f x) (hfm : StronglyMeasurableAtFilter f (nhds x) μ) {s : ιSet X} {li : Filter ι} (hs : Filter.Tendsto s li (Filter.smallSets (nhds x))) (m : optParam (ι) fun (i : ι) => (μ (s i)).toReal) (hsμ : autoParam ((fun (i : ι) => (μ (s i)).toReal) =ᶠ[li] m) _auto✝) :
    (fun (i : ι) => ∫ (x : X) in s i, f xμ - m i f x) =o[li] m

    Fundamental theorem of calculus for set integrals, nhds version: if μ is a locally finite measure and f is an almost everywhere measurable function that is continuous at a point a, then ∫ x in s i, f x ∂μ = μ (s i) • f a + o(μ (s i)) at li provided that s tends to (𝓝 a).smallSets along li. Since μ (s i) is an ℝ≥0∞ number, we use (μ (s i)).toReal in the actual statement.

    Often there is a good formula for (μ (s i)).toReal, so the formalization can take an optional argument m with this formula and a proof of (fun i => (μ (s i)).toReal) =ᶠ[li] m. Without these arguments, m i = (μ (s i)).toReal is used in the output.

    theorem ContinuousOn.integral_sub_linear_isLittleO_ae {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {ι : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] [TopologicalSpace X] [OpensMeasurableSpace X] [SecondCountableTopologyEither X E] {μ : MeasureTheory.Measure X} [MeasureTheory.IsLocallyFiniteMeasure μ] {x : X} {t : Set X} {f : XE} (hft : ContinuousOn f t) (hx : x t) (ht : MeasurableSet t) {s : ιSet X} {li : Filter ι} (hs : Filter.Tendsto s li (Filter.smallSets (nhdsWithin x t))) (m : optParam (ι) fun (i : ι) => (μ (s i)).toReal) (hsμ : autoParam ((fun (i : ι) => (μ (s i)).toReal) =ᶠ[li] m) _auto✝) :
    (fun (i : ι) => ∫ (x : X) in s i, f xμ - m i f x) =o[li] m

    Fundamental theorem of calculus for set integrals, nhdsWithin version: if μ is a locally finite measure, f is continuous on a measurable set t, and a ∈ t, then ∫ x in (s i), f x ∂μ = μ (s i) • f a + o(μ (s i)) at li provided that s i tends to (𝓝[t] a).smallSets along li. Since μ (s i) is an ℝ≥0∞ number, we use (μ (s i)).toReal in the actual statement.

    Often there is a good formula for (μ (s i)).toReal, so the formalization can take an optional argument m with this formula and a proof of (fun i => (μ (s i)).toReal) =ᶠ[li] m. Without these arguments, m i = (μ (s i)).toReal is used in the output.

    Continuous linear maps composed with integration #

    The goal of this section is to prove that integration commutes with continuous linear maps. This holds for simple functions. The general result follows from the continuity of all involved operations on the space . Note that composition by a continuous linear map on is not just the composition, as we are dealing with classes of functions, but it has already been defined as ContinuousLinearMap.compLp. We take advantage of this construction here.

    theorem ContinuousLinearMap.integral_compLp {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] {p : ENNReal} [NormedSpace F] (L : E →L[𝕜] F) (φ : (MeasureTheory.Lp E p μ)) :
    ∫ (x : X), (ContinuousLinearMap.compLp L φ) xμ = ∫ (x : X), L (φ x)μ
    theorem ContinuousLinearMap.set_integral_compLp {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] {p : ENNReal} [NormedSpace F] (L : E →L[𝕜] F) (φ : (MeasureTheory.Lp E p μ)) {s : Set X} (hs : MeasurableSet s) :
    ∫ (x : X) in s, (ContinuousLinearMap.compLp L φ) xμ = ∫ (x : X) in s, L (φ x)μ
    theorem ContinuousLinearMap.continuous_integral_comp_L1 {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace F] (L : E →L[𝕜] F) :
    Continuous fun (φ : (MeasureTheory.Lp E 1 μ)) => ∫ (x : X), L (φ x)μ
    theorem ContinuousLinearMap.integral_comp_comm {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace F] [CompleteSpace F] [NormedSpace E] [CompleteSpace E] (L : E →L[𝕜] F) {φ : XE} (φ_int : MeasureTheory.Integrable φ μ) :
    ∫ (x : X), L (φ x)μ = L (∫ (x : X), φ xμ)
    theorem ContinuousLinearMap.integral_apply {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedSpace E] {H : Type u_6} [NormedAddCommGroup H] [NormedSpace 𝕜 H] {φ : XH →L[𝕜] E} (φ_int : MeasureTheory.Integrable φ μ) (v : H) :
    (∫ (x : X), φ xμ) v = ∫ (x : X), (φ x) vμ
    theorem ContinuousMultilinearMap.integral_apply {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedSpace E] {ι : Type u_6} [Fintype ι] {M : ιType u_7} [(i : ι) → NormedAddCommGroup (M i)] [(i : ι) → NormedSpace 𝕜 (M i)] {φ : XContinuousMultilinearMap 𝕜 M E} (φ_int : MeasureTheory.Integrable φ μ) (m : (i : ι) → M i) :
    (∫ (x : X), φ xμ) m = ∫ (x : X), (φ x) mμ
    theorem ContinuousLinearMap.integral_comp_comm' {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace F] [CompleteSpace F] [NormedSpace E] [CompleteSpace E] (L : E →L[𝕜] F) {K : NNReal} (hL : AntilipschitzWith K L) (φ : XE) :
    ∫ (x : X), L (φ x)μ = L (∫ (x : X), φ xμ)
    theorem ContinuousLinearMap.integral_comp_L1_comm {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace F] [CompleteSpace F] [NormedSpace E] [CompleteSpace E] (L : E →L[𝕜] F) (φ : (MeasureTheory.Lp E 1 μ)) :
    ∫ (x : X), L (φ x)μ = L (∫ (x : X), φ xμ)
    theorem LinearIsometry.integral_comp_comm {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [CompleteSpace F] [NormedSpace F] [CompleteSpace E] [NormedSpace E] (L : E →ₗᵢ[𝕜] F) (φ : XE) :
    ∫ (x : X), L (φ x)μ = L (∫ (x : X), φ xμ)
    theorem ContinuousLinearEquiv.integral_comp_comm {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace F] [NormedSpace E] (L : E ≃L[𝕜] F) (φ : XE) :
    ∫ (x : X), L (φ x)μ = L (∫ (x : X), φ xμ)
    theorem integral_ofReal {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X} :
    ∫ (x : X), (f x)μ = (∫ (x : X), f xμ)
    theorem integral_re {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} (hf : MeasureTheory.Integrable f μ) :
    ∫ (x : X), RCLike.re (f x)μ = RCLike.re (∫ (x : X), f xμ)
    theorem integral_im {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} (hf : MeasureTheory.Integrable f μ) :
    ∫ (x : X), RCLike.im (f x)μ = RCLike.im (∫ (x : X), f xμ)
    theorem integral_conj {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} :
    ∫ (x : X), (starRingEnd 𝕜) (f x)μ = (starRingEnd 𝕜) (∫ (x : X), f xμ)
    theorem integral_coe_re_add_coe_im {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} (hf : MeasureTheory.Integrable f μ) :
    ∫ (x : X), (RCLike.re (f x))μ + (∫ (x : X), (RCLike.im (f x))μ) * RCLike.I = ∫ (x : X), f xμ
    theorem integral_re_add_im {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} (hf : MeasureTheory.Integrable f μ) :
    (∫ (x : X), RCLike.re (f x)μ) + (∫ (x : X), RCLike.im (f x)μ) * RCLike.I = ∫ (x : X), f xμ
    theorem set_integral_re_add_im {X : Type u_1} [MeasurableSpace X] {μ : MeasureTheory.Measure X} {𝕜 : Type u_5} [RCLike 𝕜] {f : X𝕜} {i : Set X} (hf : MeasureTheory.IntegrableOn f i μ) :
    (∫ (x : X) in i, RCLike.re (f x)μ) + (∫ (x : X) in i, RCLike.im (f x)μ) * RCLike.I = ∫ (x : X) in i, f xμ
    theorem swap_integral {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace E] [NormedSpace F] (f : XE × F) :
    Prod.swap (∫ (x : X), f xμ) = ∫ (x : X), Prod.swap (f x)μ
    theorem fst_integral {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace E] [NormedSpace F] [CompleteSpace F] {f : XE × F} (hf : MeasureTheory.Integrable f μ) :
    (∫ (x : X), f xμ).1 = ∫ (x : X), (f x).1μ
    theorem snd_integral {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace E] [NormedSpace F] [CompleteSpace E] {f : XE × F} (hf : MeasureTheory.Integrable f μ) :
    (∫ (x : X), f xμ).2 = ∫ (x : X), (f x).2μ
    theorem integral_pair {X : Type u_1} {E : Type u_3} {F : Type u_4} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedSpace E] [NormedSpace F] [CompleteSpace E] [CompleteSpace F] {f : XE} {g : XF} (hf : MeasureTheory.Integrable f μ) (hg : MeasureTheory.Integrable g μ) :
    ∫ (x : X), (f x, g x)μ = (∫ (x : X), f xμ, ∫ (x : X), g xμ)
    theorem integral_smul_const {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] {𝕜 : Type u_6} [RCLike 𝕜] [NormedSpace 𝕜 E] [CompleteSpace E] (f : X𝕜) (c : E) :
    ∫ (x : X), f x cμ = (∫ (x : X), f xμ) c
    theorem integral_withDensity_eq_integral_smul {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] {f : XNNReal} (f_meas : Measurable f) (g : XE) :
    (∫ (x : X), g xμ.withDensity fun (x : X) => (f x)) = ∫ (x : X), f x g xμ
    theorem integral_withDensity_eq_integral_smul₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] {f : XNNReal} (hf : AEMeasurable f μ) (g : XE) :
    (∫ (x : X), g xμ.withDensity fun (x : X) => (f x)) = ∫ (x : X), f x g xμ
    theorem set_integral_withDensity_eq_set_integral_smul {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] {f : XNNReal} (f_meas : Measurable f) (g : XE) {s : Set X} (hs : MeasurableSet s) :
    (∫ (x : X) in s, g xμ.withDensity fun (x : X) => (f x)) = ∫ (x : X) in s, f x g xμ
    theorem set_integral_withDensity_eq_set_integral_smul₀ {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] {f : XNNReal} {s : Set X} (hf : AEMeasurable f (μ.restrict s)) (g : XE) (hs : MeasurableSet s) :
    (∫ (x : X) in s, g xμ.withDensity fun (x : X) => (f x)) = ∫ (x : X) in s, f x g xμ
    theorem set_integral_withDensity_eq_set_integral_smul₀' {X : Type u_1} {E : Type u_3} [MeasurableSpace X] {μ : MeasureTheory.Measure X} [NormedAddCommGroup E] [NormedSpace E] [MeasureTheory.SFinite μ] {f : XNNReal} (s : Set X) (hf : AEMeasurable f (μ.restrict s)) (g : XE) :
    (∫ (x : X) in s, g xμ.withDensity fun (x : X) => (f x)) = ∫ (x : X) in s, f x g xμ
    theorem measure_le_lintegral_thickenedIndicatorAux {X : Type u_1} [MeasurableSpace X] [PseudoEMetricSpace X] (μ : MeasureTheory.Measure X) {E : Set X} (E_mble : MeasurableSet E) (δ : ) :
    μ E ∫⁻ (x : X), thickenedIndicatorAux δ E xμ
    theorem measure_le_lintegral_thickenedIndicator {X : Type u_1} [MeasurableSpace X] [PseudoEMetricSpace X] (μ : MeasureTheory.Measure X) {E : Set X} (E_mble : MeasurableSet E) {δ : } (δ_pos : 0 < δ) :
    μ E ∫⁻ (x : X), ((thickenedIndicator δ_pos E) x)μ

    The parametric integral over a continuous function on a compact set is continuous, under mild assumptions on the topologies involved.

    theorem continuousOn_integral_bilinear_of_locally_integrable_of_compact_support {X : Type u_1} {Y : Type u_2} {E : Type u_3} {F : Type u_4} {G : Type u_5} {𝕜 : Type u_6} [TopologicalSpace X] [TopologicalSpace Y] [MeasurableSpace Y] [OpensMeasurableSpace Y] {μ : MeasureTheory.Measure Y} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedSpace 𝕜 E] (L : F →L[𝕜] G →L[𝕜] E) {f : XYG} {s : Set X} {k : Set Y} {g : YF} (hk : IsCompact k) (hf : ContinuousOn (Function.uncurry f) (s ×ˢ Set.univ)) (hfs : ∀ (p : X) (x : Y), p sxkf p x = 0) (hg : MeasureTheory.IntegrableOn g k μ) :
    ContinuousOn (fun (x : X) => ∫ (y : Y), (L (g y)) (f x y)μ) s

    Consider a parameterized integral x ↦ ∫ y, L (g y) (f x y) where L is bilinear, g is locally integrable and f is continuous and uniformly compactly supported. Then the integral depends continuously on x.

    theorem continuousOn_integral_of_compact_support {X : Type u_1} {Y : Type u_2} {E : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [MeasurableSpace Y] [OpensMeasurableSpace Y] {μ : MeasureTheory.Measure Y} [NormedAddCommGroup E] [NormedSpace E] {f : XYE} {s : Set X} {k : Set Y} [MeasureTheory.IsFiniteMeasureOnCompacts μ] (hk : IsCompact k) (hf : ContinuousOn (Function.uncurry f) (s ×ˢ Set.univ)) (hfs : ∀ (p : X) (x : Y), p sxkf p x = 0) :
    ContinuousOn (fun (x : X) => ∫ (y : Y), f x yμ) s

    Consider a parameterized integral x ↦ ∫ y, f x y where f is continuous and uniformly compactly supported. Then the integral depends continuously on x.