Independence of families of sets with respect to a kernel and a measure #
A family of sets of sets π : ι → Set (Set Ω) is independent with respect to a kernel
κ : Kernel α Ω and a measure μ on α if for any finite set of indices s = {i_1, ..., i_n},
for any sets f i_1 ∈ π i_1, ..., f i_n ∈ π i_n, then for μ-almost every a : α,
κ a (⋂ i in s, f i) = ∏ i ∈ s, κ a (f i).
This notion of independence is a generalization of both independence and conditional independence.
For conditional independence, κ is the conditional kernel ProbabilityTheory.condExpKernel and
μ is the ambient measure. For (non-conditional) independence, κ = Kernel.const Unit μ and the
measure is the Dirac measure on Unit.
The main purpose of this file is to prove only once the properties that hold for both conditional and non-conditional independence.
This file contains results about independence of families of sets and σ-algebras.
See the IndepFun file for results about independence of random variables.
Main definitions #
ProbabilityTheory.Kernel.iIndepSets: independence of a family of sets of sets. Variant for two sets of sets:ProbabilityTheory.Kernel.IndepSets.ProbabilityTheory.Kernel.iIndep: independence of a family of σ-algebras. Variant for two σ-algebras:Indep.ProbabilityTheory.Kernel.iIndepSet: independence of a family of sets. Variant for two sets:ProbabilityTheory.Kernel.IndepSet.
See the file Mathlib/Probability/Kernel/Basic.lean for a more detailed discussion of these
definitions in the particular case of the usual independence notion.
Main statements #
ProbabilityTheory.Kernel.iIndepSets.iIndep: if π-systems are independent as sets of sets, then the measurable space structures they generate are independent.ProbabilityTheory.Kernel.IndepSets.Indep: variant with two π-systems.
A family of sets of sets π : ι → Set (Set Ω) is independent with respect to a kernel κ and
a measure μ if for any finite set of indices s = {i_1, ..., i_n}, for any sets
f i_1 ∈ π i_1, ..., f i_n ∈ π i_n, then ∀ᵐ a ∂μ, κ a (⋂ i in s, f i) = ∏ i ∈ s, κ a (f i).
It will be used for families of pi_systems.
Equations
Instances For
Two sets of sets s₁, s₂ are independent with respect to a kernel κ and a measure μ if for
any sets t₁ ∈ s₁, t₂ ∈ s₂, then ∀ᵐ a ∂μ, κ a (t₁ ∩ t₂) = κ a (t₁) * κ a (t₂)
Equations
Instances For
A family of measurable space structures (i.e. of σ-algebras) is independent with respect to a
kernel κ and a measure μ if the family of sets of measurable sets they define is independent.
Equations
- ProbabilityTheory.Kernel.iIndep m κ μ = ProbabilityTheory.Kernel.iIndepSets (fun (x : ι) => {s : Set Ω | MeasurableSet s}) κ μ
Instances For
Two measurable space structures (or σ-algebras) m₁, m₂ are independent with respect to a
kernel κ and a measure μ if for any sets t₁ ∈ m₁, t₂ ∈ m₂,
∀ᵐ a ∂μ, κ a (t₁ ∩ t₂) = κ a (t₁) * κ a (t₂)
Equations
- ProbabilityTheory.Kernel.Indep m₁ m₂ κ μ = ProbabilityTheory.Kernel.IndepSets {s : Set Ω | MeasurableSet s} {s : Set Ω | MeasurableSet s} κ μ
Instances For
A family of sets is independent if the family of measurable space structures they generate is
independent. For a set s, the generated measurable space has measurable sets ∅, s, sᶜ, univ.
Equations
- ProbabilityTheory.Kernel.iIndepSet s κ μ = ProbabilityTheory.Kernel.iIndep (fun (i : ι) => MeasurableSpace.generateFrom {s i}) κ μ
Instances For
Two sets are independent if the two measurable space structures they generate are independent.
For a set s, the generated measurable space structure has measurable sets ∅, s, sᶜ, univ.
Equations
Instances For
Alias of the forward direction of ProbabilityTheory.Kernel.iIndepSets_congr.
Alias of the forward direction of ProbabilityTheory.Kernel.indepSets_congr.
Alias of the forward direction of ProbabilityTheory.Kernel.iIndep_congr.
Alias of the forward direction of ProbabilityTheory.Kernel.indep_congr.
Alias of the forward direction of ProbabilityTheory.Kernel.iIndepSet_congr.
Alias of the forward direction of ProbabilityTheory.Kernel.indepSet_congr.
π-system lemma #
Independence of measurable spaces is equivalent to independence of generating π-systems.
Independence of measurable space structures implies independence of generating π-systems #
Independence of generating π-systems implies independence of measurable space structures #
The measurable space structures generated by independent pi-systems are independent.
Alias of ProbabilityTheory.Kernel.iIndepSets.piiUnionInter_of_notMem.
The measurable space structures generated by independent pi-systems are independent.
Independence of measurable sets #
We prove the following equivalences on IndepSet, for measurable sets s, t.