mathlib documentation

measure_theory.​giry_monad

measure_theory.​giry_monad

The Giry monad

Let X be a measurable space. The collection of all measures on X again forms a measurable space. This construction forms a monad on measurable spaces and measurable functions, called the Giry monad.

Note that most sources use the term "Giry monad" for the restriction to probability measures. Here we include all measures on X.

See also measure_theory/category/Meas.lean, containing an upgrade of the type-level monad to an honest monad of the functor Measure : MeasMeas.

References

Tags

giry monad

@[instance]

Measurability structure on measure: Measures are measurable w.r.t. all projections

Equations
theorem measure_theory.​measure.​measurable_of_measurable_coe {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] (f : β → measure_theory.measure α) :
(∀ (s : set α), is_measurable smeasurable (λ (b : β), (f b) s))measurable f

theorem measure_theory.​measure.​measurable_map {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] (f : α → β) :

theorem measure_theory.​measure.​measurable_lintegral {α : Type u_1} [measurable_space α] {f : α → ennreal} :
measurable fmeasurable (λ (μ : measure_theory.measure α), ∫⁻ (x : α), f x μ)

Monadic join on measure in the category of measurable spaces and measurable functions.

Equations

Monadic bind on measure, only works in the category of measurable spaces and measurable functions. When the function f is not measurable the result is not well defined.

Equations
@[simp]
theorem measure_theory.​measure.​bind_apply {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] {m : measure_theory.measure α} {f : α → measure_theory.measure β} {s : set β} :
is_measurable smeasurable f(m.bind f) s = ∫⁻ (a : α), (f a) s m

theorem measure_theory.​measure.​measurable_bind' {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] {g : α → measure_theory.measure β} :

theorem measure_theory.​measure.​lintegral_bind {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] {m : measure_theory.measure α} {μ : α → measure_theory.measure β} {f : β → ennreal} :
measurable μmeasurable f∫⁻ (x : β), f x m.bind μ = ∫⁻ (a : α), ∫⁻ (x : β), f x μ a m

theorem measure_theory.​measure.​bind_bind {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] {γ : Type u_3} [measurable_space γ] {m : measure_theory.measure α} {f : α → measure_theory.measure β} {g : β → measure_theory.measure γ} :
measurable fmeasurable g(m.bind f).bind g = m.bind (λ (a : α), (f a).bind g)

theorem measure_theory.​measure.​bind_dirac {α : Type u_1} {β : Type u_2} [measurable_space α] [measurable_space β] {f : α → measure_theory.measure β} (hf : measurable f) (a : α) :