mathlib documentation

measure_theory.integration

Lebesgue integral for ennreal-valued functions

We define simple functions and show that each Borel measurable function on ennreal can be approximated by a sequence of simple functions.

To prove something for an arbitrary measurable function into ennreal, the theorem measurable.ennreal_induction shows that is it sufficient to show that the property holds for (multiples of) characteristic functions and is closed under addition and supremum of increasing sequences of functions.

Notation

We introduce the following notation for the lower Lebesgue integral of a function f : α → ennreal.

structure measure_theory.simple_func (α : Type u) [measurable_space α] (β : Type v) :
Type (max u v)

A function f from a measurable space to any type is called simple, if every preimage f ⁻¹' {x} is measurable, and the range is finite. This structure bundles a function with these properties.

theorem measure_theory.simple_func.coe_injective {α : Type u_1} {β : Type u_2} [measurable_space α] ⦃f g : measure_theory.simple_func α β⦄ (H : f = g) :
f = g

@[ext]
theorem measure_theory.simple_func.ext {α : Type u_1} {β : Type u_2} [measurable_space α] {f g : measure_theory.simple_func α β} (H : ∀ (a : α), f a = g a) :
f = g

theorem measure_theory.simple_func.finite_range {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) :

theorem measure_theory.simple_func.is_measurable_fiber {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) (x : β) :

def measure_theory.simple_func.range {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) :

Range of a simple function α →ₛ β as a finset β.

Equations
@[simp]
theorem measure_theory.simple_func.mem_range {α : Type u_1} {β : Type u_2} [measurable_space α] {f : measure_theory.simple_func α β} {b : β} :

theorem measure_theory.simple_func.mem_range_self {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) (x : α) :

@[simp]
theorem measure_theory.simple_func.coe_range {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) :

theorem measure_theory.simple_func.mem_range_of_measure_ne_zero {α : Type u_1} {β : Type u_2} [measurable_space α] {f : measure_theory.simple_func α β} {x : β} {μ : measure_theory.measure α} (H : μ (f ⁻¹' {x}) 0) :

theorem measure_theory.simple_func.forall_range_iff {α : Type u_1} {β : Type u_2} [measurable_space α] {f : measure_theory.simple_func α β} {p : β → Prop} :
(∀ (y : β), y f.rangep y) ∀ (x : α), p (f x)

theorem measure_theory.simple_func.exists_range_iff {α : Type u_1} {β : Type u_2} [measurable_space α] {f : measure_theory.simple_func α β} {p : β → Prop} :
(∃ (y : β) (H : y f.range), p y) ∃ (x : α), p (f x)

theorem measure_theory.simple_func.preimage_eq_empty_iff {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) (b : β) :

def measure_theory.simple_func.const (α : Type u_1) {β : Type u_2} [measurable_space α] (b : β) :

Constant function as a simple_func.

Equations
theorem measure_theory.simple_func.const_apply {α : Type u_1} {β : Type u_2} [measurable_space α] (a : α) (b : β) :

@[simp]
theorem measure_theory.simple_func.coe_const {α : Type u_1} {β : Type u_2} [measurable_space α] (b : β) :

@[simp]
theorem measure_theory.simple_func.range_const {β : Type u_2} (α : Type u_1) [measurable_space α] [nonempty α] (b : β) :

theorem measure_theory.simple_func.is_measurable_cut {α : Type u_1} {β : Type u_2} [measurable_space α] (r : α → β → Prop) (f : measure_theory.simple_func α β) (h : ∀ (b : β), is_measurable {a : α | r a b}) :
is_measurable {a : α | r a (f a)}

theorem measure_theory.simple_func.is_measurable_preimage {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) (s : set β) :

theorem measure_theory.simple_func.measurable {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] (f : measure_theory.simple_func α β) :

A simple function is measurable

theorem measure_theory.simple_func.sum_measure_preimage_singleton {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) {μ : measure_theory.measure α} (s : finset β) :
∑ (y : β) in s, μ (f ⁻¹' {y}) = μ (f ⁻¹' s)

theorem measure_theory.simple_func.sum_range_measure_preimage_singleton {α : Type u_1} {β : Type u_2} [measurable_space α] (f : measure_theory.simple_func α β) (μ : measure_theory.measure α) :
∑ (y : β) in f.range, μ (f ⁻¹' {y}) = μ set.univ

def measure_theory.simple_func.piecewise {α : Type u_1} {β : Type u_2} [measurable_space α] (s : set α) (hs : is_measurable s) (f g : measure_theory.simple_func α β) :

If-then-else as a simple_func.

Equations
@[simp]
theorem measure_theory.simple_func.coe_piecewise {α : Type u_1} {β : Type u_2} [measurable_space α] {s : set α} (hs : is_measurable s) (f g : measure_theory.simple_func α β) :

theorem measure_theory.simple_func.piecewise_apply {α : Type u_1} {β : Type u_2} [measurable_space α] {s : set α} (hs : is_measurable s) (f g : measure_theory.simple_func α β) (a : α) :

theorem measure_theory.simple_func.measurable_bind {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [measurable_space γ] (f : measure_theory.simple_func α β) (g : β → α → γ) (hg : ∀ (b : β), measurable (g b)) :
measurable (λ (a : α), g (f a) a)

def measure_theory.simple_func.bind {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : β → measure_theory.simple_func α γ) :

If f : α →ₛ β is a simple function and g : β → α →ₛ γ is a family of simple functions, then f.bind g binds the first argument of g to f. In other words, f.bind g a = g (f a) a.

Equations
@[simp]
theorem measure_theory.simple_func.bind_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : β → measure_theory.simple_func α γ) (a : α) :
(f.bind g) a = (g (f a)) a

def measure_theory.simple_func.map {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (g : β → γ) (f : measure_theory.simple_func α β) :

Given a function g : β → γ and a simple function f : α →ₛ β, f.map g return the simple function g ∘ f : α →ₛ γ

Equations
theorem measure_theory.simple_func.map_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (g : β → γ) (f : measure_theory.simple_func α β) (a : α) :

theorem measure_theory.simple_func.map_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [measurable_space α] (g : β → γ) (h : γ → δ) (f : measure_theory.simple_func α β) :

@[simp]
theorem measure_theory.simple_func.coe_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (g : β → γ) (f : measure_theory.simple_func α β) :

@[simp]
theorem measure_theory.simple_func.range_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [decidable_eq γ] (g : β → γ) (f : measure_theory.simple_func α β) :

@[simp]
theorem measure_theory.simple_func.map_const {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (g : β → γ) (b : β) :

theorem measure_theory.simple_func.map_preimage {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : β → γ) (s : set γ) :

theorem measure_theory.simple_func.map_preimage_singleton {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : β → γ) (c : γ) :

def measure_theory.simple_func.comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [measurable_space β] (f : measure_theory.simple_func β γ) (g : α → β) (hgm : measurable g) :

Composition of a simple_fun and a measurable function is a simple_func.

Equations
@[simp]
theorem measure_theory.simple_func.coe_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [measurable_space β] (f : measure_theory.simple_func β γ) {g : α → β} (hgm : measurable g) :
(f.comp g hgm) = f g

theorem measure_theory.simple_func.range_comp_subset_range {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [measurable_space β] (f : measure_theory.simple_func β γ) {g : α → β} (hgm : measurable g) :
(f.comp g hgm).range f.range

def measure_theory.simple_func.seq {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α (β → γ)) (g : measure_theory.simple_func α β) :

If f is a simple function taking values in β → γ and g is another simple function with the same domain and codomain β, then f.seq g = f a (g a).

Equations
@[simp]
theorem measure_theory.simple_func.seq_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α (β → γ)) (g : measure_theory.simple_func α β) (a : α) :
(f.seq g) a = f a (g a)

def measure_theory.simple_func.pair {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : measure_theory.simple_func α γ) :

Combine two simple functions f : α →ₛ β and g : α →ₛ β into λ a, (f a, g a).

Equations
@[simp]
theorem measure_theory.simple_func.pair_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : measure_theory.simple_func α γ) (a : α) :
(f.pair g) a = (f a, g a)

theorem measure_theory.simple_func.pair_preimage {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : measure_theory.simple_func α γ) (s : set β) (t : set γ) :

theorem measure_theory.simple_func.pair_preimage_singleton {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] (f : measure_theory.simple_func α β) (g : measure_theory.simple_func α γ) (b : β) (c : γ) :
(f.pair g) ⁻¹' {(b, c)} = f ⁻¹' {b} g ⁻¹' {c}

@[instance]
def measure_theory.simple_func.has_le {α : Type u_1} {β : Type u_2} [measurable_space α] [has_le β] :

Equations
@[simp]
theorem measure_theory.simple_func.coe_zero {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] :
0 = 0

@[simp]
theorem measure_theory.simple_func.const_zero {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] :

@[simp]
theorem measure_theory.simple_func.coe_add {α : Type u_1} {β : Type u_2} [measurable_space α] [has_add β] (f g : measure_theory.simple_func α β) :
(f + g) = f + g

@[simp]
theorem measure_theory.simple_func.coe_mul {α : Type u_1} {β : Type u_2} [measurable_space α] [has_mul β] (f g : measure_theory.simple_func α β) :
f * g = (f) * g

@[simp]
theorem measure_theory.simple_func.coe_le {α : Type u_1} {β : Type u_2} [measurable_space α] [preorder β] {f g : measure_theory.simple_func α β} :
f g f g

@[simp]
theorem measure_theory.simple_func.range_zero {α : Type u_1} {β : Type u_2} [measurable_space α] [nonempty α] [has_zero β] :
0.range = {0}

theorem measure_theory.simple_func.eq_zero_of_mem_range_zero {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {y : β} (a : y 0.range) :
y = 0

theorem measure_theory.simple_func.sup_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [has_sup β] (f g : measure_theory.simple_func α β) (a : α) :
(f g) a = f a g a

theorem measure_theory.simple_func.mul_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [has_mul β] (f g : measure_theory.simple_func α β) (a : α) :
(f * g) a = (f a) * g a

theorem measure_theory.simple_func.add_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [has_add β] (f g : measure_theory.simple_func α β) (a : α) :
(f + g) a = f a + g a

theorem measure_theory.simple_func.add_eq_map₂ {α : Type u_1} {β : Type u_2} [measurable_space α] [has_add β] (f g : measure_theory.simple_func α β) :
f + g = measure_theory.simple_func.map (λ (p : β × β), p.fst + p.snd) (f.pair g)

theorem measure_theory.simple_func.mul_eq_map₂ {α : Type u_1} {β : Type u_2} [measurable_space α] [has_mul β] (f g : measure_theory.simple_func α β) :
f * g = measure_theory.simple_func.map (λ (p : β × β), (p.fst) * p.snd) (f.pair g)

theorem measure_theory.simple_func.sup_eq_map₂ {α : Type u_1} {β : Type u_2} [measurable_space α] [has_sup β] (f g : measure_theory.simple_func α β) :
f g = measure_theory.simple_func.map (λ (p : β × β), p.fst p.snd) (f.pair g)

theorem measure_theory.simple_func.const_mul_eq_map {α : Type u_1} {β : Type u_2} [measurable_space α] [has_mul β] (f : measure_theory.simple_func α β) (b : β) :

@[instance]

Equations
@[instance]

Equations
@[simp]
theorem measure_theory.simple_func.coe_neg {α : Type u_1} {β : Type u_2} [measurable_space α] [has_neg β] (f : measure_theory.simple_func α β) :

@[instance]

Equations
@[simp]
theorem measure_theory.simple_func.coe_sub {α : Type u_1} {β : Type u_2} [measurable_space α] [add_group β] (f g : measure_theory.simple_func α β) :
(f - g) = f - g

@[instance]

Equations
@[simp]
theorem measure_theory.simple_func.coe_smul {α : Type u_1} {β : Type u_2} [measurable_space α] {K : Type u_5} [has_scalar K β] (c : K) (f : measure_theory.simple_func α β) :
(c f) = c f

theorem measure_theory.simple_func.smul_apply {α : Type u_1} {β : Type u_2} [measurable_space α] {K : Type u_5} [has_scalar K β] (k : K) (f : measure_theory.simple_func α β) (a : α) :
(k f) a = k f a

@[instance]
def measure_theory.simple_func.semimodule {α : Type u_1} {β : Type u_2} [measurable_space α] {K : Type u_5} [semiring K] [add_comm_monoid β] [semimodule K β] :

Equations
theorem measure_theory.simple_func.smul_eq_map {α : Type u_1} {β : Type u_2} [measurable_space α] {K : Type u_5} [has_scalar K β] (k : K) (f : measure_theory.simple_func α β) :

theorem measure_theory.simple_func.finset_sup_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [semilattice_sup_bot β] {f : γ → measure_theory.simple_func α β} (s : finset γ) (a : α) :
(s.sup f) a = s.sup (λ (c : γ), (f c) a)

def measure_theory.simple_func.restrict {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) (s : set α) :

Restrict a simple function f : α →ₛ β to a set s. If s is measurable, then f.restrict s a = if a ∈ s then f a else 0, otherwise f.restrict s = const α 0.

Equations
theorem measure_theory.simple_func.restrict_of_not_measurable {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {f : measure_theory.simple_func α β} {s : set α} (hs : ¬is_measurable s) :
f.restrict s = 0

@[simp]
theorem measure_theory.simple_func.coe_restrict {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) {s : set α} (hs : is_measurable s) :

@[simp]
theorem measure_theory.simple_func.restrict_univ {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) :

@[simp]
theorem measure_theory.simple_func.restrict_empty {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) :

theorem measure_theory.simple_func.map_restrict_of_zero {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [has_zero β] [has_zero γ] {g : β → γ} (hg : g 0 = 0) (f : measure_theory.simple_func α β) (s : set α) :

theorem measure_theory.simple_func.restrict_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) {s : set α} (hs : is_measurable s) (a : α) :
(f.restrict s) a = ite (a s) (f a) 0

theorem measure_theory.simple_func.restrict_preimage {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) {s : set α} (hs : is_measurable s) {t : set β} (ht : 0 t) :

theorem measure_theory.simple_func.restrict_preimage_singleton {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) {s : set α} (hs : is_measurable s) {r : β} (hr : r 0) :
(f.restrict s) ⁻¹' {r} = s f ⁻¹' {r}

theorem measure_theory.simple_func.mem_restrict_range {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {r : β} {s : set α} {f : measure_theory.simple_func α β} (hs : is_measurable s) :
r (f.restrict s).range r = 0 s set.univ r f '' s

theorem measure_theory.simple_func.mem_image_of_mem_range_restrict {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {r : β} {s : set α} {f : measure_theory.simple_func α β} (hr : r (f.restrict s).range) (h0 : r 0) :
r f '' s

theorem measure_theory.simple_func.restrict_mono {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] [preorder β] (s : set α) {f g : measure_theory.simple_func α β} (H : f g) :

def measure_theory.simple_func.approx {α : Type u_1} {β : Type u_2} [measurable_space α] [semilattice_sup_bot β] [has_zero β] (i : → β) (f : α → β) (n : ) :

Fix a sequence i : ℕ → β. Given a function α → β, its n-th approximation by simple functions is defined so that in case β = ennreal it sends each a to the supremum of the set {i k | k ≤ n ∧ i k ≤ f a}, see approx_apply and supr_approx_apply for details.

Equations
theorem measure_theory.simple_func.approx_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [semilattice_sup_bot β] [has_zero β] [topological_space β] [order_closed_topology β] [measurable_space β] [opens_measurable_space β] {i : → β} {f : α → β} {n : } (a : α) (hf : measurable f) :
(measure_theory.simple_func.approx i f n) a = (finset.range n).sup (λ (k : ), ite (i k f a) (i k) 0)

theorem measure_theory.simple_func.monotone_approx {α : Type u_1} {β : Type u_2} [measurable_space α] [semilattice_sup_bot β] [has_zero β] (i : → β) (f : α → β) :

theorem measure_theory.simple_func.approx_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [semilattice_sup_bot β] [has_zero β] [topological_space β] [order_closed_topology β] [measurable_space β] [opens_measurable_space β] [measurable_space γ] {i : → β} {f : γ → β} {g : α → γ} {n : } (a : α) (hf : measurable f) (hg : measurable g) :

theorem measure_theory.simple_func.supr_approx_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [topological_space β] [complete_lattice β] [order_closed_topology β] [has_zero β] [measurable_space β] [opens_measurable_space β] (i : → β) (f : α → β) (a : α) (hf : measurable f) (h_zero : 0 = ) :
(⨆ (n : ), (measure_theory.simple_func.approx i f n) a) = ⨆ (k : ) (h : i k f a), i k

A sequence of ennreals such that its range is the set of non-negative rational numbers.

Equations
theorem measure_theory.simple_func.supr_eapprox_apply {α : Type u_1} [measurable_space α] (f : α → ennreal) (hf : measurable f) (a : α) :

theorem measure_theory.simple_func.eapprox_comp {α : Type u_1} {γ : Type u_3} [measurable_space α] [measurable_space γ] {f : γ → ennreal} {g : α → γ} {n : } (hf : measurable f) (hg : measurable g) :

Integral of a simple function whose codomain is ennreal.

Equations
theorem measure_theory.simple_func.lintegral_eq_of_subset {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (f : measure_theory.simple_func α ennreal) {s : finset ennreal} (hs : ∀ (x : α), f x 0μ (f ⁻¹' {f x}) 0f x s) :
f.lintegral μ = ∑ (x : ennreal) in s, x * μ (f ⁻¹' {x})

theorem measure_theory.simple_func.map_lintegral {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} (g : β → ennreal) (f : measure_theory.simple_func α β) :
(measure_theory.simple_func.map g f).lintegral μ = ∑ (x : β) in f.range, (g x) * μ (f ⁻¹' {x})

Calculate the integral of (g ∘ f), where g : β → ennreal and f : α →ₛ β.

@[simp]

theorem measure_theory.simple_func.lintegral_sum {α : Type u_1} [measurable_space α] {ι : Type u_2} (f : measure_theory.simple_func α ennreal) (μ : ι → measure_theory.measure α) :
f.lintegral (measure_theory.measure.sum μ) = ∑' (i : ι), f.lintegral (μ i)

theorem measure_theory.simple_func.restrict_lintegral {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (f : measure_theory.simple_func α ennreal) {s : set α} (hs : is_measurable s) :
(f.restrict s).lintegral μ = ∑ (r : ennreal) in f.range, r * μ (f ⁻¹' {r} s)

theorem measure_theory.simple_func.lintegral_restrict {α : Type u_1} [measurable_space α] (f : measure_theory.simple_func α ennreal) (s : set α) (μ : measure_theory.measure α) :
f.lintegral (μ.restrict s) = ∑ (y : ennreal) in f.range, y * μ (f ⁻¹' {y} s)

theorem measure_theory.simple_func.lintegral_mono {α : Type u_1} [measurable_space α] {f g : measure_theory.simple_func α ennreal} (hfg : f g) {μ ν : measure_theory.measure α} (hμν : μ ν) :

simple_func.lintegral is monotone both in function and in measure.

simple_func.lintegral depends only on the measures of f ⁻¹' {y}.

If two simple functions are equal a.e., then their lintegrals are equal.

theorem measure_theory.simple_func.lintegral_map {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {β : Type u_2} [measurable_space β] {μ' : measure_theory.measure β} (f : measure_theory.simple_func α ennreal) (g : measure_theory.simple_func β ennreal) (m : α → β) (eq : ∀ (a : α), f a = g (m a)) (h : ∀ (s : set β), is_measurable sμ' s = μ (m ⁻¹' s)) :
f.lintegral μ = g.lintegral μ'

theorem measure_theory.simple_func.support_eq {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) :
function.support f = ⋃ (y : β) (H : y finset.filter (λ (y : β), y 0) f.range), f ⁻¹' {y}

def measure_theory.simple_func.fin_meas_supp {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] (f : measure_theory.simple_func α β) (μ : measure_theory.measure α) :
Prop

A simple_func has finite measure support if it is equal to 0 outside of a set of finite measure.

Equations
theorem measure_theory.simple_func.fin_meas_supp_iff {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {f : measure_theory.simple_func α β} {μ : measure_theory.measure α} :
f.fin_meas_supp μ ∀ (y : β), y 0μ (f ⁻¹' {y}) <

theorem measure_theory.simple_func.fin_meas_supp.meas_preimage_singleton_ne_zero {α : Type u_1} {β : Type u_2} [measurable_space α] [has_zero β] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} (h : f.fin_meas_supp μ) {y : β} (hy : y 0) :
μ (f ⁻¹' {y}) <

theorem measure_theory.simple_func.fin_meas_supp.map {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [has_zero β] [has_zero γ] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} {g : β → γ} (hf : f.fin_meas_supp μ) (hg : g 0 = 0) :

theorem measure_theory.simple_func.fin_meas_supp.of_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [has_zero β] [has_zero γ] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} {g : β → γ} (h : (measure_theory.simple_func.map g f).fin_meas_supp μ) (hg : ∀ (b : β), g b = 0b = 0) :

theorem measure_theory.simple_func.fin_meas_supp.map_iff {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [has_zero β] [has_zero γ] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} {g : β → γ} (hg : ∀ {b : β}, g b = 0 b = 0) :

theorem measure_theory.simple_func.fin_meas_supp.pair {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] [has_zero β] [has_zero γ] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} {g : measure_theory.simple_func α γ} (hf : f.fin_meas_supp μ) (hg : g.fin_meas_supp μ) :

theorem measure_theory.simple_func.fin_meas_supp.map₂ {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [measurable_space α] [has_zero β] [has_zero γ] [has_zero δ] {μ : measure_theory.measure α} {f : measure_theory.simple_func α β} (hf : f.fin_meas_supp μ) {g : measure_theory.simple_func α γ} (hg : g.fin_meas_supp μ) {op : β → γ → δ} (H : op 0 0 = 0) :

theorem measure_theory.simple_func.fin_meas_supp.add {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {β : Type u_2} [add_monoid β] {f g : measure_theory.simple_func α β} (hf : f.fin_meas_supp μ) (hg : g.fin_meas_supp μ) :
(f + g).fin_meas_supp μ

theorem measure_theory.simple_func.fin_meas_supp.mul {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {β : Type u_2} [monoid_with_zero β] {f g : measure_theory.simple_func α β} (hf : f.fin_meas_supp μ) (hg : g.fin_meas_supp μ) :
(f * g).fin_meas_supp μ

theorem measure_theory.simple_func.fin_meas_supp.lintegral_lt_top {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : measure_theory.simple_func α ennreal} (hm : f.fin_meas_supp μ) (hf : ∀ᵐ (a : α) ∂μ, f a < ) :

theorem measure_theory.simple_func.induction {α : Type u_1} {γ : Type u_2} [measurable_space α] [add_monoid γ] {P : measure_theory.simple_func α γ → Prop} (h_ind : ∀ (c : γ) {s : set α} (hs : is_measurable s), P (measure_theory.simple_func.piecewise s hs (measure_theory.simple_func.const α c) (measure_theory.simple_func.const α 0))) (h_sum : ∀ ⦃f g : measure_theory.simple_func α γ⦄, set.univ f ⁻¹' {0} g ⁻¹' {0}P fP gP (f + g)) (f : measure_theory.simple_func α γ) :
P f

To prove something for an arbitrary simple function, it suffices to show that the property holds for (multiples of) characteristic functions and is closed under addition (of functions with disjoint support).

It is possible to make the hypotheses in h_sum a bit stronger, and such conditions can be added once we need them (for example it is only necessary to consider the case where g is a multiple of a characteristic function, and that this multiple doesn't appear in the image of f)

def measure_theory.lintegral {α : Type u_1} [measurable_space α] (μ : measure_theory.measure α) (f : α → ennreal) :

The lower Lebesgue integral of a function f with respect to a measure μ.

Equations
theorem measure_theory.lintegral_mono' {α : Type u_1} [measurable_space α] ⦃μ ν : measure_theory.measure α⦄ (hμν : μ ν) ⦃f g : α → ennreal (hfg : f g) :
∫⁻ (a : α), f a μ ∫⁻ (a : α), g a ν

theorem measure_theory.lintegral_mono {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} ⦃f g : α → ennreal (hfg : f g) :
∫⁻ (a : α), f a μ ∫⁻ (a : α), g a μ

@[simp]
theorem measure_theory.lintegral_const {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (c : ennreal) :
∫⁻ (a : α), c μ = c * μ set.univ

theorem measure_theory.set_lintegral_one {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (s : set α) :
∫⁻ (a : α) in s, 1 μ = μ s

theorem measure_theory.lintegral_eq_nnreal {α : Type u_1} [measurable_space α] (f : α → ennreal) (μ : measure_theory.measure α) :
∫⁻ (a : α), f a μ = ⨆ (φ : measure_theory.simple_func α ℝ≥0) (hf : ∀ (x : α), (φ x) f x), (measure_theory.simple_func.map coe φ).lintegral μ

∫⁻ a in s, f a ∂μ is defined as the supremum of integrals of simple functions φ : α →ₛ ennreal such that φ ≤ f. This lemma says that it suffices to take functions φ : α →ₛ ℝ≥0.

theorem measure_theory.supr_lintegral_le {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {ι : Sort u_2} (f : ι → α → ennreal) :
(⨆ (i : ι), ∫⁻ (a : α), f i a μ) ∫⁻ (a : α), (⨆ (i : ι), f i a) μ

theorem measure_theory.supr2_lintegral_le {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {ι : Sort u_2} {ι' : ι → Sort u_3} (f : Π (i : ι), ι' iα → ennreal) :
(⨆ (i : ι) (h : ι' i), ∫⁻ (a : α), f i h a μ) ∫⁻ (a : α), (⨆ (i : ι) (h : ι' i), f i h a) μ

theorem measure_theory.le_infi_lintegral {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {ι : Sort u_2} (f : ι → α → ennreal) :
∫⁻ (a : α), (⨅ (i : ι), f i a) μ ⨅ (i : ι), ∫⁻ (a : α), f i a μ

theorem measure_theory.le_infi2_lintegral {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {ι : Sort u_2} {ι' : ι → Sort u_3} (f : Π (i : ι), ι' iα → ennreal) :
∫⁻ (a : α), (⨅ (i : ι) (h : ι' i), f i h a) μ ⨅ (i : ι) (h : ι' i), ∫⁻ (a : α), f i h a μ

theorem measure_theory.lintegral_supr {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : ∀ (n : ), measurable (f n)) (h_mono : monotone f) :
∫⁻ (a : α), (⨆ (n : ), f n a) μ = ⨆ (n : ), ∫⁻ (a : α), f n a μ

Monotone convergence theorem -- sometimes called Beppo-Levi convergence.

See lintegral_supr_directed for a more general form.

theorem measure_theory.lintegral_eq_supr_eapprox_lintegral {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : measurable f) :
∫⁻ (a : α), f a μ = ⨆ (n : ), (measure_theory.simple_func.eapprox f n).lintegral μ

theorem measure_theory.lintegral_add {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f g : α → ennreal} (hf : measurable f) (hg : measurable g) :
∫⁻ (a : α), f a + g a μ = ∫⁻ (a : α), f a μ + ∫⁻ (a : α), g a μ

theorem measure_theory.lintegral_zero {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} :
∫⁻ (a : α), 0 μ = 0

theorem measure_theory.lintegral_smul_measure {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (c : ennreal) (f : α → ennreal) :
∫⁻ (a : α), f a c μ = c * ∫⁻ (a : α), f a μ

theorem measure_theory.lintegral_sum_measure {α : Type u_1} [measurable_space α] {ι : Type u_2} (f : α → ennreal) (μ : ι → measure_theory.measure α) :
∫⁻ (a : α), f a measure_theory.measure.sum μ = ∑' (i : ι), ∫⁻ (a : α), f a μ i

theorem measure_theory.lintegral_add_measure {α : Type u_1} [measurable_space α] (f : α → ennreal) (μ ν : measure_theory.measure α) :
∫⁻ (a : α), f a + ν) = ∫⁻ (a : α), f a μ + ∫⁻ (a : α), f a ν

@[simp]
theorem measure_theory.lintegral_zero_measure {α : Type u_1} [measurable_space α] (f : α → ennreal) :
∫⁻ (a : α), f a 0 = 0

theorem measure_theory.lintegral_finset_sum {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} (s : finset β) {f : β → α → ennreal} (hf : ∀ (b : β), measurable (f b)) :
∫⁻ (a : α), ∑ (b : β) in s, f b a μ = ∑ (b : β) in s, ∫⁻ (a : α), f b a μ

theorem measure_theory.lintegral_const_mul {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (r : ennreal) {f : α → ennreal} (hf : measurable f) :
∫⁻ (a : α), r * f a μ = r * ∫⁻ (a : α), f a μ

theorem measure_theory.lintegral_const_mul_le {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (r : ennreal) (f : α → ennreal) :
r * ∫⁻ (a : α), f a μ ∫⁻ (a : α), r * f a μ

theorem measure_theory.lintegral_const_mul' {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (r : ennreal) (f : α → ennreal) (hr : r ) :
∫⁻ (a : α), r * f a μ = r * ∫⁻ (a : α), f a μ

theorem measure_theory.lintegral_mono_ae {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f g : α → ennreal} (h : ∀ᵐ (a : α) ∂μ, f a g a) :
∫⁻ (a : α), f a μ ∫⁻ (a : α), g a μ

theorem measure_theory.lintegral_congr_ae {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f g : α → ennreal} (h : f =ᵐ[μ] g) :
∫⁻ (a : α), f a μ = ∫⁻ (a : α), g a μ

theorem measure_theory.lintegral_congr {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f g : α → ennreal} (h : ∀ (a : α), f a = g a) :
∫⁻ (a : α), f a μ = ∫⁻ (a : α), g a μ

theorem measure_theory.set_lintegral_congr {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} {s t : set α} (h : s =ᵐ[μ] t) :
∫⁻ (x : α) in s, f x μ = ∫⁻ (x : α) in t, f x μ

theorem measure_theory.lintegral_rw₁ {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} {f f' : α → β} (h : f =ᵐ[μ] f') (g : β → ennreal) :
∫⁻ (a : α), g (f a) μ = ∫⁻ (a : α), g (f' a) μ

theorem measure_theory.lintegral_rw₂ {α : Type u_1} {β : Type u_2} {γ : Type u_3} [measurable_space α] {μ : measure_theory.measure α} {f₁ f₁' : α → β} {f₂ f₂' : α → γ} (h₁ : f₁ =ᵐ[μ] f₁') (h₂ : f₂ =ᵐ[μ] f₂') (g : β → γ → ennreal) :
∫⁻ (a : α), g (f₁ a) (f₂ a) μ = ∫⁻ (a : α), g (f₁' a) (f₂' a) μ

@[simp]
theorem measure_theory.lintegral_indicator {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (f : α → ennreal) {s : set α} (hs : is_measurable s) :
∫⁻ (a : α), s.indicator f a μ = ∫⁻ (a : α) in s, f a μ

theorem measure_theory.mul_meas_ge_le_lintegral {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : measurable f) (ε : ennreal) :
ε * μ {x : α | ε f x} ∫⁻ (a : α), f a μ

Chebyshev's inequality

theorem measure_theory.meas_ge_le_lintegral_div {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : measurable f) {ε : ennreal} (hε : ε 0) (hε' : ε ) :
μ {x : α | ε f x} ∫⁻ (a : α), f a μ / ε

theorem measure_theory.lintegral_eq_zero_iff {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : measurable f) :
∫⁻ (a : α), f a μ = 0 f =ᵐ[μ] 0

theorem measure_theory.lintegral_supr_ae {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (hf : ∀ (n : ), measurable (f n)) (h_mono : ∀ (n : ), ∀ᵐ (a : α) ∂μ, f n a f n.succ a) :
∫⁻ (a : α), (⨆ (n : ), f n a) μ = ⨆ (n : ), ∫⁻ (a : α), f n a μ

Weaker version of the monotone convergence theorem

theorem measure_theory.lintegral_sub {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f g : α → ennreal} (hf : measurable f) (hg : measurable g) (hg_fin : ∫⁻ (a : α), g a μ < ) (h_le : g ≤ᵐ[μ] f) :
∫⁻ (a : α), f a - g a μ = ∫⁻ (a : α), f a μ - ∫⁻ (a : α), g a μ

theorem measure_theory.lintegral_infi_ae {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (h_meas : ∀ (n : ), measurable (f n)) (h_mono : ∀ (n : ), f n.succ ≤ᵐ[μ] f n) (h_fin : ∫⁻ (a : α), f 0 a μ < ) :
∫⁻ (a : α), (⨅ (n : ), f n a) μ = ⨅ (n : ), ∫⁻ (a : α), f n a μ

Monotone convergence theorem for nonincreasing sequences of functions

theorem measure_theory.lintegral_infi {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (h_meas : ∀ (n : ), measurable (f n)) (h_mono : ∀ ⦃m n : ⦄, m nf n f m) (h_fin : ∫⁻ (a : α), f 0 a μ < ) :
∫⁻ (a : α), (⨅ (n : ), f n a) μ = ⨅ (n : ), ∫⁻ (a : α), f n a μ

Monotone convergence theorem for nonincreasing sequences of functions

theorem measure_theory.lintegral_liminf_le {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} (h_meas : ∀ (n : ), measurable (f n)) :
∫⁻ (a : α), filter.at_top.liminf (λ (n : ), f n a) μ filter.at_top.liminf (λ (n : ), ∫⁻ (a : α), f n a μ)

Known as Fatou's lemma

theorem measure_theory.limsup_lintegral_le {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {f : α → ennreal} {g : α → ennreal} (hf_meas : ∀ (n : ), measurable (f n)) (h_bound : ∀ (n : ), f n ≤ᵐ[μ] g) (h_fin : ∫⁻ (a : α), g a μ < ) :
filter.at_top.limsup (λ (n : ), ∫⁻ (a : α), f n a μ) ∫⁻ (a : α), filter.at_top.limsup (λ (n : ), f n a) μ

theorem measure_theory.tendsto_lintegral_of_dominated_convergence {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {F : α → ennreal} {f : α → ennreal} (bound : α → ennreal) (hF_meas : ∀ (n : ), measurable (F n)) (h_bound : ∀ (n : ), F n ≤ᵐ[μ] bound) (h_fin : ∫⁻ (a : α), bound a μ < ) (h_lim : ∀ᵐ (a : α) ∂μ, filter.tendsto (λ (n : ), F n a) filter.at_top (𝓝 (f a))) :
filter.tendsto (λ (n : ), ∫⁻ (a : α), F n a μ) filter.at_top (𝓝 (∫⁻ (a : α), f a μ))

Dominated convergence theorem for nonnegative functions

theorem measure_theory.tendsto_lintegral_filter_of_dominated_convergence {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} {ι : Type u_2} {l : filter ι} {F : ι → α → ennreal} {f : α → ennreal} (bound : α → ennreal) (hl_cb : l.is_countably_generated) (hF_meas : ∀ᶠ (n : ι) in l, measurable (F n)) (h_bound : ∀ᶠ (n : ι) in l, ∀ᵐ (a : α) ∂μ, F n a bound a) (h_fin : ∫⁻ (a : α), bound a μ < ) (h_lim : ∀ᵐ (a : α) ∂μ, filter.tendsto (λ (n : ι), F n a) l (𝓝 (f a))) :
filter.tendsto (λ (n : ι), ∫⁻ (a : α), F n a μ) l (𝓝 (∫⁻ (a : α), f a μ))

Dominated convergence theorem for filters with a countable basis

theorem measure_theory.lintegral_supr_directed {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [encodable β] {f : β → α → ennreal} (hf : ∀ (b : β), measurable (f b)) (h_directed : directed has_le.le f) :
∫⁻ (a : α), (⨆ (b : β), f b a) μ = ⨆ (b : β), ∫⁻ (a : α), f b a μ

Monotone convergence for a suprema over a directed family and indexed by an encodable type

theorem measure_theory.lintegral_tsum {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [encodable β] {f : β → α → ennreal} (hf : ∀ (i : β), measurable (f i)) :
∫⁻ (a : α), (∑' (i : β), f i a) μ = ∑' (i : β), ∫⁻ (a : α), f i a μ

theorem measure_theory.lintegral_Union {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [encodable β] {s : β → set α} (hm : ∀ (i : β), is_measurable (s i)) (hd : pairwise (disjoint on s)) (f : α → ennreal) :
∫⁻ (a : α) in ⋃ (i : β), s i, f a μ = ∑' (i : β), ∫⁻ (a : α) in s i, f a μ

theorem measure_theory.lintegral_Union_le {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [encodable β] (s : β → set α) (f : α → ennreal) :
∫⁻ (a : α) in ⋃ (i : β), s i, f a μ ∑' (i : β), ∫⁻ (a : α) in s i, f a μ

theorem measure_theory.lintegral_map {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [measurable_space β] {f : β → ennreal} {g : α → β} (hf : measurable f) (hg : measurable g) :
∫⁻ (a : β), f a (measure_theory.measure.map g) μ = ∫⁻ (a : α), f (g a) μ

theorem measure_theory.set_lintegral_map {α : Type u_1} {β : Type u_2} [measurable_space α] {μ : measure_theory.measure α} [measurable_space β] {f : β → ennreal} {g : α → β} {s : set β} (hs : is_measurable s) (hf : measurable f) (hg : measurable g) :
∫⁻ (y : β) in s, f y (measure_theory.measure.map g) μ = ∫⁻ (x : α) in g ⁻¹' s, f (g x) μ

theorem measure_theory.lintegral_dirac {α : Type u_1} [measurable_space α] (a : α) {f : α → ennreal} (hf : measurable f) :

Given a measure μ : measure α and a function f : α → ennreal, μ.with_density f is the measure such that for a measurable set s we have μ.with_density f s = ∫⁻ a in s, f a ∂μ.

Equations
theorem measure_theory.with_density_apply {α : Type u_1} [measurable_space α] {μ : measure_theory.measure α} (f : α → ennreal) {s : set α} (hs : is_measurable s) :
(μ.with_density f) s = ∫⁻ (a : α) in s, f a μ

theorem measurable.ennreal_induction {α : Type u_1} [measurable_space α] {P : (α → ennreal) → Prop} (h_ind : ∀ (c : ennreal) ⦃s : set α⦄, is_measurable sP (s.indicator (λ (_x : α), c))) (h_sum : ∀ ⦃f g : α → ennreal⦄, set.univ f ⁻¹' {0} g ⁻¹' {0}measurable fmeasurable gP fP gP (f + g)) (h_supr : ∀ ⦃f : α → ennreal⦄, (∀ (n : ), measurable (f n))monotone f(∀ (n : ), P (f n))P (λ (x : α), ⨆ (n : ), f n x)) ⦃f : α → ennreal (hf : measurable f) :
P f

To prove something for an arbitrary measurable function into ennreal, it suffices to show that the property holds for (multiples of) characteristic functions and is closed under addition and supremum of increasing sequences of functions.

It is possible to make the hypotheses in the induction steps a bit stronger, and such conditions can be added once we need them (for example in h_sum it is only necessary to consider the sum of a simple function with a multiple of a characteristic function and that the intersection of their images is a subset of {0}.