Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore(probability/kernel/composition): swap the order of the arguments of prod_mk_left #18929

Closed
wants to merge 1 commit into from
Closed
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
20 changes: 10 additions & 10 deletions src/probability/kernel/composition.lean
Original file line number Diff line number Diff line change
Expand Up @@ -461,32 +461,32 @@ open_locale probability_theory
section fst_snd

/-- Define a `kernel (γ × α) β` from a `kernel α β` by taking the comap of the projection. -/
def prod_mk_left (κ : kernel α β) (γ : Type*) [measurable_space γ] : kernel (γ × α) β :=
def prod_mk_left (γ : Type*) [measurable_space γ] (κ : kernel α β) : kernel (γ × α) β :=
comap κ prod.snd measurable_snd

variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → α}

include mγ

lemma prod_mk_left_apply (κ : kernel α β) (ca : γ × α) :
prod_mk_left κ γ ca = κ ca.snd := rfl
prod_mk_left γ κ ca = κ ca.snd := rfl

lemma prod_mk_left_apply' (κ : kernel α β) (ca : γ × α) (s : set β) :
prod_mk_left κ γ ca s = κ ca.snd s := rfl
prod_mk_left γ κ ca s = κ ca.snd s := rfl

lemma lintegral_prod_mk_left (κ : kernel α β) (ca : γ × α) (g : β → ℝ≥0∞) :
∫⁻ b, g b ∂(prod_mk_left κ γ ca) = ∫⁻ b, g b ∂(κ ca.snd) := rfl
∫⁻ b, g b ∂(prod_mk_left γ κ ca) = ∫⁻ b, g b ∂(κ ca.snd) := rfl

instance is_markov_kernel.prod_mk_left (κ : kernel α β) [is_markov_kernel κ] :
is_markov_kernel (prod_mk_left κ γ) :=
is_markov_kernel (prod_mk_left γ κ) :=
by { rw prod_mk_left, apply_instance, }

instance is_finite_kernel.prod_mk_left (κ : kernel α β) [is_finite_kernel κ] :
is_finite_kernel (prod_mk_left κ γ) :=
is_finite_kernel (prod_mk_left γ κ) :=
by { rw prod_mk_left, apply_instance, }

instance is_s_finite_kernel.prod_mk_left (κ : kernel α β) [is_s_finite_kernel κ] :
is_s_finite_kernel (prod_mk_left κ γ) :=
is_s_finite_kernel (prod_mk_left γ κ) :=
by { rw prod_mk_left, apply_instance, }

/-- Define a `kernel (β × α) γ` from a `kernel (α × β) γ` by taking the comap of `prod.swap`. -/
Expand Down Expand Up @@ -612,7 +612,7 @@ include mγ
noncomputable
def comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] :
kernel α γ :=
snd (κ ⊗ₖ prod_mk_left η α)
snd (κ ⊗ₖ prod_mk_left α η)

localized "infix (name := kernel.comp) ` ∘ₖ `:100 := probability_theory.kernel.comp" in
probability_theory
Expand All @@ -631,7 +631,7 @@ lemma lintegral_comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α
∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a) :=
begin
rw [comp, lintegral_snd _ _ hg],
change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left η α) a)
change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left α η) a)
= ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a),
exact lintegral_comp_prod _ _ _ (hg.comp measurable_snd),
end
Expand Down Expand Up @@ -690,7 +690,7 @@ include mγ
noncomputable
def prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] :
kernel α (β × γ) :=
κ ⊗ₖ (swap_left (prod_mk_left η β))
κ ⊗ₖ (swap_left (prod_mk_left β η))

lemma prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η]
(a : α) {s : set (β × γ)} (hs : measurable_set s) :
Expand Down