Documentation

PFR.Mathlib.Probability.Independence.Conditional

theorem ProbabilityTheory.IndepFun.identDistrib_cond {Ω : Type u_1} {α : Type u_2} {β : Type u_3} :
∀ {x : MeasurableSpace Ω} {x_1 : MeasurableSpace α} {x_2 : MeasurableSpace β} {μ : MeasureTheory.Measure Ω} {A : Ωα} {B : Ωβ} [inst : MeasureTheory.IsProbabilityMeasure μ], ProbabilityTheory.IndepFun A B μ∀ {s : Set β}, MeasurableSet sMeasurable AMeasurable Bμ (B ⁻¹' s) 0ProbabilityTheory.IdentDistrib A A μ (ProbabilityTheory.cond μ (B ⁻¹' s))

If A is independent from B, then conditioning on an event given by B does not change the distribution of A.

theorem ProbabilityTheory.IndepFun.cond_left {Ω : Type u_1} {α : Type u_2} {β : Type u_3} :
∀ {x : MeasurableSpace Ω} {x_1 : MeasurableSpace α} {x_2 : MeasurableSpace β} {μ : MeasureTheory.Measure Ω} {A : Ωα} {B : Ωβ}, ProbabilityTheory.IndepFun A B μ∀ {s : Set α}, MeasurableSet sMeasurable AProbabilityTheory.IndepFun A B (ProbabilityTheory.cond μ (A ⁻¹' s))

If A is independent of B, then they remain independent when conditioning on an event of the form A ∈ s of positive probability.

theorem ProbabilityTheory.IndepFun.cond_right {Ω : Type u_1} {α : Type u_2} {β : Type u_3} :
∀ {x : MeasurableSpace Ω} {x_1 : MeasurableSpace α} {x_2 : MeasurableSpace β} {μ : MeasureTheory.Measure Ω} {A : Ωα} {B : Ωβ}, ProbabilityTheory.IndepFun A B μ∀ {t : Set β}, MeasurableSet tMeasurable BProbabilityTheory.IndepFun A B (ProbabilityTheory.cond μ (B ⁻¹' t))

If A is independent of B, then they remain independent when conditioning on an event of the form B ∈ t of positive probability.

theorem ProbabilityTheory.IndepFun.cond {Ω : Type u_1} {α : Type u_2} {β : Type u_3} :
∀ {x : MeasurableSpace Ω} {x_1 : MeasurableSpace α} {x_2 : MeasurableSpace β} {μ : MeasureTheory.Measure Ω} {A : Ωα} {B : Ωβ}, ProbabilityTheory.IndepFun A B μ∀ {s : Set α} {t : Set β}, MeasurableSet sMeasurableSet tMeasurable AMeasurable BProbabilityTheory.IndepFun A B (ProbabilityTheory.cond μ (A ⁻¹' s B ⁻¹' t))

If A is independent of B, then they remain independent when conditioning on an event of the form A ∈ s ∩ B ∈ t of positive probability.

def ProbabilityTheory.CondIndepFun {Ω : Type u_1} {α : Type u_3} {β : Type u_4} {γ : Type u_5} [MeasurableSpace Ω] [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] (f : Ωα) (g : Ωβ) (h : Ωγ) (μ : autoParam (MeasureTheory.Measure Ω) _auto✝) :

The assertion that f and g are conditionally independent relative to h.

Equations
Instances For
    theorem ProbabilityTheory.condIndepFun_iff {Ω : Type u_1} {α : Type u_3} {β : Type u_4} {γ : Type u_5} [MeasurableSpace Ω] [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] {μ : MeasureTheory.Measure Ω} {f : Ωα} {g : Ωβ} {h : Ωγ} :
    theorem ProbabilityTheory.CondIndepFun.comp_right {Ω : Type u_1} {Ω' : Type u_2} {α : Type u_3} {β : Type u_4} {γ : Type u_5} [MeasurableSpace Ω] [MeasurableSpace Ω'] [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] {μ : MeasureTheory.Measure Ω} {f : Ωα} {g : Ωβ} {h : Ωγ} [MeasurableSingletonClass γ] {i : Ω'Ω} (hi : MeasurableEmbedding i) (hi' : ∀ᵐ (a : Ω) ∂μ, a Set.range i) (hf : Measurable f) (hg : Measurable g) (hh : Measurable h) (hfg : ProbabilityTheory.CondIndepFun f g h μ) :

    Composing independent functions with a measurable embedding of conull range gives independent functions.

    theorem ProbabilityTheory.condIndep_copies {Ω : Type u_1} {α : Type u} {β : Type u} [MeasurableSpace Ω] [MeasurableSpace α] [MeasurableSpace β] [MeasurableSingletonClass β] [Countable β] (X : Ωα) (Y : Ωβ) (hX : Measurable X) (hY : Measurable Y) [finY : FiniteRange Y] (μ : MeasureTheory.Measure Ω) [MeasureTheory.IsProbabilityMeasure μ] :
    ∃ (Ω' : Type u) (mΩ' : MeasurableSpace Ω') (X₁ : Ω'α) (X₂ : Ω'α) (Y' : Ω'β) (ν : MeasureTheory.Measure Ω'), MeasureTheory.IsProbabilityMeasure ν Measurable X₁ Measurable X₂ Measurable Y' ProbabilityTheory.CondIndepFun X₁ X₂ Y' ν ProbabilityTheory.IdentDistrib (X₁, Y') (X, Y) ν μ ProbabilityTheory.IdentDistrib (X₂, Y') (X, Y) ν μ

    For X, Y random variables, there exist conditionally independent trials X_1, X_2, Y'.

    theorem ProbabilityTheory.condIndep_copies' {Ω : Type u_1} {α : Type u} {β : Type u} [MeasurableSpace Ω] [MeasurableSpace α] [MeasurableSpace β] [MeasurableSingletonClass β] [Countable β] (X : Ωα) (Y : Ωβ) (hX : Measurable X) (hY : Measurable Y) [FiniteRange Y] (μ : MeasureTheory.Measure Ω) [MeasureTheory.IsProbabilityMeasure μ] (p : αβProp) (hp : Measurable (Function.uncurry p)) (hp' : ∀ᵐ (ω : Ω) ∂μ, p (X ω) (Y ω)) :
    ∃ (Ω' : Type u) (mΩ' : MeasurableSpace Ω') (X₁ : Ω'α) (X₂ : Ω'α) (Y' : Ω'β) (ν : MeasureTheory.Measure Ω'), MeasureTheory.IsProbabilityMeasure ν Measurable X₁ Measurable X₂ Measurable Y' ProbabilityTheory.CondIndepFun X₁ X₂ Y' ν ProbabilityTheory.IdentDistrib (X₁, Y') (X, Y) ν μ ProbabilityTheory.IdentDistrib (X₂, Y') (X, Y) ν μ (∀ (ω : Ω'), p (X₁ ω) (Y' ω)) ∀ (ω : Ω'), p (X₂ ω) (Y' ω)

    For X, Y random variables, there exist conditionally independent trials X₁, X₂, Y'.