Theory PDF_Semantics
theory PDF_Semantics
imports PDF_Values
begin
lemma measurable_subprob_algebra_density:
assumes "sigma_finite_measure N"
assumes "space N ≠ {}"
assumes [measurable]: "case_prod f ∈ borel_measurable (M ⨂⇩M N)"
assumes "⋀x. x ∈ space M ⟹ (∫⇧+y. f x y ∂N) ≤ 1"
shows "(λx. density N (f x)) ∈ measurable M (subprob_algebra N)"
proof (rule measurable_subprob_algebra)
fix x assume "x ∈ space M"
with assms show "subprob_space (density N (f x))"
by (intro subprob_spaceI) (auto simp: emeasure_density cong: nn_integral_cong')
next
interpret sigma_finite_measure N by fact
fix X assume X: "X ∈ sets N"
hence "(λx. (∫⇧+y. f x y * indicator X y ∂N)) ∈ borel_measurable M" by simp
moreover from X and assms have
"⋀x. x ∈ space M ⟹ emeasure (density N (f x)) X = (∫⇧+y. f x y * indicator X y ∂N)"
by (simp add: emeasure_density)
ultimately show "(λx. emeasure (density N (f x)) X) ∈ borel_measurable M"
by (simp only: cong: measurable_cong)
qed simp_all
section ‹Built-in Probability Distributions›
subsection ‹Bernoulli›
definition bernoulli_density :: "real ⇒ bool ⇒ ennreal" where
"bernoulli_density p b = (if p ∈ {0..1} then (if b then p else 1 - p) else 0)"
definition bernoulli :: "val ⇒ val measure" where
"bernoulli p = density BOOL (bernoulli_density (extract_real p) o extract_bool)"
lemma measurable_bernoulli_density[measurable]:
"case_prod bernoulli_density ∈ borel_measurable (borel ⨂⇩M count_space UNIV)"
unfolding bernoulli_density_def[abs_def] by measurable
lemma measurable_bernoulli[measurable]: "bernoulli ∈ measurable REAL (subprob_algebra BOOL)"
unfolding bernoulli_def[abs_def]
by (auto intro!: measurable_subprob_algebra_density
simp: measurable_split_conv nn_integral_BoolVal bernoulli_density_def
ennreal_plus[symmetric]
simp del: ennreal_plus)
subsection ‹Uniform›
definition uniform_real_density :: "real × real ⇒ real ⇒ ennreal" where
"uniform_real_density ≡ λ(a,b) x. ennreal (if a < b ∧ x ∈ {a..b} then inverse (b - a) else 0)"
definition uniform_int_density :: "int × int ⇒ int ⇒ ennreal" where
"uniform_int_density ≡ λ(a,b) x. (if x ∈ {a..b} then inverse (nat (b - a + 1)) else 0)"
lemma measurable_uniform_density_int[measurable]:
"(case_prod uniform_int_density)
∈ borel_measurable ((count_space UNIV ⨂⇩M count_space UNIV) ⨂⇩M count_space UNIV)"
by (simp add: pair_measure_countable)
lemma measurable_uniform_density_real[measurable]:
"(case_prod uniform_real_density) ∈ borel_measurable (borel ⨂⇩M borel)"
proof-
have "(case_prod uniform_real_density) =
(λx. uniform_real_density (fst (fst x), snd (fst x)) (snd x))"
by (rule ext) (simp split: prod.split)
also have "... ∈ borel_measurable (borel ⨂⇩M borel)"
unfolding uniform_real_density_def
by (simp only: prod.case) (simp add: borel_prod[symmetric])
finally show ?thesis .
qed
definition uniform_int :: "val ⇒ val measure" where
"uniform_int = map_int_pair (λl u. density INTEG (uniform_int_density (l,u) o extract_int)) (λ_. undefined)"
definition uniform_real :: "val ⇒ val measure" where
"uniform_real = map_real_pair (λl u. density REAL (uniform_real_density (l,u) o extract_real)) (λ_. undefined)"
lemma if_bounded: "(if a ≤ i ∧ i ≤ b then v else 0) = (v::real) * indicator {a .. b} i"
by auto
lemma measurable_uniform_int[measurable]:
"uniform_int ∈ measurable (PRODUCT INTEG INTEG) (subprob_algebra INTEG)"
unfolding uniform_int_def
proof (rule measurable measurable_subprob_algebra_density)+
fix x :: "int × int"
show "integral⇧N INTEG (uniform_int_density (fst x, snd x) ∘ extract_int) ≤ 1"
proof cases
assume "fst x ≤ snd x" then show ?thesis
by (cases x)
(simp add: uniform_int_density_def comp_def nn_integral_IntVal nn_integral_cmult
nn_integral_set_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat
if_bounded[where 'a=int] ennreal_mult[symmetric]
del: ennreal_plus)
qed (simp add: uniform_int_density_def comp_def split_beta' if_bounded[where 'a=int])
qed (auto simp: comp_def)
lemma density_cong':
"(⋀x. x ∈ space M ⟹ f x = g x) ⟹ density M f = density M g"
unfolding density_def
by (auto dest: sets.sets_into_space intro!: nn_integral_cong measure_of_eq)
lemma measurable_uniform_real[measurable]:
"uniform_real ∈ measurable (PRODUCT REAL REAL) (subprob_algebra REAL)"
unfolding uniform_real_def
proof (rule measurable measurable_subprob_algebra_density)+
fix x :: "real × real"
obtain l u where [simp]: "x = (l, u)"
by (cases x) auto
show "(∫⇧+y. (uniform_real_density (fst x, snd x) o extract_real) y ∂REAL) ≤ 1"
proof cases
assume "l < u" then show ?thesis
by (simp add: nn_integral_RealVal uniform_real_density_def if_bounded nn_integral_cmult
nn_integral_set_ennreal[symmetric] ennreal_mult[symmetric])
qed (simp add: uniform_real_density_def comp_def)
qed (auto simp: comp_def borel_prod)
subsection ‹Gaussian›
definition gaussian_density :: "real × real ⇒ real ⇒ ennreal" where
"gaussian_density ≡
λ(m,s) x. (if s > 0 then exp (-(x - m)⇧2 / (2 * s⇧2)) / sqrt (2 * pi * s⇧2) else 0)"
lemma measurable_gaussian_density[measurable]:
"case_prod gaussian_density ∈ borel_measurable (borel ⨂⇩M borel)"
proof-
have "case_prod gaussian_density =
(λ(x,y). (if snd x > 0 then exp (-((y - fst x)^2) / (2 * snd x^2)) /
sqrt (2 * pi * snd x^2) else 0))"
unfolding gaussian_density_def by (intro ext) (simp split: prod.split)
also have "... ∈ borel_measurable (borel ⨂⇩M borel)"
by (simp add: borel_prod[symmetric])
finally show ?thesis .
qed
definition gaussian :: "val ⇒ val measure" where
"gaussian = map_real_pair (λm s. density REAL (gaussian_density (m,s) o extract_real)) undefined"
lemma measurable_gaussian[measurable]: "gaussian ∈ measurable (PRODUCT REAL REAL) (subprob_algebra REAL)"
unfolding gaussian_def
proof (rule measurable measurable_subprob_algebra_density)+
fix x :: "real × real"
show "integral⇧N (stock_measure REAL) (gaussian_density (fst x, snd x) ∘ extract_real) ≤ 1"
proof cases
assume "snd x > 0"
then have "integral⇧N lborel (gaussian_density x) = (∫⇧+y. normal_density (fst x) (snd x) y ∂lborel)"
by (auto simp add: gaussian_density_def normal_density_def split_beta' intro!: nn_integral_cong)
also have "… = 1"
using ‹snd x > 0›
by (subst nn_integral_eq_integral) (auto intro!: normal_density_nonneg)
finally show ?thesis
by (cases x) (simp add: nn_integral_RealVal comp_def)
next
assume "¬ snd x > 0" then show ?thesis
by (cases x)
(simp add: nn_integral_RealVal comp_def gaussian_density_def zero_ennreal_def[symmetric])
qed
qed (auto simp: comp_def borel_prod)
subsection ‹Poisson›
definition poisson_density' :: "real ⇒ int ⇒ ennreal" where
"poisson_density' rate k = pmf (poisson_pmf rate) (nat k) * indicator ({0 <..} × {0..}) (rate, k)"
lemma measurable_poisson_density'[measurable]:
"case_prod poisson_density' ∈ borel_measurable (borel ⨂⇩M count_space UNIV)"
proof -
have "case_prod poisson_density' =
(λ(rate, k). rate ^ nat k / real_of_nat (fact (nat k)) * exp (-rate) * indicator ({0 <..} × {0..}) (rate, k))"
by (auto split: split_indicator simp: fun_eq_iff poisson_density'_def)
then show ?thesis
by simp
qed
definition poisson :: "val ⇒ val measure" where
"poisson rate = density INTEG (poisson_density' (extract_real rate) o extract_int)"
lemma measurable_poisson[measurable]: "poisson ∈ measurable REAL (subprob_algebra INTEG)"
unfolding poisson_def[abs_def]
proof (rule measurable measurable_subprob_algebra_density)+
fix r :: real
have [simp]: "nat ` {0..} = UNIV"
by (auto simp: image_iff intro!: bexI[of _ "int x" for x])
{ assume "0 < r"
then have "(∫⇧+ x. ennreal (r ^ nat x * exp (- r) * indicator ({0<..} × {0..}) (r, x) / (fact (nat x))) ∂count_space UNIV)
= (∫⇧+ x. ennreal (pmf (poisson_pmf r) (nat x)) ∂count_space {0 ..})"
by (auto intro!: nn_integral_cong simp add: nn_integral_count_space_indicator split: split_indicator)
also have "… = 1"
using measure_pmf.emeasure_space_1[of "poisson_pmf r"]
by (subst nn_integral_pmf') (auto simp: inj_on_def)
finally have "(∫⇧+ x. ennreal (r ^ nat x * exp (- r) * indicator ({0<..} × {0..}) (r, x) / (fact (nat x))) ∂count_space UNIV) = 1"
. }
then show "integral⇧N INTEG (poisson_density' r ∘ extract_int) ≤ 1"
by (cases "0 < r")
(auto simp: nn_integral_IntVal poisson_density'_def zero_ennreal_def[symmetric])
qed (auto simp: comp_def)
section ‹Source Language Syntax and Semantics›
subsection ‹Expressions›
class expr = fixes free_vars :: "'a ⇒ vname set"