Markov Kernels #
A kernel from a measurable space α
to another measurable space β
is a measurable map
α → MeasureTheory.Measure β
, where the measurable space instance on measure β
is the one defined
in MeasureTheory.Measure.instMeasurableSpace
. That is, a kernel κ
verifies that for all
measurable sets s
of β
, a ↦ κ a s
is measurable.
Main definitions #
Classes of kernels:
ProbabilityTheory.kernel α β
: kernels fromα
toβ
, defined as theAddSubmonoid
of the measurable functions inα → Measure β
.ProbabilityTheory.IsMarkovKernel κ
: a kernel fromα
toβ
is said to be a Markov kernel if for alla : α
,k a
is a probability measure.ProbabilityTheory.IsFiniteKernel κ
: a kernel fromα
toβ
is said to be finite if there existsC : ℝ≥0∞
such thatC < ∞
and for alla : α
,κ a univ ≤ C
. This implies in particular that all measures in the image ofκ
are finite, but is stronger since it requires a uniform bound. This stronger condition is necessary to ensure that the composition of two finite kernels is finite.ProbabilityTheory.IsSFiniteKernel κ
: a kernel is called s-finite if it is a countable sum of finite kernels.
Particular kernels:
ProbabilityTheory.kernel.deterministic (f : α → β) (hf : Measurable f)
: kernela ↦ Measure.dirac (f a)
.ProbabilityTheory.kernel.const α (μβ : measure β)
: constant kernela ↦ μβ
.ProbabilityTheory.kernel.restrict κ (hs : MeasurableSet s)
: kernel for which the image ofa : α
is(κ a).restrict s
. Integral:∫⁻ b, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in s, f b ∂(κ a)
Main statements #
ProbabilityTheory.kernel.ext_fun
: if∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)
for all measurable functionsf
and alla
, then the two kernelsκ
andη
are equal.
A kernel from a measurable space α
to another measurable space β
is a measurable function
κ : α → Measure β
. The measurable space structure on MeasureTheory.Measure β
is given by
MeasureTheory.Measure.instMeasurableSpace
. A map κ : α → MeasureTheory.Measure β
is measurable
iff ∀ s : Set β, MeasurableSet s → Measurable (fun a ↦ κ a s)
.
Equations
- ProbabilityTheory.kernel α β = { carrier := Measurable, add_mem' := ⋯, zero_mem' := ⋯ }
Instances For
Equations
- ProbabilityTheory.instFunLikeSubtypeForallMeasureMemAddSubmonoidKernel = { coe := Subtype.val, coe_injective' := ⋯ }
Equations
- ⋯ = ⋯
Equations
- ProbabilityTheory.kernel.instOrderBot = OrderBot.mk ⋯
Coercion to a function as an additive monoid homomorphism.
Equations
- ProbabilityTheory.kernel.coeAddHom α β = (ProbabilityTheory.kernel α β).subtype
Instances For
A kernel is a Markov kernel if every measure in its image is a probability measure.
- isProbabilityMeasure : ∀ (a : α), MeasureTheory.IsProbabilityMeasure (κ a)
Instances
A kernel is finite if every measure in its image is finite, with a uniform bound.
Instances
A constant C : ℝ≥0∞
such that C < ∞
(ProbabilityTheory.IsFiniteKernel.bound_lt_top κ
) and
for all a : α
and s : Set β
, κ a s ≤ C
(ProbabilityTheory.kernel.measure_le_bound κ a s
).
Porting note (#11215): TODO: does it make sense to
-- make ProbabilityTheory.IsFiniteKernel.bound
the least possible bound?
-- Should it be an NNReal
number?
Equations
- ProbabilityTheory.IsFiniteKernel.bound κ = ⋯.choose
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Sum of an indexed family of kernels.
Equations
- ProbabilityTheory.kernel.sum κ = ⟨fun (a : α) => MeasureTheory.Measure.sum fun (n : ι) => (κ n) a, ⋯⟩
Instances For
A kernel is s-finite if it can be written as the sum of countably many finite kernels.
- tsum_finite : ∃ (κs : ℕ → ↥(ProbabilityTheory.kernel α β)), (∀ (n : ℕ), ProbabilityTheory.IsFiniteKernel (κs n)) ∧ κ = ProbabilityTheory.kernel.sum κs
Instances
Equations
- ⋯ = ⋯
A sequence of finite kernels such that κ = ProbabilityTheory.kernel.sum (seq κ)
. See
ProbabilityTheory.kernel.isFiniteKernel_seq
and ProbabilityTheory.kernel.kernel_sum_seq
.
Equations
- ProbabilityTheory.kernel.seq κ = ⋯.choose
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Kernel which to a
associates the dirac measure at f a
. This is a Markov kernel.
Equations
- ProbabilityTheory.kernel.deterministic f hf = ⟨fun (a : α) => MeasureTheory.Measure.dirac (f a), ⋯⟩
Instances For
Equations
- ⋯ = ⋯
Alias of ProbabilityTheory.kernel.setLIntegral_deterministic'
.
Alias of ProbabilityTheory.kernel.setLIntegral_deterministic
.
Alias of ProbabilityTheory.kernel.setIntegral_deterministic'
.
Alias of ProbabilityTheory.kernel.setIntegral_deterministic
.
Constant kernel, which always returns the same measure.
Equations
- ProbabilityTheory.kernel.const α μβ = ⟨fun (x : α) => μβ, ⋯⟩
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
In a countable space with measurable singletons, every function α → MeasureTheory.Measure β
defines a kernel.
Equations
Instances For
Kernel given by the restriction of the measures in the image of a kernel to a set.
Equations
- ProbabilityTheory.kernel.restrict κ hs = ⟨fun (a : α) => (κ a).restrict s, ⋯⟩
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Kernel with value (κ a).comap f
, for a measurable embedding f
. That is, for a measurable set
t : Set β
, ProbabilityTheory.kernel.comapRight κ hf a t = κ a (f '' t)
.
Equations
- ProbabilityTheory.kernel.comapRight κ hf = ⟨fun (a : α) => MeasureTheory.Measure.comap f (κ a), ⋯⟩
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
ProbabilityTheory.kernel.piecewise hs κ η
is the kernel equal to κ
on the measurable set s
and to η
on its complement.
Equations
- ProbabilityTheory.kernel.piecewise hs κ η = ⟨fun (a : α) => if a ∈ s then κ a else η a, ⋯⟩
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯