Measurable spaces and measurable functions #
This file provides properties of measurable spaces and the functions and isomorphisms between them.
The definition of a measurable space is in Mathlib/MeasureTheory/MeasurableSpace/Defs.lean.
A measurable space is a set equipped with a σ-algebra, a collection of subsets closed under complementation and countable union. A function between measurable spaces is measurable if the preimage of each measurable subset is measurable.
σ-algebras on a fixed set α form a complete lattice. Here we order
σ-algebras by writing m₁ ≤ m₂ if every set which is m₁-measurable is
also m₂-measurable (that is, m₁ is a subset of m₂). In particular, any
collection of subsets of α generates a smallest σ-algebra which
contains all of them. A function f : α → β induces a Galois connection
between the lattices of σ-algebras on α and β.
We say that a filter f is measurably generated if every set s ∈ f includes a measurable
set t ∈ f. This property is useful, e.g., to extract a measurable witness of Filter.Eventually.
Implementation notes #
Measurability of a function f : α → β between measurable spaces is
defined in terms of the Galois connection induced by f.
References #
- https://en.wikipedia.org/wiki/Measurable_space
- https://en.wikipedia.org/wiki/Sigma-algebra
- https://en.wikipedia.org/wiki/Dynkin_system
Tags #
measurable space, σ-algebra, measurable function, dynkin system, π-λ theorem, π-system
The forward image of a measurable space under a function. map f m contains the sets
s : Set β whose preimage under f is measurable.
Equations
- MeasurableSpace.map f m = { MeasurableSet' := fun (s : Set β) => MeasurableSet (f ⁻¹' s), measurableSet_empty := ⋯, measurableSet_compl := ⋯, measurableSet_iUnion := ⋯ }
Instances For
The reverse image of a measurable space under a function. comap f m contains the sets
s : Set α such that s is the f-preimage of a measurable set in β.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Alias of the reverse direction of measurable_iff_le_map.
Alias of the forward direction of measurable_iff_le_map.
Alias of the reverse direction of measurable_iff_comap_le.
Alias of the forward direction of measurable_iff_comap_le.
A version of measurable_const that assumes f x = f y for all x, y. This version works
for functions between empty types.
This is slightly different from Measurable.piecewise. It can be used to show
Measurable (ite (x=0) 0 1) by
exact Measurable.ite (measurableSet_singleton 0) measurable_const measurable_const,
but replacing Measurable.ite by Measurable.piecewise in that example proof does not work.
The measurability of a set A is equivalent to the measurability of the indicator function
which takes a constant value b ≠ 0 on a set A and 0 elsewhere.
If a function coincides with a measurable function outside of a countable set, it is measurable.
Equations
- ULift.instMeasurableSpace = MeasurableSpace.map ULift.up inst
Equations
- Quot.instMeasurableSpace = MeasurableSpace.map (Quot.mk r) m
Equations
- Quotient.instMeasurableSpace = MeasurableSpace.map Quotient.mk'' m
Equations
- QuotientAddGroup.measurableSpace S = Quotient.instMeasurableSpace
Equations
- QuotientGroup.measurableSpace S = Quotient.instMeasurableSpace
Equations
- Subtype.instMeasurableSpace = MeasurableSpace.comap Subtype.val m
Equations
- ⋯ = ⋯
Alias of Measurable.subtype_coe.
The measurable atom of x is the intersection of all the measurable sets countaining x.
It is measurable when the space is countable (or more generally when the measurable space is
countably generated).
Equations
- measurableAtom x = ⋂ (s : Set β), ⋂ (_ : x ∈ s), ⋂ (_ : MeasurableSet s), s
Instances For
A MeasurableSpace structure on the product of two measurable spaces.
Equations
- m₁.prod m₂ = MeasurableSpace.comap Prod.fst m₁ ⊔ MeasurableSpace.comap Prod.snd m₂
Instances For
Equations
- Prod.instMeasurableSpace = m₁.prod m₂
Equations
- ⋯ = ⋯
A piecewise function on countably many pieces is measurable if all the data is measurable.
Let t i be a countable covering of a set T by measurable sets. Let f i : t i → β be a
family of functions that agree on the intersections t i ∩ t j. Then the function
Set.iUnionLift t f _ _ : T → β, defined as f i ⟨x, hx⟩ for hx : x ∈ t i, is measurable.
Let t i be a countable covering of α by measurable sets. Let f i : t i → β be a family of
functions that agree on the intersections t i ∩ t j. Then the function Set.liftCover t f _ _,
defined as f i ⟨x, hx⟩ for hx : x ∈ t i, is measurable.
Let t i be a nonempty countable family of measurable sets in α. Let g i : α → β be a
family of measurable functions such that g i agrees with g j on t i ∩ t j. Then there exists
a measurable function f : α → β that agrees with each g i on t i.
We only need the assumption [Nonempty ι] to prove [Nonempty (α → β)].
Equations
- MeasurableSpace.pi = ⨆ (a : δ), MeasurableSpace.comap (fun (b : (a : δ) → π a) => b a) (m a)
The function (f, x) ↦ update f a x : (Π a, π a) × π a → Π a, π a is measurable.
The function update f a : π a → Π a, π a is always measurable.
This doesn't require f to be measurable.
This should not be confused with the statement that update f a x is measurable.
Equations
- ⋯ = ⋯
Equations
- TProd.instMeasurableSpace π [] = PUnit.instMeasurableSpace
- TProd.instMeasurableSpace π (head :: is) = Prod.instMeasurableSpace
Equations
- Sum.instMeasurableSpace = MeasurableSpace.map Sum.inl m₁ ⊓ MeasurableSpace.map Sum.inr m₂
Alias of the reverse direction of measurableSet_inl_image.
Alias of the reverse direction of measurableSet_inr_image.
Equations
- Sigma.instMeasurableSpace = ⨅ (a : α), MeasurableSpace.map (Sigma.mk a) (m a)
Alias of the reverse direction of measurableSet_setOf.
Alias of the reverse direction of measurable_mem.
This instance is useful when talking about Bernoulli sequences of random variables or binomial random graphs.
Equations
- ⋯ = ⋯
The sigma-algebra generated by a single set s is {∅, s, sᶜ, univ}.
A filter f is measurably generates if each s ∈ f includes a measurable t ∈ f.
- exists_measurable_subset : ∀ ⦃s : Set α⦄, s ∈ f → ∃ t ∈ f, MeasurableSet t ∧ t ⊆ s
Instances
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Alias of the reverse direction of Filter.principal_isMeasurablyGenerated_iff.
Equations
- ⋯ = ⋯
The set of points for which a sequence of measurable functions converges to a given value is measurable.
We say that a collection of sets is countably spanning if a countable subset spans the
whole type. This is a useful condition in various parts of measure theory. For example, it is
a needed condition to show that the product of two collections generate the product sigma algebra,
see generateFrom_prod_eq.
Equations
Instances For
Typeclasses on Subtype MeasurableSet #
Equations
- MeasurableSet.Subtype.instSingleton = { singleton := fun (a : α) => ⟨{a}, ⋯⟩ }
Equations
- ⋯ = ⋯
Equations
- MeasurableSet.Subtype.instTop = { top := ⟨Set.univ, ⋯⟩ }
Equations
- MeasurableSet.Subtype.instBooleanAlgebra = Function.Injective.booleanAlgebra (fun (a : Subtype MeasurableSet) => ↑a) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯