| author | wenzelm | 
| Wed, 23 Aug 2023 11:00:30 +0200 | |
| changeset 78568 | a97d2b6b5c3e | 
| parent 75607 | 3c544d64c218 | 
| child 78517 | 28c1f4f5335f | 
| permissions | -rw-r--r-- | 
| 58606 | 1 | (* Title: HOL/Probability/Giry_Monad.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | Author: Manuel Eberl, TU München | |
| 4 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 5 | Defines subprobability spaces, the subprobability functor and the Giry monad on subprobability | 
| 58606 | 6 | spaces. | 
| 7 | *) | |
| 8 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 9 | section \<open>The Giry monad\<close> | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 10 | |
| 58606 | 11 | theory Giry_Monad | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64320diff
changeset | 12 | imports Probability_Measure "HOL-Library.Monad_Syntax" | 
| 58606 | 13 | begin | 
| 14 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 15 | subsection \<open>Sub-probability spaces\<close> | 
| 58606 | 16 | |
| 17 | locale subprob_space = finite_measure + | |
| 18 | assumes emeasure_space_le_1: "emeasure M (space M) \<le> 1" | |
| 19 |   assumes subprob_not_empty: "space M \<noteq> {}"
 | |
| 20 | ||
| 21 | lemma subprob_spaceI[Pure.intro!]: | |
| 22 | assumes *: "emeasure M (space M) \<le> 1" | |
| 23 |   assumes "space M \<noteq> {}"
 | |
| 24 | shows "subprob_space M" | |
| 25 | proof - | |
| 26 | interpret finite_measure M | |
| 27 | proof | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 28 | show "emeasure M (space M) \<noteq> \<infinity>" using * by (auto simp: top_unique) | 
| 58606 | 29 | qed | 
| 61169 | 30 | show "subprob_space M" by standard fact+ | 
| 58606 | 31 | qed | 
| 32 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 33 | lemma (in subprob_space) emeasure_subprob_space_less_top: "emeasure M A \<noteq> top" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 34 | using emeasure_finite[of A] . | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 35 | |
| 58606 | 36 | lemma prob_space_imp_subprob_space: | 
| 37 | "prob_space M \<Longrightarrow> subprob_space M" | |
| 38 | by (rule subprob_spaceI) (simp_all add: prob_space.emeasure_space_1 prob_space.not_empty) | |
| 39 | ||
| 59425 | 40 | lemma subprob_space_imp_sigma_finite: "subprob_space M \<Longrightarrow> sigma_finite_measure M" | 
| 41 | unfolding subprob_space_def finite_measure_def by simp | |
| 42 | ||
| 58606 | 43 | sublocale prob_space \<subseteq> subprob_space | 
| 44 | by (rule subprob_spaceI) (simp_all add: emeasure_space_1 not_empty) | |
| 45 | ||
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 46 | lemma subprob_space_sigma [simp]: "\<Omega> \<noteq> {} \<Longrightarrow> subprob_space (sigma \<Omega> X)"
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 47 | by(rule subprob_spaceI)(simp_all add: emeasure_sigma space_measure_of_conv) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 48 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 49 | lemma subprob_space_null_measure: "space M \<noteq> {} \<Longrightarrow> subprob_space (null_measure M)"
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 50 | by(simp add: null_measure_def) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 51 | |
| 58606 | 52 | lemma (in subprob_space) subprob_space_distr: | 
| 53 |   assumes f: "f \<in> measurable M M'" and "space M' \<noteq> {}" shows "subprob_space (distr M M' f)"
 | |
| 54 | proof (rule subprob_spaceI) | |
| 55 | have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space) | |
| 56 | with f show "emeasure (distr M M' f) (space (distr M M' f)) \<le> 1" | |
| 57 | by (auto simp: emeasure_distr emeasure_space_le_1) | |
| 58 |   show "space (distr M M' f) \<noteq> {}" by (simp add: assms)
 | |
| 59 | qed | |
| 60 | ||
| 59000 | 61 | lemma (in subprob_space) subprob_emeasure_le_1: "emeasure M X \<le> 1" | 
| 58606 | 62 | by (rule order.trans[OF emeasure_space emeasure_space_le_1]) | 
| 63 | ||
| 59000 | 64 | lemma (in subprob_space) subprob_measure_le_1: "measure M X \<le> 1" | 
| 65 | using subprob_emeasure_le_1[of X] by (simp add: emeasure_eq_measure) | |
| 66 | ||
| 59427 | 67 | lemma (in subprob_space) nn_integral_le_const: | 
| 68 | assumes "0 \<le> c" "AE x in M. f x \<le> c" | |
| 69 | shows "(\<integral>\<^sup>+x. f x \<partial>M) \<le> c" | |
| 70 | proof - | |
| 71 | have "(\<integral>\<^sup>+ x. f x \<partial>M) \<le> (\<integral>\<^sup>+ x. c \<partial>M)" | |
| 72 | by(rule nn_integral_mono_AE) fact | |
| 73 | also have "\<dots> \<le> c * emeasure M (space M)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 74 | using \<open>0 \<le> c\<close> by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 75 | also have "\<dots> \<le> c * 1" using emeasure_space_le_1 \<open>0 \<le> c\<close> by(rule mult_left_mono) | 
| 59427 | 76 | finally show ?thesis by simp | 
| 77 | qed | |
| 78 | ||
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 79 | lemma emeasure_density_distr_interval: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 80 | fixes h :: "real \<Rightarrow> real" and g :: "real \<Rightarrow> real" and g' :: "real \<Rightarrow> real" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 81 | assumes [simp]: "a \<le> b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 82 | assumes Mf[measurable]: "f \<in> borel_measurable borel" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 83 | assumes Mg[measurable]: "g \<in> borel_measurable borel" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 84 | assumes Mg'[measurable]: "g' \<in> borel_measurable borel" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 85 | assumes Mh[measurable]: "h \<in> borel_measurable borel" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 86 | assumes prob: "subprob_space (density lborel f)" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 87 | assumes nonnegf: "\<And>x. f x \<ge> 0" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 88 |   assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 89 |   assumes contg': "continuous_on {a..b} g'"
 | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
73253diff
changeset | 90 |   assumes mono: "strict_mono_on {a..b} g" and inv: "\<And>x. h x \<in> {a..b} \<Longrightarrow> g (h x) = x"
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 91 |   assumes range: "{a..b} \<subseteq> range h"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 92 |   shows "emeasure (distr (density lborel f) lborel h) {a..b} =
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 93 |              emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 94 | proof (cases "a < b") | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 95 | assume "a < b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 96 |   from mono have inj: "inj_on g {a..b}" by (rule strict_mono_on_imp_inj_on)
 | 
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
73253diff
changeset | 97 |   from mono have mono': "mono_on {a..b} g" by (rule strict_mono_on_imp_mono_on)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 98 |   from mono' derivg have "\<And>x. x \<in> {a<..<b} \<Longrightarrow> g' x \<ge> 0"
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 99 | by (rule mono_on_imp_deriv_nonneg) auto | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 100 |   from contg' this have derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
 | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 101 | by (rule continuous_ge_on_Ioo) (simp_all add: \<open>a < b\<close>) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 102 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 103 |   from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 104 |   have A: "h -` {a..b} = {g a..g b}"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 105 | proof (intro equalityI subsetI) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 106 |     fix x assume x: "x \<in> h -` {a..b}"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 107 |     hence "g (h x) \<in> {g a..g b}" by (auto intro: mono_onD[OF mono'])
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 108 |     with inv and x show "x \<in> {g a..g b}" by simp
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 109 | next | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 110 |     fix y assume y: "y \<in> {g a..g b}"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 111 |     with IVT'[OF _ _ _ contg, of y] obtain x where "x \<in> {a..b}" "y = g x" by auto
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 112 |     with range and inv show "y \<in> h -` {a..b}" by auto
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 113 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 114 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 115 | have prob': "subprob_space (distr (density lborel f) lborel h)" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 116 | by (rule subprob_space.subprob_space_distr[OF prob]) (simp_all add: Mh) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 117 |   have B: "emeasure (distr (density lborel f) lborel h) {a..b} =
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 118 |             \<integral>\<^sup>+x. f x * indicator (h -` {a..b}) x \<partial>lborel"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 119 | by (subst emeasure_distr) (simp_all add: emeasure_density Mf Mh measurable_sets_borel[OF Mh]) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 120 | also note A | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 121 |   also have "emeasure (distr (density lborel f) lborel h) {a..b} \<le> 1"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 122 | by (rule subprob_space.subprob_emeasure_le_1) (rule prob') | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 123 |   hence "emeasure (distr (density lborel f) lborel h) {a..b} \<noteq> \<infinity>" by (auto simp: top_unique)
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 124 |   with assms have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 125 |                       (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 126 | by (intro nn_integral_substitution_aux) | 
| 61808 | 127 | (auto simp: derivg_nonneg A B emeasure_density mult.commute \<open>a < b\<close>) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 128 |   also have "... = emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}"
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 129 | by (simp add: emeasure_density) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 130 | finally show ?thesis . | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 131 | next | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 132 | assume "\<not>a < b" | 
| 61808 | 133 | with \<open>a \<le> b\<close> have [simp]: "b = a" by (simp add: not_less del: \<open>a \<le> b\<close>) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 134 |   from inv and range have "h -` {a} = {g a}" by auto
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 135 | thus ?thesis by (simp_all add: emeasure_distr emeasure_density measurable_sets_borel[OF Mh]) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 136 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59048diff
changeset | 137 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 138 | locale pair_subprob_space = | 
| 58606 | 139 | pair_sigma_finite M1 M2 + M1: subprob_space M1 + M2: subprob_space M2 for M1 M2 | 
| 140 | ||
| 61565 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 ballarin parents: 
61424diff
changeset | 141 | sublocale pair_subprob_space \<subseteq> P?: subprob_space "M1 \<Otimes>\<^sub>M M2" | 
| 58606 | 142 | proof | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 143 | from mult_le_one[OF M1.emeasure_space_le_1 _ M2.emeasure_space_le_1] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 144 | show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) \<le> 1" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 145 | by (simp add: M2.emeasure_pair_measure_Times space_pair_measure) | 
| 58606 | 146 |   from M1.subprob_not_empty and M2.subprob_not_empty show "space (M1 \<Otimes>\<^sub>M M2) \<noteq> {}"
 | 
| 147 | by (simp add: space_pair_measure) | |
| 148 | qed | |
| 149 | ||
| 59425 | 150 | lemma subprob_space_null_measure_iff: | 
| 151 |     "subprob_space (null_measure M) \<longleftrightarrow> space M \<noteq> {}"
 | |
| 152 | by (auto intro!: subprob_spaceI dest: subprob_space.subprob_not_empty) | |
| 153 | ||
| 59525 | 154 | lemma subprob_space_restrict_space: | 
| 155 | assumes M: "subprob_space M" | |
| 156 |   and A: "A \<inter> space M \<in> sets M" "A \<inter> space M \<noteq> {}"
 | |
| 157 | shows "subprob_space (restrict_space M A)" | |
| 158 | proof(rule subprob_spaceI) | |
| 159 | have "emeasure (restrict_space M A) (space (restrict_space M A)) = emeasure M (A \<inter> space M)" | |
| 160 | using A by(simp add: emeasure_restrict_space space_restrict_space) | |
| 161 | also have "\<dots> \<le> 1" by(rule subprob_space.subprob_emeasure_le_1)(rule M) | |
| 162 | finally show "emeasure (restrict_space M A) (space (restrict_space M A)) \<le> 1" . | |
| 163 | next | |
| 164 |   show "space (restrict_space M A) \<noteq> {}"
 | |
| 165 | using A by(simp add: space_restrict_space) | |
| 166 | qed | |
| 167 | ||
| 58606 | 168 | definition subprob_algebra :: "'a measure \<Rightarrow> 'a measure measure" where | 
| 169 | "subprob_algebra K = | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67649diff
changeset | 170 |     (SUP A \<in> sets K. vimage_algebra {M. subprob_space M \<and> sets M = sets K} (\<lambda>M. emeasure M A) borel)"
 | 
| 58606 | 171 | |
| 172 | lemma space_subprob_algebra: "space (subprob_algebra A) = {M. subprob_space M \<and> sets M = sets A}"
 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63092diff
changeset | 173 | by (auto simp add: subprob_algebra_def space_Sup_eq_UN) | 
| 58606 | 174 | |
| 175 | lemma subprob_algebra_cong: "sets M = sets N \<Longrightarrow> subprob_algebra M = subprob_algebra N" | |
| 176 | by (simp add: subprob_algebra_def) | |
| 177 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 178 | lemma measurable_emeasure_subprob_algebra[measurable]: | 
| 58606 | 179 | "a \<in> sets A \<Longrightarrow> (\<lambda>M. emeasure M a) \<in> borel_measurable (subprob_algebra A)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63092diff
changeset | 180 | by (auto intro!: measurable_Sup1 measurable_vimage_algebra1 simp: subprob_algebra_def) | 
| 58606 | 181 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 182 | lemma measurable_measure_subprob_algebra[measurable]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 183 | "a \<in> sets A \<Longrightarrow> (\<lambda>M. measure M a) \<in> borel_measurable (subprob_algebra A)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 184 | unfolding measure_def by measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 185 | |
| 59000 | 186 | lemma subprob_measurableD: | 
| 187 | assumes N: "N \<in> measurable M (subprob_algebra S)" and x: "x \<in> space M" | |
| 188 | shows "space (N x) = space S" | |
| 189 | and "sets (N x) = sets S" | |
| 190 | and "measurable (N x) K = measurable S K" | |
| 191 | and "measurable K (N x) = measurable K S" | |
| 192 | using measurable_space[OF N x] | |
| 193 | by (auto simp: space_subprob_algebra intro!: measurable_cong_sets dest: sets_eq_imp_space_eq) | |
| 194 | ||
| 61808 | 195 | ML \<open> | 
| 59048 | 196 | |
| 197 | fun subprob_cong thm ctxt = ( | |
| 198 | let | |
| 67649 | 199 | val thm' = Thm.transfer' ctxt thm | 
| 59582 | 200 | val free = thm' |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_comb |> fst |> | 
| 59048 | 201 | dest_comb |> snd |> strip_abs_body |> head_of |> is_Free | 
| 202 | in | |
| 203 |     if free then ([], Measurable.add_local_cong (thm' RS @{thm subprob_measurableD(2)}) ctxt)
 | |
| 204 | else ([], ctxt) | |
| 205 | end | |
| 206 | handle THM _ => ([], ctxt) | TERM _ => ([], ctxt)) | |
| 207 | ||
| 61808 | 208 | \<close> | 
| 59048 | 209 | |
| 210 | setup \<open> | |
| 211 | Context.theory_map (Measurable.add_preprocessor "subprob_cong" subprob_cong) | |
| 212 | \<close> | |
| 213 | ||
| 58606 | 214 | context | 
| 215 | fixes K M N assumes K: "K \<in> measurable M (subprob_algebra N)" | |
| 216 | begin | |
| 217 | ||
| 218 | lemma subprob_space_kernel: "a \<in> space M \<Longrightarrow> subprob_space (K a)" | |
| 219 | using measurable_space[OF K] by (simp add: space_subprob_algebra) | |
| 220 | ||
| 221 | lemma sets_kernel: "a \<in> space M \<Longrightarrow> sets (K a) = sets N" | |
| 222 | using measurable_space[OF K] by (simp add: space_subprob_algebra) | |
| 223 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 224 | lemma measurable_emeasure_kernel[measurable]: | 
| 58606 | 225 | "A \<in> sets N \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M" | 
| 226 | using measurable_compose[OF K measurable_emeasure_subprob_algebra] . | |
| 227 | ||
| 228 | end | |
| 229 | ||
| 230 | lemma measurable_subprob_algebra: | |
| 231 | "(\<And>a. a \<in> space M \<Longrightarrow> subprob_space (K a)) \<Longrightarrow> | |
| 232 | (\<And>a. a \<in> space M \<Longrightarrow> sets (K a) = sets N) \<Longrightarrow> | |
| 233 | (\<And>A. A \<in> sets N \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M) \<Longrightarrow> | |
| 234 | K \<in> measurable M (subprob_algebra N)" | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63092diff
changeset | 235 | by (auto intro!: measurable_Sup2 measurable_vimage_algebra2 simp: subprob_algebra_def) | 
| 58606 | 236 | |
| 59778 | 237 | lemma measurable_submarkov: | 
| 238 | "K \<in> measurable M (subprob_algebra M) \<longleftrightarrow> | |
| 239 | (\<forall>x\<in>space M. subprob_space (K x) \<and> sets (K x) = sets M) \<and> | |
| 240 | (\<forall>A\<in>sets M. (\<lambda>x. emeasure (K x) A) \<in> measurable M borel)" | |
| 241 | proof | |
| 242 | assume "(\<forall>x\<in>space M. subprob_space (K x) \<and> sets (K x) = sets M) \<and> | |
| 243 | (\<forall>A\<in>sets M. (\<lambda>x. emeasure (K x) A) \<in> borel_measurable M)" | |
| 244 | then show "K \<in> measurable M (subprob_algebra M)" | |
| 245 | by (intro measurable_subprob_algebra) auto | |
| 246 | next | |
| 247 | assume "K \<in> measurable M (subprob_algebra M)" | |
| 248 | then show "(\<forall>x\<in>space M. subprob_space (K x) \<and> sets (K x) = sets M) \<and> | |
| 249 | (\<forall>A\<in>sets M. (\<lambda>x. emeasure (K x) A) \<in> borel_measurable M)" | |
| 250 | by (auto dest: subprob_space_kernel sets_kernel) | |
| 251 | qed | |
| 252 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 253 | lemma measurable_subprob_algebra_generated: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 254 | assumes eq: "sets N = sigma_sets \<Omega> G" and "Int_stable G" "G \<subseteq> Pow \<Omega>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 255 | assumes subsp: "\<And>a. a \<in> space M \<Longrightarrow> subprob_space (K a)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 256 | assumes sets: "\<And>a. a \<in> space M \<Longrightarrow> sets (K a) = sets N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 257 | assumes "\<And>A. A \<in> G \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 258 | assumes \<Omega>: "(\<lambda>a. emeasure (K a) \<Omega>) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 259 | shows "K \<in> measurable M (subprob_algebra N)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 260 | proof (rule measurable_subprob_algebra) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 261 | fix a assume "a \<in> space M" then show "subprob_space (K a)" "sets (K a) = sets N" by fact+ | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 262 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 263 | interpret G: sigma_algebra \<Omega> "sigma_sets \<Omega> G" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 264 | using \<open>G \<subseteq> Pow \<Omega>\<close> by (rule sigma_algebra_sigma_sets) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 265 | fix A assume "A \<in> sets N" with assms(2,3) show "(\<lambda>a. emeasure (K a) A) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 266 | unfolding \<open>sets N = sigma_sets \<Omega> G\<close> | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 267 | proof (induction rule: sigma_sets_induct_disjoint) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 268 | case (basic A) then show ?case by fact | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 269 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 270 | case empty then show ?case by simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 271 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 272 | case (compl A) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 273 | have "(\<lambda>a. emeasure (K a) (\<Omega> - A)) \<in> borel_measurable M \<longleftrightarrow> | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 274 | (\<lambda>a. emeasure (K a) \<Omega> - emeasure (K a) A) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 275 | using G.top G.sets_into_space sets eq compl subprob_space.emeasure_subprob_space_less_top[OF subsp] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 276 | by (intro measurable_cong emeasure_Diff) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 277 | with compl \<Omega> show ?case | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 278 | by simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 279 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 280 | case (union F) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 281 | moreover have "(\<lambda>a. emeasure (K a) (\<Union>i. F i)) \<in> borel_measurable M \<longleftrightarrow> | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 282 | (\<lambda>a. \<Sum>i. emeasure (K a) (F i)) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 283 | using sets union eq | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 284 | by (intro measurable_cong suminf_emeasure[symmetric]) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 285 | ultimately show ?case | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 286 | by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 287 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 288 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 289 | |
| 58606 | 290 | lemma space_subprob_algebra_empty_iff: | 
| 291 |   "space (subprob_algebra N) = {} \<longleftrightarrow> space N = {}"
 | |
| 292 | proof | |
| 293 | have "\<And>x. x \<in> space N \<Longrightarrow> density N (\<lambda>_. 0) \<in> space (subprob_algebra N)" | |
| 294 | by (auto simp: space_subprob_algebra emeasure_density intro!: subprob_spaceI) | |
| 295 |   then show "space (subprob_algebra N) = {} \<Longrightarrow> space N = {}"
 | |
| 296 | by auto | |
| 297 | next | |
| 298 |   assume "space N = {}"
 | |
| 299 |   hence "sets N = {{}}" by (simp add: space_empty_iff)
 | |
| 300 |   moreover have "\<And>M. subprob_space M \<Longrightarrow> sets M \<noteq> {{}}"
 | |
| 301 | by (simp add: subprob_space.subprob_not_empty space_empty_iff[symmetric]) | |
| 302 |   ultimately show "space (subprob_algebra N) = {}" by (auto simp: space_subprob_algebra)
 | |
| 303 | qed | |
| 304 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 305 | lemma nn_integral_measurable_subprob_algebra[measurable]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 306 | assumes f: "f \<in> borel_measurable N" | 
| 59000 | 307 | shows "(\<lambda>M. integral\<^sup>N M f) \<in> borel_measurable (subprob_algebra N)" (is "_ \<in> ?B") | 
| 308 | using f | |
| 309 | proof induct | |
| 310 | case (cong f g) | |
| 311 | moreover have "(\<lambda>M'. \<integral>\<^sup>+M''. f M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. \<integral>\<^sup>+M''. g M'' \<partial>M') \<in> ?B" | |
| 312 | by (intro measurable_cong nn_integral_cong cong) | |
| 313 | (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq) | |
| 314 | ultimately show ?case by simp | |
| 315 | next | |
| 316 | case (set B) | |
| 63540 | 317 | then have "(\<lambda>M'. \<integral>\<^sup>+M''. indicator B M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. emeasure M' B) \<in> ?B" | 
| 59000 | 318 | by (intro measurable_cong nn_integral_indicator) (simp add: space_subprob_algebra) | 
| 63540 | 319 | with set show ?case | 
| 59000 | 320 | by (simp add: measurable_emeasure_subprob_algebra) | 
| 321 | next | |
| 322 | case (mult f c) | |
| 63540 | 323 | then have "(\<lambda>M'. \<integral>\<^sup>+M''. c * f M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. c * \<integral>\<^sup>+M''. f M'' \<partial>M') \<in> ?B" | 
| 59048 | 324 | by (intro measurable_cong nn_integral_cmult) (auto simp add: space_subprob_algebra) | 
| 63540 | 325 | with mult show ?case | 
| 59000 | 326 | by simp | 
| 327 | next | |
| 328 | case (add f g) | |
| 63540 | 329 | then have "(\<lambda>M'. \<integral>\<^sup>+M''. f M'' + g M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. (\<integral>\<^sup>+M''. f M'' \<partial>M') + (\<integral>\<^sup>+M''. g M'' \<partial>M')) \<in> ?B" | 
| 59048 | 330 | by (intro measurable_cong nn_integral_add) (auto simp add: space_subprob_algebra) | 
| 63540 | 331 | with add show ?case | 
| 59000 | 332 | by (simp add: ac_simps) | 
| 333 | next | |
| 334 | case (seq F) | |
| 63540 | 335 | then have "(\<lambda>M'. \<integral>\<^sup>+M''. (SUP i. F i) M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. SUP i. (\<integral>\<^sup>+M''. F i M'' \<partial>M')) \<in> ?B" | 
| 59000 | 336 | unfolding SUP_apply | 
| 59048 | 337 | by (intro measurable_cong nn_integral_monotone_convergence_SUP) (auto simp add: space_subprob_algebra) | 
| 63540 | 338 | with seq show ?case | 
| 59000 | 339 | by (simp add: ac_simps) | 
| 340 | qed | |
| 341 | ||
| 58606 | 342 | lemma measurable_distr: | 
| 343 | assumes [measurable]: "f \<in> measurable M N" | |
| 344 | shows "(\<lambda>M'. distr M' N f) \<in> measurable (subprob_algebra M) (subprob_algebra N)" | |
| 345 | proof (cases "space N = {}")
 | |
| 346 |   assume not_empty: "space N \<noteq> {}"
 | |
| 347 | show ?thesis | |
| 348 | proof (rule measurable_subprob_algebra) | |
| 349 | fix A assume A: "A \<in> sets N" | |
| 350 | then have "(\<lambda>M'. emeasure (distr M' N f) A) \<in> borel_measurable (subprob_algebra M) \<longleftrightarrow> | |
| 351 | (\<lambda>M'. emeasure M' (f -` A \<inter> space M)) \<in> borel_measurable (subprob_algebra M)" | |
| 352 | by (intro measurable_cong) | |
| 59048 | 353 | (auto simp: emeasure_distr space_subprob_algebra | 
| 67399 | 354 | intro!: arg_cong2[where f=emeasure] sets_eq_imp_space_eq arg_cong2[where f="(\<inter>)"]) | 
| 58606 | 355 | also have "\<dots>" | 
| 356 | using A by (intro measurable_emeasure_subprob_algebra) simp | |
| 357 | finally show "(\<lambda>M'. emeasure (distr M' N f) A) \<in> borel_measurable (subprob_algebra M)" . | |
| 59048 | 358 | qed (auto intro!: subprob_space.subprob_space_distr simp: space_subprob_algebra not_empty cong: measurable_cong_sets) | 
| 58606 | 359 | qed (insert assms, auto simp: measurable_empty_iff space_subprob_algebra_empty_iff) | 
| 360 | ||
| 59000 | 361 | lemma emeasure_space_subprob_algebra[measurable]: | 
| 362 | "(\<lambda>a. emeasure a (space a)) \<in> borel_measurable (subprob_algebra N)" | |
| 363 | proof- | |
| 364 | have "(\<lambda>a. emeasure a (space N)) \<in> borel_measurable (subprob_algebra N)" (is "?f \<in> ?M") | |
| 365 | by (rule measurable_emeasure_subprob_algebra) simp | |
| 366 | also have "?f \<in> ?M \<longleftrightarrow> (\<lambda>a. emeasure a (space a)) \<in> ?M" | |
| 367 | by (rule measurable_cong) (auto simp: space_subprob_algebra dest: sets_eq_imp_space_eq) | |
| 368 | finally show ?thesis . | |
| 369 | qed | |
| 370 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 371 | lemma integrable_measurable_subprob_algebra[measurable]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 372 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 373 | assumes [measurable]: "f \<in> borel_measurable N" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 374 | shows "Measurable.pred (subprob_algebra N) (\<lambda>M. integrable M f)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 375 | proof (rule measurable_cong[THEN iffD2]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 376 | show "M \<in> space (subprob_algebra N) \<Longrightarrow> integrable M f \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" for M | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 377 | by (auto simp: space_subprob_algebra integrable_iff_bounded) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 378 | qed measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 379 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 380 | lemma integral_measurable_subprob_algebra[measurable]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 381 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 382 | assumes f [measurable]: "f \<in> borel_measurable N" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 383 | shows "(\<lambda>M. integral\<^sup>L M f) \<in> subprob_algebra N \<rightarrow>\<^sub>M borel" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 384 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 385 | from borel_measurable_implies_sequence_metric[OF f, of 0] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 386 | obtain F where F: "\<And>i. simple_function N (F i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 387 | "\<And>x. x \<in> space N \<Longrightarrow> (\<lambda>i. F i x) \<longlonglongrightarrow> f x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 388 | "\<And>i x. x \<in> space N \<Longrightarrow> norm (F i x) \<le> 2 * norm (f x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 389 | unfolding norm_conv_dist by blast | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 390 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 391 | have [measurable]: "F i \<in> N \<rightarrow>\<^sub>M count_space UNIV" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 392 | using F(1) by (rule measurable_simple_function) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 393 | |
| 63040 | 394 | define F' where [abs_def]: | 
| 395 | "F' M i = (if integrable M f then integral\<^sup>L M (F i) else 0)" for M i | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 396 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 397 | have "(\<lambda>M. F' M i) \<in> subprob_algebra N \<rightarrow>\<^sub>M borel" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 398 | proof (rule measurable_cong[THEN iffD2]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 399 | fix M assume "M \<in> space (subprob_algebra N)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 400 | then have [simp]: "sets M = sets N" "space M = space N" "subprob_space M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 401 | by (auto simp: space_subprob_algebra intro!: sets_eq_imp_space_eq) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 402 | interpret subprob_space M by fact | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 403 | have "F' M i = (if integrable M f then Bochner_Integration.simple_bochner_integral M (F i) else 0)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 404 | using F(1) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 405 | by (subst simple_bochner_integrable_eq_integral) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 406 | (auto simp: simple_bochner_integrable.simps simple_function_def F'_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 407 |     then show "F' M i = (if integrable M f then \<Sum>y\<in>F i ` space N. measure M {x\<in>space N. F i x = y} *\<^sub>R y else 0)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 408 | unfolding simple_bochner_integral_def by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 409 | qed measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 410 | moreover | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 411 | have "F' M \<longlonglongrightarrow> integral\<^sup>L M f" if M: "M \<in> space (subprob_algebra N)" for M | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 412 | proof cases | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 413 | from M have [simp]: "sets M = sets N" "space M = space N" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 414 | by (auto simp: space_subprob_algebra intro!: sets_eq_imp_space_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 415 | assume "integrable M f" then show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 416 | unfolding F'_def using F(1)[THEN borel_measurable_simple_function] F | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 417 | by (auto intro!: integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 418 | cong: measurable_cong_sets) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 419 | qed (auto simp: F'_def not_integrable_integral_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 420 | ultimately show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 421 | by (rule borel_measurable_LIMSEQ_metric) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 422 | qed | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 423 | |
| 59978 | 424 | (* TODO: Rename. This name is too general -- Manuel *) | 
| 59000 | 425 | lemma measurable_pair_measure: | 
| 426 | assumes f: "f \<in> measurable M (subprob_algebra N)" | |
| 427 | assumes g: "g \<in> measurable M (subprob_algebra L)" | |
| 428 | shows "(\<lambda>x. f x \<Otimes>\<^sub>M g x) \<in> measurable M (subprob_algebra (N \<Otimes>\<^sub>M L))" | |
| 429 | proof (rule measurable_subprob_algebra) | |
| 430 |   { fix x assume "x \<in> space M"
 | |
| 431 | with measurable_space[OF f] measurable_space[OF g] | |
| 432 | have fx: "f x \<in> space (subprob_algebra N)" and gx: "g x \<in> space (subprob_algebra L)" | |
| 433 | by auto | |
| 434 | interpret F: subprob_space "f x" | |
| 435 | using fx by (simp add: space_subprob_algebra) | |
| 436 | interpret G: subprob_space "g x" | |
| 437 | using gx by (simp add: space_subprob_algebra) | |
| 438 | ||
| 439 | interpret pair_subprob_space "f x" "g x" .. | |
| 440 | show "subprob_space (f x \<Otimes>\<^sub>M g x)" by unfold_locales | |
| 441 | show sets_eq: "sets (f x \<Otimes>\<^sub>M g x) = sets (N \<Otimes>\<^sub>M L)" | |
| 442 | using fx gx by (simp add: space_subprob_algebra) | |
| 443 | ||
| 444 | have 1: "\<And>A B. A \<in> sets N \<Longrightarrow> B \<in> sets L \<Longrightarrow> emeasure (f x \<Otimes>\<^sub>M g x) (A \<times> B) = emeasure (f x) A * emeasure (g x) B" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 445 | using fx gx by (intro G.emeasure_pair_measure_Times) (auto simp: space_subprob_algebra) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 446 | have "emeasure (f x \<Otimes>\<^sub>M g x) (space (f x \<Otimes>\<^sub>M g x)) = | 
| 59000 | 447 | emeasure (f x) (space (f x)) * emeasure (g x) (space (g x))" | 
| 448 | by (subst G.emeasure_pair_measure_Times[symmetric]) (simp_all add: space_pair_measure) | |
| 449 | hence 2: "\<And>A. A \<in> sets (N \<Otimes>\<^sub>M L) \<Longrightarrow> emeasure (f x \<Otimes>\<^sub>M g x) (space N \<times> space L - A) = | |
| 450 | ... - emeasure (f x \<Otimes>\<^sub>M g x) A" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 451 | using emeasure_compl[simplified, OF _ P.emeasure_finite] | 
| 59000 | 452 | unfolding sets_eq | 
| 453 | unfolding sets_eq_imp_space_eq[OF sets_eq] | |
| 454 | by (simp add: space_pair_measure G.emeasure_pair_measure_Times) | |
| 455 | note 1 2 sets_eq } | |
| 456 | note Times = this(1) and Compl = this(2) and sets_eq = this(3) | |
| 457 | ||
| 458 | fix A assume A: "A \<in> sets (N \<Otimes>\<^sub>M L)" | |
| 459 | show "(\<lambda>a. emeasure (f a \<Otimes>\<^sub>M g a) A) \<in> borel_measurable M" | |
| 460 | using Int_stable_pair_measure_generator pair_measure_closed A | |
| 461 | unfolding sets_pair_measure | |
| 462 | proof (induct A rule: sigma_sets_induct_disjoint) | |
| 463 | case (basic A) then show ?case | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 464 | by (auto intro!: borel_measurable_times_ennreal simp: Times cong: measurable_cong) | 
| 59000 | 465 | (auto intro!: measurable_emeasure_kernel f g) | 
| 466 | next | |
| 467 | case (compl A) | |
| 468 | then have A: "A \<in> sets (N \<Otimes>\<^sub>M L)" | |
| 469 | by (auto simp: sets_pair_measure) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 470 | have "(\<lambda>x. emeasure (f x) (space (f x)) * emeasure (g x) (space (g x)) - | 
| 59000 | 471 | emeasure (f x \<Otimes>\<^sub>M g x) A) \<in> borel_measurable M" (is "?f \<in> ?M") | 
| 472 | using compl(2) f g by measurable | |
| 473 | thus ?case by (simp add: Compl A cong: measurable_cong) | |
| 474 | next | |
| 475 | case (union A) | |
| 476 | then have "range A \<subseteq> sets (N \<Otimes>\<^sub>M L)" "disjoint_family A" | |
| 477 | by (auto simp: sets_pair_measure) | |
| 478 | then have "(\<lambda>a. emeasure (f a \<Otimes>\<^sub>M g a) (\<Union>i. A i)) \<in> borel_measurable M \<longleftrightarrow> | |
| 479 | (\<lambda>a. \<Sum>i. emeasure (f a \<Otimes>\<^sub>M g a) (A i)) \<in> borel_measurable M" | |
| 480 | by (intro measurable_cong suminf_emeasure[symmetric]) | |
| 481 | (auto simp: sets_eq) | |
| 482 | also have "\<dots>" | |
| 483 | using union by auto | |
| 484 | finally show ?case . | |
| 485 | qed simp | |
| 486 | qed | |
| 487 | ||
| 488 | lemma restrict_space_measurable: | |
| 489 |   assumes X: "X \<noteq> {}" "X \<in> sets K"
 | |
| 490 | assumes N: "N \<in> measurable M (subprob_algebra K)" | |
| 491 | shows "(\<lambda>x. restrict_space (N x) X) \<in> measurable M (subprob_algebra (restrict_space K X))" | |
| 492 | proof (rule measurable_subprob_algebra) | |
| 493 | fix a assume a: "a \<in> space M" | |
| 494 | from N[THEN measurable_space, OF this] | |
| 495 | have "subprob_space (N a)" and [simp]: "sets (N a) = sets K" "space (N a) = space K" | |
| 496 | by (auto simp add: space_subprob_algebra dest: sets_eq_imp_space_eq) | |
| 497 | then interpret subprob_space "N a" | |
| 498 | by simp | |
| 499 | show "subprob_space (restrict_space (N a) X)" | |
| 500 | proof | |
| 501 |     show "space (restrict_space (N a) X) \<noteq> {}"
 | |
| 502 | using X by (auto simp add: space_restrict_space) | |
| 503 | show "emeasure (restrict_space (N a) X) (space (restrict_space (N a) X)) \<le> 1" | |
| 504 | using X by (simp add: emeasure_restrict_space space_restrict_space subprob_emeasure_le_1) | |
| 505 | qed | |
| 506 | show "sets (restrict_space (N a) X) = sets (restrict_space K X)" | |
| 507 | by (intro sets_restrict_space_cong) fact | |
| 508 | next | |
| 509 | fix A assume A: "A \<in> sets (restrict_space K X)" | |
| 510 | show "(\<lambda>a. emeasure (restrict_space (N a) X) A) \<in> borel_measurable M" | |
| 511 | proof (subst measurable_cong) | |
| 512 | fix a assume "a \<in> space M" | |
| 513 | from N[THEN measurable_space, OF this] | |
| 514 | have [simp]: "sets (N a) = sets K" "space (N a) = space K" | |
| 515 | by (auto simp add: space_subprob_algebra dest: sets_eq_imp_space_eq) | |
| 516 | show "emeasure (restrict_space (N a) X) A = emeasure (N a) (A \<inter> X)" | |
| 517 | using X A by (subst emeasure_restrict_space) (auto simp add: sets_restrict_space ac_simps) | |
| 518 | next | |
| 519 | show "(\<lambda>w. emeasure (N w) (A \<inter> X)) \<in> borel_measurable M" | |
| 520 | using A X | |
| 521 | by (intro measurable_compose[OF N measurable_emeasure_subprob_algebra]) | |
| 522 | (auto simp: sets_restrict_space) | |
| 523 | qed | |
| 524 | qed | |
| 525 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 526 | subsection \<open>Properties of ``return''\<close> | 
| 58606 | 527 | |
| 528 | definition return :: "'a measure \<Rightarrow> 'a \<Rightarrow> 'a measure" where | |
| 529 | "return R x = measure_of (space R) (sets R) (\<lambda>A. indicator A x)" | |
| 530 | ||
| 531 | lemma space_return[simp]: "space (return M x) = space M" | |
| 532 | by (simp add: return_def) | |
| 533 | ||
| 534 | lemma sets_return[simp]: "sets (return M x) = sets M" | |
| 535 | by (simp add: return_def) | |
| 536 | ||
| 537 | lemma measurable_return1[simp]: "measurable (return N x) L = measurable N L" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 538 | by (simp cong: measurable_cong_sets) | 
| 58606 | 539 | |
| 540 | lemma measurable_return2[simp]: "measurable L (return N x) = measurable L N" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 541 | by (simp cong: measurable_cong_sets) | 
| 58606 | 542 | |
| 59000 | 543 | lemma return_sets_cong: "sets M = sets N \<Longrightarrow> return M = return N" | 
| 544 | by (auto dest: sets_eq_imp_space_eq simp: fun_eq_iff return_def) | |
| 545 | ||
| 546 | lemma return_cong: "sets A = sets B \<Longrightarrow> return A x = return B x" | |
| 547 | by (auto simp add: return_def dest: sets_eq_imp_space_eq) | |
| 548 | ||
| 58606 | 549 | lemma emeasure_return[simp]: | 
| 550 | assumes "A \<in> sets M" | |
| 551 | shows "emeasure (return M x) A = indicator A x" | |
| 552 | proof (rule emeasure_measure_of[OF return_def]) | |
| 553 | show "sets M \<subseteq> Pow (space M)" by (rule sets.space_closed) | |
| 554 | show "positive (sets (return M x)) (\<lambda>A. indicator A x)" by (simp add: positive_def) | |
| 555 | from assms show "A \<in> sets (return M x)" unfolding return_def by simp | |
| 556 | show "countably_additive (sets (return M x)) (\<lambda>A. indicator A x)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 557 | by (auto intro!: countably_additiveI suminf_indicator) | 
| 58606 | 558 | qed | 
| 559 | ||
| 560 | lemma prob_space_return: "x \<in> space M \<Longrightarrow> prob_space (return M x)" | |
| 561 | by rule simp | |
| 562 | ||
| 563 | lemma subprob_space_return: "x \<in> space M \<Longrightarrow> subprob_space (return M x)" | |
| 564 | by (intro prob_space_return prob_space_imp_subprob_space) | |
| 565 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 566 | lemma subprob_space_return_ne: | 
| 59000 | 567 |   assumes "space M \<noteq> {}" shows "subprob_space (return M x)"
 | 
| 568 | proof | |
| 569 | show "emeasure (return M x) (space (return M x)) \<le> 1" | |
| 570 | by (subst emeasure_return) (auto split: split_indicator) | |
| 571 | qed (simp, fact) | |
| 572 | ||
| 573 | lemma measure_return: assumes X: "X \<in> sets M" shows "measure (return M x) X = indicator X x" | |
| 574 | unfolding measure_def emeasure_return[OF X, of x] by (simp split: split_indicator) | |
| 575 | ||
| 58606 | 576 | lemma AE_return: | 
| 577 | assumes [simp]: "x \<in> space M" and [measurable]: "Measurable.pred M P" | |
| 578 | shows "(AE y in return M x. P y) \<longleftrightarrow> P x" | |
| 579 | proof - | |
| 580 |   have "(AE y in return M x. y \<notin> {x\<in>space M. \<not> P x}) \<longleftrightarrow> P x"
 | |
| 581 | by (subst AE_iff_null_sets[symmetric]) (simp_all add: null_sets_def split: split_indicator) | |
| 582 |   also have "(AE y in return M x. y \<notin> {x\<in>space M. \<not> P x}) \<longleftrightarrow> (AE y in return M x. P y)"
 | |
| 583 | by (rule AE_cong) auto | |
| 584 | finally show ?thesis . | |
| 585 | qed | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 586 | |
| 58606 | 587 | lemma nn_integral_return: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 588 | assumes "x \<in> space M" "g \<in> borel_measurable M" | 
| 58606 | 589 | shows "(\<integral>\<^sup>+ a. g a \<partial>return M x) = g x" | 
| 590 | proof- | |
| 61808 | 591 | interpret prob_space "return M x" by (rule prob_space_return[OF \<open>x \<in> space M\<close>]) | 
| 58606 | 592 | have "(\<integral>\<^sup>+ a. g a \<partial>return M x) = (\<integral>\<^sup>+ a. g x \<partial>return M x)" using assms | 
| 593 | by (intro nn_integral_cong_AE) (auto simp: AE_return) | |
| 594 | also have "... = g x" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 595 | using nn_integral_const[of "return M x"] emeasure_space_1 by simp | 
| 58606 | 596 | finally show ?thesis . | 
| 597 | qed | |
| 598 | ||
| 59000 | 599 | lemma integral_return: | 
| 600 |   fixes g :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | |
| 601 | assumes "x \<in> space M" "g \<in> borel_measurable M" | |
| 602 | shows "(\<integral>a. g a \<partial>return M x) = g x" | |
| 603 | proof- | |
| 61808 | 604 | interpret prob_space "return M x" by (rule prob_space_return[OF \<open>x \<in> space M\<close>]) | 
| 59000 | 605 | have "(\<integral>a. g a \<partial>return M x) = (\<integral>a. g x \<partial>return M x)" using assms | 
| 606 | by (intro integral_cong_AE) (auto simp: AE_return) | |
| 607 | then show ?thesis | |
| 608 | using prob_space by simp | |
| 609 | qed | |
| 610 | ||
| 611 | lemma return_measurable[measurable]: "return N \<in> measurable N (subprob_algebra N)" | |
| 58606 | 612 | by (rule measurable_subprob_algebra) (auto simp: subprob_space_return) | 
| 613 | ||
| 614 | lemma distr_return: | |
| 615 | assumes "f \<in> measurable M N" and "x \<in> space M" | |
| 616 | shows "distr (return M x) N f = return N (f x)" | |
| 617 | using assms by (intro measure_eqI) (simp_all add: indicator_def emeasure_distr) | |
| 618 | ||
| 59000 | 619 | lemma return_restrict_space: | 
| 620 | "\<Omega> \<in> sets M \<Longrightarrow> return (restrict_space M \<Omega>) x = restrict_space (return M x) \<Omega>" | |
| 621 | by (auto intro!: measure_eqI simp: sets_restrict_space emeasure_restrict_space) | |
| 622 | ||
| 623 | lemma measurable_distr2: | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61359diff
changeset | 624 | assumes f[measurable]: "case_prod f \<in> measurable (L \<Otimes>\<^sub>M M) N" | 
| 59000 | 625 | assumes g[measurable]: "g \<in> measurable L (subprob_algebra M)" | 
| 626 | shows "(\<lambda>x. distr (g x) N (f x)) \<in> measurable L (subprob_algebra N)" | |
| 627 | proof - | |
| 628 | have "(\<lambda>x. distr (g x) N (f x)) \<in> measurable L (subprob_algebra N) | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61359diff
changeset | 629 | \<longleftrightarrow> (\<lambda>x. distr (return L x \<Otimes>\<^sub>M g x) N (case_prod f)) \<in> measurable L (subprob_algebra N)" | 
| 59000 | 630 | proof (rule measurable_cong) | 
| 631 | fix x assume x: "x \<in> space L" | |
| 632 | have gx: "g x \<in> space (subprob_algebra M)" | |
| 633 | using measurable_space[OF g x] . | |
| 634 | then have [simp]: "sets (g x) = sets M" | |
| 635 | by (simp add: space_subprob_algebra) | |
| 636 | then have [simp]: "space (g x) = space M" | |
| 637 | by (rule sets_eq_imp_space_eq) | |
| 638 | let ?R = "return L x" | |
| 639 | from measurable_compose_Pair1[OF x f] have f_M': "f x \<in> measurable M N" | |
| 640 | by simp | |
| 641 | interpret subprob_space "g x" | |
| 642 | using gx by (simp add: space_subprob_algebra) | |
| 643 | have space_pair_M'[simp]: "\<And>X. space (X \<Otimes>\<^sub>M g x) = space (X \<Otimes>\<^sub>M M)" | |
| 644 | by (simp add: space_pair_measure) | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61359diff
changeset | 645 | show "distr (g x) N (f x) = distr (?R \<Otimes>\<^sub>M g x) N (case_prod f)" (is "?l = ?r") | 
| 59000 | 646 | proof (rule measure_eqI) | 
| 647 | show "sets ?l = sets ?r" | |
| 648 | by simp | |
| 649 | next | |
| 650 | fix A assume "A \<in> sets ?l" | |
| 651 | then have A[measurable]: "A \<in> sets N" | |
| 652 | by simp | |
| 653 | then have "emeasure ?r A = emeasure (?R \<Otimes>\<^sub>M g x) ((\<lambda>(x, y). f x y) -` A \<inter> space (?R \<Otimes>\<^sub>M g x))" | |
| 654 | by (auto simp add: emeasure_distr f_M' cong: measurable_cong_sets) | |
| 655 | also have "\<dots> = (\<integral>\<^sup>+M''. emeasure (g x) (f M'' -` A \<inter> space M) \<partial>?R)" | |
| 656 | apply (subst emeasure_pair_measure_alt) | |
| 657 | apply (rule measurable_sets[OF _ A]) | |
| 658 | apply (auto simp add: f_M' cong: measurable_cong_sets) | |
| 659 | apply (intro nn_integral_cong arg_cong[where f="emeasure (g x)"]) | |
| 660 | apply (auto simp: space_subprob_algebra space_pair_measure) | |
| 661 | done | |
| 662 | also have "\<dots> = emeasure (g x) (f x -` A \<inter> space M)" | |
| 663 | by (subst nn_integral_return) | |
| 664 | (auto simp: x intro!: measurable_emeasure) | |
| 665 | also have "\<dots> = emeasure ?l A" | |
| 666 | by (simp add: emeasure_distr f_M' cong: measurable_cong_sets) | |
| 667 | finally show "emeasure ?l A = emeasure ?r A" .. | |
| 668 | qed | |
| 669 | qed | |
| 670 | also have "\<dots>" | |
| 671 | apply (intro measurable_compose[OF measurable_pair_measure measurable_distr]) | |
| 672 | apply (rule return_measurable) | |
| 673 | apply measurable | |
| 674 | done | |
| 675 | finally show ?thesis . | |
| 676 | qed | |
| 677 | ||
| 678 | lemma nn_integral_measurable_subprob_algebra2: | |
| 59048 | 679 | assumes f[measurable]: "(\<lambda>(x, y). f x y) \<in> borel_measurable (M \<Otimes>\<^sub>M N)" | 
| 59000 | 680 | assumes N[measurable]: "L \<in> measurable M (subprob_algebra N)" | 
| 681 | shows "(\<lambda>x. integral\<^sup>N (L x) (f x)) \<in> borel_measurable M" | |
| 682 | proof - | |
| 59048 | 683 | note nn_integral_measurable_subprob_algebra[measurable] | 
| 684 | note measurable_distr2[measurable] | |
| 59000 | 685 | have "(\<lambda>x. integral\<^sup>N (distr (L x) (M \<Otimes>\<^sub>M N) (\<lambda>y. (x, y))) (\<lambda>(x, y). f x y)) \<in> borel_measurable M" | 
| 59048 | 686 | by measurable | 
| 59000 | 687 | then show "(\<lambda>x. integral\<^sup>N (L x) (f x)) \<in> borel_measurable M" | 
| 59048 | 688 | by (rule measurable_cong[THEN iffD1, rotated]) | 
| 689 | (simp add: nn_integral_distr) | |
| 59000 | 690 | qed | 
| 691 | ||
| 692 | lemma emeasure_measurable_subprob_algebra2: | |
| 693 | assumes A[measurable]: "(SIGMA x:space M. A x) \<in> sets (M \<Otimes>\<^sub>M N)" | |
| 694 | assumes L[measurable]: "L \<in> measurable M (subprob_algebra N)" | |
| 695 | shows "(\<lambda>x. emeasure (L x) (A x)) \<in> borel_measurable M" | |
| 696 | proof - | |
| 697 |   { fix x assume "x \<in> space M"
 | |
| 698 | then have "Pair x -` Sigma (space M) A = A x" | |
| 699 | by auto | |
| 700 | with sets_Pair1[OF A, of x] have "A x \<in> sets N" | |
| 701 | by auto } | |
| 702 | note ** = this | |
| 703 | ||
| 704 | have *: "\<And>x. fst x \<in> space M \<Longrightarrow> snd x \<in> A (fst x) \<longleftrightarrow> x \<in> (SIGMA x:space M. A x)" | |
| 705 | by (auto simp: fun_eq_iff) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 706 | have "(\<lambda>(x, y). indicator (A x) y::ennreal) \<in> borel_measurable (M \<Otimes>\<^sub>M N)" | 
| 59000 | 707 | apply measurable | 
| 708 | apply (subst measurable_cong) | |
| 709 | apply (rule *) | |
| 710 | apply (auto simp: space_pair_measure) | |
| 711 | done | |
| 712 | then have "(\<lambda>x. integral\<^sup>N (L x) (indicator (A x))) \<in> borel_measurable M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 713 | by (intro nn_integral_measurable_subprob_algebra2[where N=N] L) | 
| 59000 | 714 | then show "(\<lambda>x. emeasure (L x) (A x)) \<in> borel_measurable M" | 
| 715 | apply (rule measurable_cong[THEN iffD1, rotated]) | |
| 716 | apply (rule nn_integral_indicator) | |
| 717 | apply (simp add: subprob_measurableD[OF L] **) | |
| 718 | done | |
| 719 | qed | |
| 720 | ||
| 721 | lemma measure_measurable_subprob_algebra2: | |
| 722 | assumes A[measurable]: "(SIGMA x:space M. A x) \<in> sets (M \<Otimes>\<^sub>M N)" | |
| 723 | assumes L[measurable]: "L \<in> measurable M (subprob_algebra N)" | |
| 724 | shows "(\<lambda>x. measure (L x) (A x)) \<in> borel_measurable M" | |
| 725 | unfolding measure_def | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 726 | by (intro borel_measurable_enn2real emeasure_measurable_subprob_algebra2[OF assms]) | 
| 59000 | 727 | |
| 58606 | 728 | definition "select_sets M = (SOME N. sets M = sets (subprob_algebra N))" | 
| 729 | ||
| 730 | lemma select_sets1: | |
| 731 | "sets M = sets (subprob_algebra N) \<Longrightarrow> sets M = sets (subprob_algebra (select_sets M))" | |
| 732 | unfolding select_sets_def by (rule someI) | |
| 733 | ||
| 734 | lemma sets_select_sets[simp]: | |
| 735 | assumes sets: "sets M = sets (subprob_algebra N)" | |
| 736 | shows "sets (select_sets M) = sets N" | |
| 737 | unfolding select_sets_def | |
| 738 | proof (rule someI2) | |
| 739 | show "sets M = sets (subprob_algebra N)" | |
| 740 | by fact | |
| 741 | next | |
| 742 | fix L assume "sets M = sets (subprob_algebra L)" | |
| 743 | with sets have eq: "space (subprob_algebra N) = space (subprob_algebra L)" | |
| 744 | by (intro sets_eq_imp_space_eq) simp | |
| 745 | show "sets L = sets N" | |
| 746 | proof cases | |
| 747 |     assume "space (subprob_algebra N) = {}"
 | |
| 748 | with space_subprob_algebra_empty_iff[of N] space_subprob_algebra_empty_iff[of L] | |
| 749 | show ?thesis | |
| 750 | by (simp add: eq space_empty_iff) | |
| 751 | next | |
| 752 |     assume "space (subprob_algebra N) \<noteq> {}"
 | |
| 753 | with eq show ?thesis | |
| 754 | by (fastforce simp add: space_subprob_algebra) | |
| 755 | qed | |
| 756 | qed | |
| 757 | ||
| 758 | lemma space_select_sets[simp]: | |
| 759 | "sets M = sets (subprob_algebra N) \<Longrightarrow> space (select_sets M) = space N" | |
| 760 | by (intro sets_eq_imp_space_eq sets_select_sets) | |
| 761 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 762 | subsection \<open>Join\<close> | 
| 58606 | 763 | |
| 764 | definition join :: "'a measure measure \<Rightarrow> 'a measure" where | |
| 765 | "join M = measure_of (space (select_sets M)) (sets (select_sets M)) (\<lambda>B. \<integral>\<^sup>+ M'. emeasure M' B \<partial>M)" | |
| 766 | ||
| 767 | lemma | |
| 768 | shows space_join[simp]: "space (join M) = space (select_sets M)" | |
| 769 | and sets_join[simp]: "sets (join M) = sets (select_sets M)" | |
| 770 | by (simp_all add: join_def) | |
| 771 | ||
| 772 | lemma emeasure_join: | |
| 59048 | 773 | assumes M[simp, measurable_cong]: "sets M = sets (subprob_algebra N)" and A: "A \<in> sets N" | 
| 58606 | 774 | shows "emeasure (join M) A = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)" | 
| 775 | proof (rule emeasure_measure_of[OF join_def]) | |
| 776 | show "countably_additive (sets (join M)) (\<lambda>B. \<integral>\<^sup>+ M'. emeasure M' B \<partial>M)" | |
| 777 | proof (rule countably_additiveI) | |
| 778 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (join M)" "disjoint_family A" | |
| 779 | have "(\<Sum>i. \<integral>\<^sup>+ M'. emeasure M' (A i) \<partial>M) = (\<integral>\<^sup>+M'. (\<Sum>i. emeasure M' (A i)) \<partial>M)" | |
| 59048 | 780 | using A by (subst nn_integral_suminf) (auto simp: measurable_emeasure_subprob_algebra) | 
| 58606 | 781 | also have "\<dots> = (\<integral>\<^sup>+M'. emeasure M' (\<Union>i. A i) \<partial>M)" | 
| 782 | proof (rule nn_integral_cong) | |
| 783 | fix M' assume "M' \<in> space M" | |
| 784 | then show "(\<Sum>i. emeasure M' (A i)) = emeasure M' (\<Union>i. A i)" | |
| 785 | using A sets_eq_imp_space_eq[OF M] by (simp add: suminf_emeasure space_subprob_algebra) | |
| 786 | qed | |
| 787 | finally show "(\<Sum>i. \<integral>\<^sup>+M'. emeasure M' (A i) \<partial>M) = (\<integral>\<^sup>+M'. emeasure M' (\<Union>i. A i) \<partial>M)" . | |
| 788 | qed | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 789 | qed (auto simp: A sets.space_closed positive_def) | 
| 58606 | 790 | |
| 791 | lemma measurable_join: | |
| 792 | "join \<in> measurable (subprob_algebra (subprob_algebra N)) (subprob_algebra N)" | |
| 793 | proof (cases "space N \<noteq> {}", rule measurable_subprob_algebra)
 | |
| 794 | fix A assume "A \<in> sets N" | |
| 795 | let ?B = "borel_measurable (subprob_algebra (subprob_algebra N))" | |
| 796 | have "(\<lambda>M'. emeasure (join M') A) \<in> ?B \<longleftrightarrow> (\<lambda>M'. (\<integral>\<^sup>+ M''. emeasure M'' A \<partial>M')) \<in> ?B" | |
| 797 | proof (rule measurable_cong) | |
| 798 | fix M' assume "M' \<in> space (subprob_algebra (subprob_algebra N))" | |
| 799 | then show "emeasure (join M') A = (\<integral>\<^sup>+ M''. emeasure M'' A \<partial>M')" | |
| 61808 | 800 | by (intro emeasure_join) (auto simp: space_subprob_algebra \<open>A\<in>sets N\<close>) | 
| 58606 | 801 | qed | 
| 802 | also have "(\<lambda>M'. \<integral>\<^sup>+M''. emeasure M'' A \<partial>M') \<in> ?B" | |
| 61808 | 803 | using measurable_emeasure_subprob_algebra[OF \<open>A\<in>sets N\<close>] | 
| 58606 | 804 | by (rule nn_integral_measurable_subprob_algebra) | 
| 805 | finally show "(\<lambda>M'. emeasure (join M') A) \<in> borel_measurable (subprob_algebra (subprob_algebra N))" . | |
| 806 | next | |
| 807 |   assume [simp]: "space N \<noteq> {}"
 | |
| 808 | fix M assume M: "M \<in> space (subprob_algebra (subprob_algebra N))" | |
| 809 | then have "(\<integral>\<^sup>+M'. emeasure M' (space N) \<partial>M) \<le> (\<integral>\<^sup>+M'. 1 \<partial>M)" | |
| 810 | apply (intro nn_integral_mono) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 811 | apply (auto simp: space_subprob_algebra | 
| 58606 | 812 | dest!: sets_eq_imp_space_eq subprob_space.emeasure_space_le_1) | 
| 813 | done | |
| 814 | with M show "subprob_space (join M)" | |
| 815 | by (intro subprob_spaceI) | |
| 63092 | 816 | (auto simp: emeasure_join space_subprob_algebra M dest: subprob_space.emeasure_space_le_1) | 
| 58606 | 817 | next | 
| 818 |   assume "\<not>(space N \<noteq> {})"
 | |
| 819 | thus ?thesis by (simp add: measurable_empty_iff space_subprob_algebra_empty_iff) | |
| 820 | qed (auto simp: space_subprob_algebra) | |
| 821 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 822 | lemma nn_integral_join: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 823 | assumes f: "f \<in> borel_measurable N" | 
| 59048 | 824 | and M[measurable_cong]: "sets M = sets (subprob_algebra N)" | 
| 58606 | 825 | shows "(\<integral>\<^sup>+x. f x \<partial>join M) = (\<integral>\<^sup>+M'. \<integral>\<^sup>+x. f x \<partial>M' \<partial>M)" | 
| 826 | using f | |
| 827 | proof induct | |
| 828 | case (cong f g) | |
| 829 | moreover have "integral\<^sup>N (join M) f = integral\<^sup>N (join M) g" | |
| 830 | by (intro nn_integral_cong cong) (simp add: M) | |
| 831 | moreover from M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' f \<partial>M) = (\<integral>\<^sup>+ M'. integral\<^sup>N M' g \<partial>M)" | |
| 832 | by (intro nn_integral_cong cong) | |
| 833 | (auto simp add: space_subprob_algebra dest!: sets_eq_imp_space_eq) | |
| 834 | ultimately show ?case | |
| 835 | by simp | |
| 836 | next | |
| 837 | case (set A) | |
| 63540 | 838 | with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)" | 
| 58606 | 839 | by (intro nn_integral_cong nn_integral_indicator) | 
| 840 | (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq) | |
| 63540 | 841 | with set show ?case | 
| 58606 | 842 | using M by (simp add: emeasure_join) | 
| 843 | next | |
| 844 | case (mult f c) | |
| 845 | have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. c * f x \<partial>M' \<partial>M) = (\<integral>\<^sup>+ M'. c * \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)" | |
| 59048 | 846 | using mult M M[THEN sets_eq_imp_space_eq] | 
| 847 | by (intro nn_integral_cong nn_integral_cmult) (auto simp add: space_subprob_algebra) | |
| 58606 | 848 | also have "\<dots> = c * (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 849 | using nn_integral_measurable_subprob_algebra[OF mult(2)] | 
| 58606 | 850 | by (intro nn_integral_cmult mult) (simp add: M) | 
| 851 | also have "\<dots> = c * (integral\<^sup>N (join M) f)" | |
| 852 | by (simp add: mult) | |
| 853 | also have "\<dots> = (\<integral>\<^sup>+ x. c * f x \<partial>join M)" | |
| 59048 | 854 | using mult(2,3) by (intro nn_integral_cmult[symmetric] mult) (simp add: M cong: measurable_cong_sets) | 
| 58606 | 855 | finally show ?case by simp | 
| 856 | next | |
| 857 | case (add f g) | |
| 858 | have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x + g x \<partial>M' \<partial>M) = (\<integral>\<^sup>+ M'. (\<integral>\<^sup>+ x. f x \<partial>M') + (\<integral>\<^sup>+ x. g x \<partial>M') \<partial>M)" | |
| 59048 | 859 | using add M M[THEN sets_eq_imp_space_eq] | 
| 860 | by (intro nn_integral_cong nn_integral_add) (auto simp add: space_subprob_algebra) | |
| 58606 | 861 | also have "\<dots> = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M) + (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. g x \<partial>M' \<partial>M)" | 
| 59048 | 862 | using nn_integral_measurable_subprob_algebra[OF add(1)] | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 863 | using nn_integral_measurable_subprob_algebra[OF add(4)] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 864 | by (intro nn_integral_add add) (simp_all add: M) | 
| 58606 | 865 | also have "\<dots> = (integral\<^sup>N (join M) f) + (integral\<^sup>N (join M) g)" | 
| 866 | by (simp add: add) | |
| 867 | also have "\<dots> = (\<integral>\<^sup>+ x. f x + g x \<partial>join M)" | |
| 59048 | 868 | using add by (intro nn_integral_add[symmetric] add) (simp_all add: M cong: measurable_cong_sets) | 
| 58606 | 869 | finally show ?case by (simp add: ac_simps) | 
| 870 | next | |
| 871 | case (seq F) | |
| 872 | have "(\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. (SUP i. F i) x \<partial>M' \<partial>M) = (\<integral>\<^sup>+ M'. (SUP i. \<integral>\<^sup>+ x. F i x \<partial>M') \<partial>M)" | |
| 59048 | 873 | using seq M M[THEN sets_eq_imp_space_eq] unfolding SUP_apply | 
| 58606 | 874 | by (intro nn_integral_cong nn_integral_monotone_convergence_SUP) | 
| 59048 | 875 | (auto simp add: space_subprob_algebra) | 
| 58606 | 876 | also have "\<dots> = (SUP i. \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. F i x \<partial>M' \<partial>M)" | 
| 59048 | 877 | using nn_integral_measurable_subprob_algebra[OF seq(1)] seq | 
| 58606 | 878 | by (intro nn_integral_monotone_convergence_SUP) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 879 | (simp_all add: M incseq_nn_integral incseq_def le_fun_def nn_integral_mono ) | 
| 58606 | 880 | also have "\<dots> = (SUP i. integral\<^sup>N (join M) (F i))" | 
| 881 | by (simp add: seq) | |
| 882 | also have "\<dots> = (\<integral>\<^sup>+ x. (SUP i. F i x) \<partial>join M)" | |
| 59048 | 883 | using seq by (intro nn_integral_monotone_convergence_SUP[symmetric] seq) | 
| 884 | (simp_all add: M cong: measurable_cong_sets) | |
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69546diff
changeset | 885 | finally show ?case by (simp add: ac_simps image_comp) | 
| 58606 | 886 | qed | 
| 887 | ||
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 888 | lemma measurable_join1: | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 889 | "\<lbrakk> f \<in> measurable N K; sets M = sets (subprob_algebra N) \<rbrakk> | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 890 | \<Longrightarrow> f \<in> measurable (join M) K" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 891 | by(simp add: measurable_def) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 892 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 893 | lemma | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 894 | fixes f :: "_ \<Rightarrow> real" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 895 | assumes f_measurable [measurable]: "f \<in> borel_measurable N" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 896 | and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> \<bar>f x\<bar> \<le> B" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 897 | and M [measurable_cong]: "sets M = sets (subprob_algebra N)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 898 | and fin: "finite_measure M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 899 | and M_bounded: "AE M' in M. emeasure M' (space M') \<le> ennreal B'" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 900 | shows integrable_join: "integrable (join M) f" (is ?integrable) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 901 | and integral_join: "integral\<^sup>L (join M) f = \<integral> M'. integral\<^sup>L M' f \<partial>M" (is ?integral) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 902 | proof(case_tac [!] "space N = {}")
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 903 |   assume *: "space N = {}"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 904 | show ?integrable | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 905 | using M * by(simp add: real_integrable_def measurable_def nn_integral_empty) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 906 | have "(\<integral> M'. integral\<^sup>L M' f \<partial>M) = (\<integral> M'. 0 \<partial>M)" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 907 | proof(rule Bochner_Integration.integral_cong) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 908 | fix M' | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 909 | assume "M' \<in> space M" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 910 | with sets_eq_imp_space_eq[OF M] have "space M' = space N" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 911 | by(auto simp add: space_subprob_algebra dest: sets_eq_imp_space_eq) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 912 | with * show "(\<integral> x. f x \<partial>M') = 0" by(simp add: Bochner_Integration.integral_empty) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 913 | qed simp | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 914 | then show ?integral | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 915 | using M * by(simp add: Bochner_Integration.integral_empty) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 916 | next | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 917 |   assume *: "space N \<noteq> {}"
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 918 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 919 | from * have B [simp]: "0 \<le> B" by(auto dest: f_bounded) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 920 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 921 | have [measurable]: "f \<in> borel_measurable (join M)" using f_measurable M | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 922 | by(rule measurable_join1) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 923 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 924 |   { fix f M'
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 925 | assume [measurable]: "f \<in> borel_measurable N" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 926 | and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> f x \<le> B" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 927 | and "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 928 | have "AE x in M'. ennreal (f x) \<le> ennreal B" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 929 | proof(rule AE_I2) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 930 | fix x | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 931 | assume "x \<in> space M'" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 932 | with \<open>M' \<in> space M\<close> sets_eq_imp_space_eq[OF M] | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 933 | have "x \<in> space N" by(auto simp add: space_subprob_algebra dest: sets_eq_imp_space_eq) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 934 | from f_bounded[OF this] show "ennreal (f x) \<le> ennreal B" by simp | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 935 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 936 | then have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> (\<integral>\<^sup>+ x. ennreal B \<partial>M')" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 937 | by(rule nn_integral_mono_AE) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 938 | also have "\<dots> = ennreal B * emeasure M' (space M')" by(simp) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 939 | also have "\<dots> \<le> ennreal B * ennreal B'" by(rule mult_left_mono)(fact, simp) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 940 | also have "\<dots> \<le> ennreal B * ennreal \<bar>B'\<bar>" by(rule mult_left_mono)(simp_all) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 941 | finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)" by (simp add: ennreal_mult) } | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 942 | note bounded1 = this | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 943 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 944 | have bounded: | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 945 | "\<And>f. \<lbrakk> f \<in> borel_measurable N; \<And>x. x \<in> space N \<Longrightarrow> f x \<le> B \<rbrakk> | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 946 | \<Longrightarrow> (\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) \<noteq> top" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 947 | proof - | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 948 | fix f | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 949 | assume [measurable]: "f \<in> borel_measurable N" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 950 | and f_bounded: "\<And>x. x \<in> space N \<Longrightarrow> f x \<le> B" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 951 | have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. ennreal (f x) \<partial>M' \<partial>M)" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 952 | by(rule nn_integral_join[OF _ M]) simp | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 953 | also have "\<dots> \<le> \<integral>\<^sup>+ M'. B * \<bar>B'\<bar> \<partial>M" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 954 | using bounded1[OF \<open>f \<in> borel_measurable N\<close> f_bounded] | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 955 | by(rule nn_integral_mono_AE[OF AE_mp[OF M_bounded AE_I2], rule_format]) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 956 | also have "\<dots> = B * \<bar>B'\<bar> * emeasure M (space M)" by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 957 | also have "\<dots> < \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 958 | using finite_measure.finite_emeasure_space[OF fin] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 959 | by(simp add: ennreal_mult_less_top less_top) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 960 | finally show "?thesis f" by simp | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 961 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 962 | have f_pos: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>join M) \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 963 | and f_neg: "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>join M) \<noteq> \<infinity>" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 964 | using f_bounded by(auto del: notI intro!: bounded simp add: abs_le_iff) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 965 | |
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 966 | show ?integrable using f_pos f_neg by(simp add: real_integrable_def) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 967 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 968 | note [measurable] = nn_integral_measurable_subprob_algebra | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 969 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 970 | have int_f: "(\<integral>\<^sup>+ x. f x \<partial>join M) = \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 971 | by(simp add: nn_integral_join[OF _ M]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 972 | have int_mf: "(\<integral>\<^sup>+ x. - f x \<partial>join M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 973 | by(simp add: nn_integral_join[OF _ M]) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 974 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 975 | have pos_finite: "AE M' in M. (\<integral>\<^sup>+ x. f x \<partial>M') \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 976 | using AE_space M_bounded | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 977 | proof eventually_elim | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 978 | fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 979 | then have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 980 | using f_measurable by(auto intro!: bounded1 dest: f_bounded) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 981 | then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M') \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 982 | by (auto simp: top_unique) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 983 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 984 | hence [simp]: "(\<integral>\<^sup>+ M'. ennreal (enn2real (\<integral>\<^sup>+ x. f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 985 | by (rule nn_integral_cong_AE[OF AE_mp]) (simp add: less_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 986 | from f_pos have [simp]: "integrable M (\<lambda>M'. enn2real (\<integral>\<^sup>+ x. f x \<partial>M'))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 987 | by(simp add: int_f real_integrable_def nn_integral_0_iff_AE[THEN iffD2] ennreal_neg enn2real_nonneg) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 988 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 989 | have neg_finite: "AE M' in M. (\<integral>\<^sup>+ x. - f x \<partial>M') \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 990 | using AE_space M_bounded | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 991 | proof eventually_elim | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 992 | fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> ennreal B'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 993 | then have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M') \<le> ennreal (B * \<bar>B'\<bar>)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 994 | using f_measurable by(auto intro!: bounded1 dest: f_bounded) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 995 | then show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M') \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 996 | by (auto simp: top_unique) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 997 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 998 | hence [simp]: "(\<integral>\<^sup>+ M'. ennreal (enn2real (\<integral>\<^sup>+ x. - f x \<partial>M')) \<partial>M) = (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. - f x \<partial>M' \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 999 | by (rule nn_integral_cong_AE[OF AE_mp]) (simp add: less_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1000 | from f_neg have [simp]: "integrable M (\<lambda>M'. enn2real (\<integral>\<^sup>+ x. - f x \<partial>M'))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1001 | by(simp add: int_mf real_integrable_def nn_integral_0_iff_AE[THEN iffD2] ennreal_neg enn2real_nonneg) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1002 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1003 | have "(\<integral> x. f x \<partial>join M) = enn2real (\<integral>\<^sup>+ N. \<integral>\<^sup>+x. f x \<partial>N \<partial>M) - enn2real (\<integral>\<^sup>+ N. \<integral>\<^sup>+x. - f x \<partial>N \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1004 | unfolding real_lebesgue_integral_def[OF \<open>?integrable\<close>] by (simp add: nn_integral_join[OF _ M]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1005 | also have "\<dots> = (\<integral>N. enn2real (\<integral>\<^sup>+x. f x \<partial>N) \<partial>M) - (\<integral>N. enn2real (\<integral>\<^sup>+x. - f x \<partial>N) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1006 | using pos_finite neg_finite by (subst (1 2) integral_eq_nn_integral) (auto simp: enn2real_nonneg) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1007 | also have "\<dots> = (\<integral>N. enn2real (\<integral>\<^sup>+x. f x \<partial>N) - enn2real (\<integral>\<^sup>+x. - f x \<partial>N) \<partial>M)" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1008 | by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1009 | also have "\<dots> = \<integral>M'. \<integral> x. f x \<partial>M' \<partial>M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1010 | proof (rule integral_cong_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1011 | show "AE x in M. | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1012 | enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>x) - enn2real (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>x) = integral\<^sup>L x f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1013 | using AE_space M_bounded | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1014 | proof eventually_elim | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1015 | fix M' assume "M' \<in> space M" "emeasure M' (space M') \<le> B'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1016 | then interpret subprob_space M' | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1017 | by (auto simp: M[THEN sets_eq_imp_space_eq] space_subprob_algebra) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1018 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1019 | from \<open>M' \<in> space M\<close> sets_eq_imp_space_eq[OF M] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1020 | have [measurable_cong]: "sets M' = sets N" by(simp add: space_subprob_algebra) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1021 | hence [simp]: "space M' = space N" by(rule sets_eq_imp_space_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1022 | have "integrable M' f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1023 | by(rule integrable_const_bound[where B=B])(auto simp add: f_bounded) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1024 | then show "enn2real (\<integral>\<^sup>+ x. f x \<partial>M') - enn2real (\<integral>\<^sup>+ x. - f x \<partial>M') = \<integral> x. f x \<partial>M'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1025 | by(simp add: real_lebesgue_integral_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1026 | qed | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1027 | qed simp_all | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1028 | finally show ?integral by simp | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1029 | qed | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1030 | |
| 58606 | 1031 | lemma join_assoc: | 
| 59048 | 1032 | assumes M[measurable_cong]: "sets M = sets (subprob_algebra (subprob_algebra N))" | 
| 58606 | 1033 | shows "join (distr M (subprob_algebra N) join) = join (join M)" | 
| 1034 | proof (rule measure_eqI) | |
| 1035 | fix A assume "A \<in> sets (join (distr M (subprob_algebra N) join))" | |
| 1036 | then have A: "A \<in> sets N" by simp | |
| 1037 | show "emeasure (join (distr M (subprob_algebra N) join)) A = emeasure (join (join M)) A" | |
| 1038 | using measurable_join[of N] | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1039 | by (auto simp: M A nn_integral_distr emeasure_join measurable_emeasure_subprob_algebra | 
| 59048 | 1040 | sets_eq_imp_space_eq[OF M] space_subprob_algebra nn_integral_join[OF _ M] | 
| 1041 | intro!: nn_integral_cong emeasure_join) | |
| 58606 | 1042 | qed (simp add: M) | 
| 1043 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1044 | lemma join_return: | 
| 58606 | 1045 | assumes "sets M = sets N" and "subprob_space M" | 
| 1046 | shows "join (return (subprob_algebra N) M) = M" | |
| 1047 | by (rule measure_eqI) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1048 | (simp_all add: emeasure_join space_subprob_algebra | 
| 58606 | 1049 | measurable_emeasure_subprob_algebra nn_integral_return assms) | 
| 1050 | ||
| 1051 | lemma join_return': | |
| 1052 | assumes "sets N = sets M" | |
| 1053 | shows "join (distr M (subprob_algebra N) (return N)) = M" | |
| 1054 | apply (rule measure_eqI) | |
| 1055 | apply (simp add: assms) | |
| 1056 | apply (subgoal_tac "return N \<in> measurable M (subprob_algebra N)") | |
| 1057 | apply (simp add: emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra assms) | |
| 1058 | apply (subst measurable_cong_sets, rule assms[symmetric], rule refl, rule return_measurable) | |
| 1059 | done | |
| 1060 | ||
| 1061 | lemma join_distr_distr: | |
| 1062 | fixes f :: "'a \<Rightarrow> 'b" and M :: "'a measure measure" and N :: "'b measure" | |
| 1063 | assumes "sets M = sets (subprob_algebra R)" and "f \<in> measurable R N" | |
| 1064 | shows "join (distr M (subprob_algebra N) (\<lambda>M. distr M N f)) = distr (join M) N f" (is "?r = ?l") | |
| 1065 | proof (rule measure_eqI) | |
| 1066 | fix A assume "A \<in> sets ?r" | |
| 1067 | hence A_in_N: "A \<in> sets N" by simp | |
| 1068 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1069 | from assms have "f \<in> measurable (join M) N" | 
| 58606 | 1070 | by (simp cong: measurable_cong_sets) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1071 | moreover from assms and A_in_N have "f-`A \<inter> space R \<in> sets R" | 
| 58606 | 1072 | by (intro measurable_sets) simp_all | 
| 1073 | ultimately have "emeasure (distr (join M) N f) A = \<integral>\<^sup>+M'. emeasure M' (f-`A \<inter> space R) \<partial>M" | |
| 1074 | by (simp_all add: A_in_N emeasure_distr emeasure_join assms) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1075 | |
| 58606 | 1076 | also have "... = \<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M" using A_in_N | 
| 1077 | proof (intro nn_integral_cong, subst emeasure_distr) | |
| 1078 | fix M' assume "M' \<in> space M" | |
| 1079 | from assms have "space M = space (subprob_algebra R)" | |
| 1080 | using sets_eq_imp_space_eq by blast | |
| 61808 | 1081 | with \<open>M' \<in> space M\<close> have [simp]: "sets M' = sets R" using space_subprob_algebra by blast | 
| 58606 | 1082 | show "f \<in> measurable M' N" by (simp cong: measurable_cong_sets add: assms) | 
| 1083 | have "space M' = space R" by (rule sets_eq_imp_space_eq) simp | |
| 1084 | thus "emeasure M' (f -` A \<inter> space R) = emeasure M' (f -` A \<inter> space M')" by simp | |
| 1085 | qed | |
| 1086 | ||
| 1087 | also have "(\<lambda>M. distr M N f) \<in> measurable M (subprob_algebra N)" | |
| 1088 | by (simp cong: measurable_cong_sets add: assms measurable_distr) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1089 | hence "(\<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M) = | 
| 58606 | 1090 | emeasure (join (distr M (subprob_algebra N) (\<lambda>M. distr M N f))) A" | 
| 1091 | by (simp_all add: emeasure_join assms A_in_N nn_integral_distr measurable_emeasure_subprob_algebra) | |
| 1092 | finally show "emeasure ?r A = emeasure ?l A" .. | |
| 1093 | qed simp | |
| 1094 | ||
| 1095 | definition bind :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b measure) \<Rightarrow> 'b measure" where
 | |
| 1096 |   "bind M f = (if space M = {} then count_space {} else
 | |
| 1097 | join (distr M (subprob_algebra (f (SOME x. x \<in> space M))) f))" | |
| 1098 | ||
| 1099 | adhoc_overloading Monad_Syntax.bind bind | |
| 1100 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1101 | lemma bind_empty: | 
| 58606 | 1102 |   "space M = {} \<Longrightarrow> bind M f = count_space {}"
 | 
| 1103 | by (simp add: bind_def) | |
| 1104 | ||
| 1105 | lemma bind_nonempty: | |
| 1106 |   "space M \<noteq> {} \<Longrightarrow> bind M f = join (distr M (subprob_algebra (f (SOME x. x \<in> space M))) f)"
 | |
| 1107 | by (simp add: bind_def) | |
| 1108 | ||
| 1109 | lemma sets_bind_empty: "sets M = {} \<Longrightarrow> sets (bind M f) = {{}}"
 | |
| 1110 | by (auto simp: bind_def) | |
| 1111 | ||
| 1112 | lemma space_bind_empty: "space M = {} \<Longrightarrow> space (bind M f) = {}"
 | |
| 1113 | by (simp add: bind_def) | |
| 1114 | ||
| 59048 | 1115 | lemma sets_bind[simp, measurable_cong]: | 
| 1116 |   assumes f: "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N" and M: "space M \<noteq> {}"
 | |
| 58606 | 1117 | shows "sets (bind M f) = sets N" | 
| 59048 | 1118 | using f [of "SOME x. x \<in> space M"] by (simp add: bind_nonempty M some_in_eq) | 
| 58606 | 1119 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1120 | lemma space_bind[simp]: | 
| 59048 | 1121 |   assumes "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N" and "space M \<noteq> {}"
 | 
| 58606 | 1122 | shows "space (bind M f) = space N" | 
| 59048 | 1123 | using assms by (intro sets_eq_imp_space_eq sets_bind) | 
| 58606 | 1124 | |
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1125 | lemma bind_cong_All: | 
| 58606 | 1126 | assumes "\<forall>x \<in> space M. f x = g x" | 
| 1127 | shows "bind M f = bind M g" | |
| 1128 | proof (cases "space M = {}")
 | |
| 1129 |   assume "space M \<noteq> {}"
 | |
| 1130 | hence "(SOME x. x \<in> space M) \<in> space M" by (rule_tac someI_ex) blast | |
| 1131 | with assms have "f (SOME x. x \<in> space M) = g (SOME x. x \<in> space M)" by blast | |
| 61808 | 1132 |   with \<open>space M \<noteq> {}\<close> and assms show ?thesis by (simp add: bind_nonempty cong: distr_cong)
 | 
| 58606 | 1133 | qed (simp add: bind_empty) | 
| 1134 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1135 | lemma bind_cong: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1136 | "M = N \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> bind M f = bind N g" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1137 | using bind_cong_All[of M f g] by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1138 | |
| 58606 | 1139 | lemma bind_nonempty': | 
| 1140 | assumes "f \<in> measurable M (subprob_algebra N)" "x \<in> space M" | |
| 1141 | shows "bind M f = join (distr M (subprob_algebra N) f)" | |
| 1142 | using assms | |
| 1143 | apply (subst bind_nonempty, blast) | |
| 1144 | apply (subst subprob_algebra_cong[OF sets_kernel[OF assms(1) someI_ex]], blast) | |
| 1145 | apply (simp add: subprob_algebra_cong[OF sets_kernel[OF assms]]) | |
| 1146 | done | |
| 1147 | ||
| 1148 | lemma bind_nonempty'': | |
| 1149 |   assumes "f \<in> measurable M (subprob_algebra N)" "space M \<noteq> {}"
 | |
| 1150 | shows "bind M f = join (distr M (subprob_algebra N) f)" | |
| 1151 | using assms by (auto intro: bind_nonempty') | |
| 1152 | ||
| 1153 | lemma emeasure_bind: | |
| 1154 |     "\<lbrakk>space M \<noteq> {}; f \<in> measurable M (subprob_algebra N);X \<in> sets N\<rbrakk>
 | |
| 62026 | 1155 | \<Longrightarrow> emeasure (M \<bind> f) X = \<integral>\<^sup>+x. emeasure (f x) X \<partial>M" | 
| 58606 | 1156 | by (simp add: bind_nonempty'' emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra) | 
| 1157 | ||
| 59048 | 1158 | lemma nn_integral_bind: | 
| 1159 | assumes f: "f \<in> borel_measurable B" | |
| 59000 | 1160 | assumes N: "N \<in> measurable M (subprob_algebra B)" | 
| 62026 | 1161 | shows "(\<integral>\<^sup>+x. f x \<partial>(M \<bind> N)) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)" | 
| 59000 | 1162 | proof cases | 
| 1163 |   assume M: "space M \<noteq> {}" show ?thesis
 | |
| 1164 | unfolding bind_nonempty''[OF N M] nn_integral_join[OF f sets_distr] | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1165 | by (rule nn_integral_distr[OF N]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1166 | (simp add: f nn_integral_measurable_subprob_algebra) | 
| 59000 | 1167 | qed (simp add: bind_empty space_empty[of M] nn_integral_count_space) | 
| 1168 | ||
| 1169 | lemma AE_bind: | |
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1170 | assumes N[measurable]: "N \<in> measurable M (subprob_algebra B)" | 
| 59000 | 1171 | assumes P[measurable]: "Measurable.pred B P" | 
| 62026 | 1172 | shows "(AE x in M \<bind> N. P x) \<longleftrightarrow> (AE x in M. AE y in N x. P y)" | 
| 59000 | 1173 | proof cases | 
| 1174 |   assume M: "space M = {}" show ?thesis
 | |
| 1175 | unfolding bind_empty[OF M] unfolding space_empty[OF M] by (simp add: AE_count_space) | |
| 1176 | next | |
| 1177 |   assume M: "space M \<noteq> {}"
 | |
| 59048 | 1178 | note sets_kernel[OF N, simp] | 
| 62026 | 1179 |   have *: "(\<integral>\<^sup>+x. indicator {x. \<not> P x} x \<partial>(M \<bind> N)) = (\<integral>\<^sup>+x. indicator {x\<in>space B. \<not> P x} x \<partial>(M \<bind> N))"
 | 
| 59048 | 1180 | by (intro nn_integral_cong) (simp add: space_bind[OF _ M] split: split_indicator) | 
| 59000 | 1181 | |
| 62026 | 1182 |   have "(AE x in M \<bind> N. P x) \<longleftrightarrow> (\<integral>\<^sup>+ x. integral\<^sup>N (N x) (indicator {x \<in> space B. \<not> P x}) \<partial>M) = 0"
 | 
| 59048 | 1183 | by (simp add: AE_iff_nn_integral sets_bind[OF _ M] space_bind[OF _ M] * nn_integral_bind[where B=B] | 
| 59000 | 1184 | del: nn_integral_indicator) | 
| 1185 | also have "\<dots> = (AE x in M. AE y in N x. P y)" | |
| 1186 | apply (subst nn_integral_0_iff_AE) | |
| 1187 | apply (rule measurable_compose[OF N nn_integral_measurable_subprob_algebra]) | |
| 1188 | apply measurable | |
| 1189 | apply (intro eventually_subst AE_I2) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1190 | apply (auto simp add: subprob_measurableD(1)[OF N] | 
| 59048 | 1191 | intro!: AE_iff_measurable[symmetric]) | 
| 59000 | 1192 | done | 
| 1193 | finally show ?thesis . | |
| 1194 | qed | |
| 1195 | ||
| 1196 | lemma measurable_bind': | |
| 1197 | assumes M1: "f \<in> measurable M (subprob_algebra N)" and | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61359diff
changeset | 1198 | M2: "case_prod g \<in> measurable (M \<Otimes>\<^sub>M N) (subprob_algebra R)" | 
| 59000 | 1199 | shows "(\<lambda>x. bind (f x) (g x)) \<in> measurable M (subprob_algebra R)" | 
| 1200 | proof (subst measurable_cong) | |
| 1201 | fix x assume x_in_M: "x \<in> space M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1202 |   with assms have "space (f x) \<noteq> {}"
 | 
| 59000 | 1203 | by (blast dest: subprob_space_kernel subprob_space.subprob_not_empty) | 
| 1204 | moreover from M2 x_in_M have "g x \<in> measurable (f x) (subprob_algebra R)" | |
| 1205 | by (subst measurable_cong_sets[OF sets_kernel[OF M1 x_in_M] refl]) | |
| 1206 | (auto dest: measurable_Pair2) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1207 | ultimately show "bind (f x) (g x) = join (distr (f x) (subprob_algebra R) (g x))" | 
| 59000 | 1208 | by (simp_all add: bind_nonempty'') | 
| 1209 | next | |
| 1210 | show "(\<lambda>w. join (distr (f w) (subprob_algebra R) (g w))) \<in> measurable M (subprob_algebra R)" | |
| 1211 | apply (rule measurable_compose[OF _ measurable_join]) | |
| 1212 | apply (rule measurable_distr2[OF M2 M1]) | |
| 1213 | done | |
| 1214 | qed | |
| 58606 | 1215 | |
| 59048 | 1216 | lemma measurable_bind[measurable (raw)]: | 
| 59000 | 1217 | assumes M1: "f \<in> measurable M (subprob_algebra N)" and | 
| 1218 | M2: "(\<lambda>x. g (fst x) (snd x)) \<in> measurable (M \<Otimes>\<^sub>M N) (subprob_algebra R)" | |
| 1219 | shows "(\<lambda>x. bind (f x) (g x)) \<in> measurable M (subprob_algebra R)" | |
| 1220 | using assms by (auto intro: measurable_bind' simp: measurable_split_conv) | |
| 1221 | ||
| 1222 | lemma measurable_bind2: | |
| 1223 | assumes "f \<in> measurable M (subprob_algebra N)" and "g \<in> measurable N (subprob_algebra R)" | |
| 1224 | shows "(\<lambda>x. bind (f x) g) \<in> measurable M (subprob_algebra R)" | |
| 1225 | using assms by (intro measurable_bind' measurable_const) auto | |
| 1226 | ||
| 1227 | lemma subprob_space_bind: | |
| 1228 | assumes "subprob_space M" "f \<in> measurable M (subprob_algebra N)" | |
| 62026 | 1229 | shows "subprob_space (M \<bind> f)" | 
| 1230 | proof (rule subprob_space_kernel[of "\<lambda>x. x \<bind> f"]) | |
| 1231 | show "(\<lambda>x. x \<bind> f) \<in> measurable (subprob_algebra M) (subprob_algebra N)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1232 | by (rule measurable_bind, rule measurable_ident_sets, rule refl, | 
| 59000 | 1233 | rule measurable_compose[OF measurable_snd assms(2)]) | 
| 1234 | from assms(1) show "M \<in> space (subprob_algebra M)" | |
| 1235 | by (simp add: space_subprob_algebra) | |
| 1236 | qed | |
| 58606 | 1237 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1238 | lemma | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1239 | fixes f :: "_ \<Rightarrow> real" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1240 | assumes f_measurable [measurable]: "f \<in> borel_measurable K" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1241 | and f_bounded: "\<And>x. x \<in> space K \<Longrightarrow> \<bar>f x\<bar> \<le> B" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1242 | and N [measurable]: "N \<in> measurable M (subprob_algebra K)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1243 | and fin: "finite_measure M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1244 | and M_bounded: "AE x in M. emeasure (N x) (space (N x)) \<le> ennreal B'" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1245 | shows integrable_bind: "integrable (bind M N) f" (is ?integrable) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1246 | and integral_bind: "integral\<^sup>L (bind M N) f = \<integral> x. integral\<^sup>L (N x) f \<partial>M" (is ?integral) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1247 | proof(case_tac [!] "space M = {}")
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1248 |   assume [simp]: "space M \<noteq> {}"
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1249 | interpret finite_measure M by(rule fin) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1250 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1251 | have "integrable (join (distr M (subprob_algebra K) N)) f" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1252 | using f_measurable f_bounded | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1253 | by(rule integrable_join[where B'=B'])(simp_all add: finite_measure_distr AE_distr_iff M_bounded) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1254 | then show ?integrable by(simp add: bind_nonempty''[where N=K]) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1255 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1256 | have "integral\<^sup>L (join (distr M (subprob_algebra K) N)) f = \<integral> M'. integral\<^sup>L M' f \<partial>distr M (subprob_algebra K) N" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1257 | using f_measurable f_bounded | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1258 | by(rule integral_join[where B'=B'])(simp_all add: finite_measure_distr AE_distr_iff M_bounded) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1259 | also have "\<dots> = \<integral> x. integral\<^sup>L (N x) f \<partial>M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1260 | by(rule integral_distr)(simp_all add: integral_measurable_subprob_algebra[OF _]) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1261 | finally show ?integral by(simp add: bind_nonempty''[where N=K]) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 1262 | qed(simp_all add: bind_def integrable_count_space lebesgue_integral_count_space_finite Bochner_Integration.integral_empty) | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1263 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1264 | lemma (in prob_space) prob_space_bind: | 
| 59000 | 1265 | assumes ae: "AE x in M. prob_space (N x)" | 
| 1266 | and N[measurable]: "N \<in> measurable M (subprob_algebra S)" | |
| 62026 | 1267 | shows "prob_space (M \<bind> N)" | 
| 59000 | 1268 | proof | 
| 62026 | 1269 | have "emeasure (M \<bind> N) (space (M \<bind> N)) = (\<integral>\<^sup>+x. emeasure (N x) (space (N x)) \<partial>M)" | 
| 59000 | 1270 | by (subst emeasure_bind[where N=S]) | 
| 59048 | 1271 | (auto simp: not_empty space_bind[OF sets_kernel] subprob_measurableD[OF N] intro!: nn_integral_cong) | 
| 59000 | 1272 | also have "\<dots> = (\<integral>\<^sup>+x. 1 \<partial>M)" | 
| 1273 | using ae by (intro nn_integral_cong_AE, eventually_elim) (rule prob_space.emeasure_space_1) | |
| 62026 | 1274 | finally show "emeasure (M \<bind> N) (space (M \<bind> N)) = 1" | 
| 59000 | 1275 | by (simp add: emeasure_space_1) | 
| 1276 | qed | |
| 1277 | ||
| 1278 | lemma (in subprob_space) bind_in_space: | |
| 62026 | 1279 | "A \<in> measurable M (subprob_algebra N) \<Longrightarrow> (M \<bind> A) \<in> space (subprob_algebra N)" | 
| 59048 | 1280 | by (auto simp add: space_subprob_algebra subprob_not_empty sets_kernel intro!: subprob_space_bind) | 
| 59000 | 1281 | unfold_locales | 
| 1282 | ||
| 1283 | lemma (in subprob_space) measure_bind: | |
| 1284 | assumes f: "f \<in> measurable M (subprob_algebra N)" and X: "X \<in> sets N" | |
| 62026 | 1285 | shows "measure (M \<bind> f) X = \<integral>x. measure (f x) X \<partial>M" | 
| 59000 | 1286 | proof - | 
| 62026 | 1287 | interpret Mf: subprob_space "M \<bind> f" | 
| 59000 | 1288 | by (rule subprob_space_bind[OF _ f]) unfold_locales | 
| 1289 | ||
| 1290 |   { fix x assume "x \<in> space M"
 | |
| 1291 | from f[THEN measurable_space, OF this] interpret subprob_space "f x" | |
| 1292 | by (simp add: space_subprob_algebra) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1293 | have "emeasure (f x) X = ennreal (measure (f x) X)" "measure (f x) X \<le> 1" | 
| 59000 | 1294 | by (auto simp: emeasure_eq_measure subprob_measure_le_1) } | 
| 1295 | note this[simp] | |
| 1296 | ||
| 62026 | 1297 | have "emeasure (M \<bind> f) X = \<integral>\<^sup>+x. emeasure (f x) X \<partial>M" | 
| 59000 | 1298 | using subprob_not_empty f X by (rule emeasure_bind) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1299 | also have "\<dots> = \<integral>\<^sup>+x. ennreal (measure (f x) X) \<partial>M" | 
| 59000 | 1300 | by (intro nn_integral_cong) simp | 
| 1301 | also have "\<dots> = \<integral>x. measure (f x) X \<partial>M" | |
| 1302 | by (intro nn_integral_eq_integral integrable_const_bound[where B=1] | |
| 1303 | measure_measurable_subprob_algebra2[OF _ f] pair_measureI X) | |
| 1304 | (auto simp: measure_nonneg) | |
| 1305 | finally show ?thesis | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1306 | by (simp add: Mf.emeasure_eq_measure measure_nonneg integral_nonneg) | 
| 58606 | 1307 | qed | 
| 1308 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1309 | lemma emeasure_bind_const: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1310 |     "space M \<noteq> {} \<Longrightarrow> X \<in> sets N \<Longrightarrow> subprob_space N \<Longrightarrow>
 | 
| 62026 | 1311 | emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1312 | by (simp add: bind_nonempty emeasure_join nn_integral_distr | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1313 | space_subprob_algebra measurable_emeasure_subprob_algebra) | 
| 58606 | 1314 | |
| 1315 | lemma emeasure_bind_const': | |
| 1316 | assumes "subprob_space M" "subprob_space N" | |
| 62026 | 1317 | shows "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" | 
| 58606 | 1318 | using assms | 
| 1319 | proof (case_tac "X \<in> sets N") | |
| 1320 | fix X assume "X \<in> sets N" | |
| 62026 | 1321 | thus "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" using assms | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1322 | by (subst emeasure_bind_const) | 
| 58606 | 1323 | (simp_all add: subprob_space.subprob_not_empty subprob_space.emeasure_space_le_1) | 
| 1324 | next | |
| 1325 | fix X assume "X \<notin> sets N" | |
| 62026 | 1326 | with assms show "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" | 
| 58606 | 1327 | by (simp add: sets_bind[of _ _ N] subprob_space.subprob_not_empty | 
| 1328 | space_subprob_algebra emeasure_notin_sets) | |
| 1329 | qed | |
| 1330 | ||
| 1331 | lemma emeasure_bind_const_prob_space: | |
| 1332 | assumes "prob_space M" "subprob_space N" | |
| 62026 | 1333 | shows "emeasure (M \<bind> (\<lambda>x. N)) X = emeasure N X" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1334 | using assms by (simp add: emeasure_bind_const' prob_space_imp_subprob_space | 
| 58606 | 1335 | prob_space.emeasure_space_1) | 
| 1336 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1337 | lemma bind_return: | 
| 59000 | 1338 | assumes "f \<in> measurable M (subprob_algebra N)" and "x \<in> space M" | 
| 1339 | shows "bind (return M x) f = f x" | |
| 1340 | using sets_kernel[OF assms] assms | |
| 1341 | by (simp_all add: distr_return join_return subprob_space_kernel bind_nonempty' | |
| 1342 | cong: subprob_algebra_cong) | |
| 1343 | ||
| 1344 | lemma bind_return': | |
| 1345 | shows "bind M (return M) = M" | |
| 1346 |   by (cases "space M = {}")
 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1347 | (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return' | 
| 59000 | 1348 | cong: subprob_algebra_cong) | 
| 1349 | ||
| 1350 | lemma distr_bind: | |
| 1351 |   assumes N: "N \<in> measurable M (subprob_algebra K)" "space M \<noteq> {}"
 | |
| 1352 | assumes f: "f \<in> measurable K R" | |
| 62026 | 1353 | shows "distr (M \<bind> N) R f = (M \<bind> (\<lambda>x. distr (N x) R f))" | 
| 59000 | 1354 | unfolding bind_nonempty''[OF N] | 
| 1355 | apply (subst bind_nonempty''[OF measurable_compose[OF N(1) measurable_distr] N(2)]) | |
| 1356 | apply (rule f) | |
| 1357 | apply (simp add: join_distr_distr[OF _ f, symmetric]) | |
| 1358 | apply (subst distr_distr[OF measurable_distr, OF f N(1)]) | |
| 1359 | apply (simp add: comp_def) | |
| 1360 | done | |
| 1361 | ||
| 1362 | lemma bind_distr: | |
| 1363 | assumes f[measurable]: "f \<in> measurable M X" | |
| 1364 |   assumes N[measurable]: "N \<in> measurable X (subprob_algebra K)" and "space M \<noteq> {}"
 | |
| 62026 | 1365 | shows "(distr M X f \<bind> N) = (M \<bind> (\<lambda>x. N (f x)))" | 
| 59000 | 1366 | proof - | 
| 1367 |   have "space X \<noteq> {}" "space M \<noteq> {}"
 | |
| 61808 | 1368 |     using \<open>space M \<noteq> {}\<close> f[THEN measurable_space] by auto
 | 
| 59000 | 1369 | then show ?thesis | 
| 1370 | by (simp add: bind_nonempty''[where N=K] distr_distr comp_def) | |
| 1371 | qed | |
| 1372 | ||
| 1373 | lemma bind_count_space_singleton: | |
| 1374 | assumes "subprob_space (f x)" | |
| 62026 | 1375 |   shows "count_space {x} \<bind> f = f x"
 | 
| 59000 | 1376 | proof- | 
| 1377 |   have A: "\<And>A. A \<subseteq> {x} \<Longrightarrow> A = {} \<or> A = {x}" by auto
 | |
| 1378 |   have "count_space {x} = return (count_space {x}) x"
 | |
| 1379 | by (intro measure_eqI) (auto dest: A) | |
| 62026 | 1380 | also have "... \<bind> f = f x" | 
| 59000 | 1381 | by (subst bind_return[of _ _ "f x"]) (auto simp: space_subprob_algebra assms) | 
| 1382 | finally show ?thesis . | |
| 1383 | qed | |
| 1384 | ||
| 1385 | lemma restrict_space_bind: | |
| 1386 | assumes N: "N \<in> measurable M (subprob_algebra K)" | |
| 1387 |   assumes "space M \<noteq> {}"
 | |
| 1388 |   assumes X[simp]: "X \<in> sets K" "X \<noteq> {}"
 | |
| 1389 | shows "restrict_space (bind M N) X = bind M (\<lambda>x. restrict_space (N x) X)" | |
| 1390 | proof (rule measure_eqI) | |
| 59048 | 1391 | note N_sets = sets_bind[OF sets_kernel[OF N] assms(2), simp] | 
| 1392 | note N_space = sets_eq_imp_space_eq[OF N_sets, simp] | |
| 1393 | show "sets (restrict_space (bind M N) X) = sets (bind M (\<lambda>x. restrict_space (N x) X))" | |
| 1394 | by (simp add: sets_restrict_space assms(2) sets_bind[OF sets_kernel[OF restrict_space_measurable[OF assms(4,3,1)]]]) | |
| 62026 | 1395 | fix A assume "A \<in> sets (restrict_space (M \<bind> N) X)" | 
| 59000 | 1396 | with X have "A \<in> sets K" "A \<subseteq> X" | 
| 59048 | 1397 | by (auto simp: sets_restrict_space) | 
| 62026 | 1398 | then show "emeasure (restrict_space (M \<bind> N) X) A = emeasure (M \<bind> (\<lambda>x. restrict_space (N x) X)) A" | 
| 59000 | 1399 | using assms | 
| 1400 | apply (subst emeasure_restrict_space) | |
| 59048 | 1401 | apply (simp_all add: emeasure_bind[OF assms(2,1)]) | 
| 59000 | 1402 | apply (subst emeasure_bind[OF _ restrict_space_measurable[OF _ _ N]]) | 
| 1403 | apply (auto simp: sets_restrict_space emeasure_restrict_space space_subprob_algebra | |
| 1404 | intro!: nn_integral_cong dest!: measurable_space) | |
| 1405 | done | |
| 59048 | 1406 | qed | 
| 59000 | 1407 | |
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1408 | lemma bind_restrict_space: | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1409 |   assumes A: "A \<inter> space M \<noteq> {}" "A \<inter> space M \<in> sets M"
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1410 | and f: "f \<in> measurable (restrict_space M A) (subprob_algebra N)" | 
| 62026 | 1411 | shows "restrict_space M A \<bind> f = M \<bind> (\<lambda>x. if x \<in> A then f x else null_measure (f (SOME x. x \<in> A \<and> x \<in> space M)))" | 
| 1412 | (is "?lhs = ?rhs" is "_ = M \<bind> ?f") | |
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1413 | proof - | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1414 | let ?P = "\<lambda>x. x \<in> A \<and> x \<in> space M" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1415 | let ?x = "Eps ?P" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1416 | let ?c = "null_measure (f ?x)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1417 | from A have "?P ?x" by-(rule someI_ex, blast) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1418 | hence "?x \<in> space (restrict_space M A)" by(simp add: space_restrict_space) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1419 | with f have "f ?x \<in> space (subprob_algebra N)" by(rule measurable_space) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1420 | hence sps: "subprob_space (f ?x)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1421 | and sets: "sets (f ?x) = sets N" | 
| 60067 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1422 | by(simp_all add: space_subprob_algebra) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1423 |   have "space (f ?x) \<noteq> {}" using sps by(rule subprob_space.subprob_not_empty)
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1424 | moreover have "sets ?c = sets N" by(simp add: null_measure_def)(simp add: sets) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1425 | ultimately have c: "?c \<in> space (subprob_algebra N)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1426 | by(simp add: space_subprob_algebra subprob_space_null_measure) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1427 | from f A c have f': "?f \<in> measurable M (subprob_algebra N)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1428 | by(simp add: measurable_restrict_space_iff) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1429 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1430 |   from A have [simp]: "space M \<noteq> {}" by blast
 | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1431 | |
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1432 | have "?lhs = join (distr (restrict_space M A) (subprob_algebra N) f)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1433 | using assms by(simp add: space_restrict_space bind_nonempty'') | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1434 | also have "\<dots> = join (distr M (subprob_algebra N) ?f)" | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1435 | by(rule measure_eqI)(auto simp add: emeasure_join nn_integral_distr nn_integral_restrict_space f f' A intro: nn_integral_cong) | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1436 | also have "\<dots> = ?rhs" using f' by(simp add: bind_nonempty'') | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1437 | finally show ?thesis . | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1438 | qed | 
| 
f1a5bcf5658f
lemmas about integrals over bind and join on measures
 Andreas Lochbihler parents: 
59978diff
changeset | 1439 | |
| 62026 | 1440 | lemma bind_const': "\<lbrakk>prob_space M; subprob_space N\<rbrakk> \<Longrightarrow> M \<bind> (\<lambda>x. N) = N" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1441 | by (intro measure_eqI) | 
| 58606 | 1442 | (simp_all add: space_subprob_algebra prob_space.not_empty emeasure_bind_const_prob_space) | 
| 1443 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1444 | lemma bind_return_distr: | 
| 58606 | 1445 |     "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (return N \<circ> f) = distr M N f"
 | 
| 1446 | apply (simp add: bind_nonempty) | |
| 1447 | apply (subst subprob_algebra_cong) | |
| 1448 | apply (rule sets_return) | |
| 1449 | apply (subst distr_distr[symmetric]) | |
| 1450 | apply (auto intro!: return_measurable simp: distr_distr[symmetric] join_return') | |
| 1451 | done | |
| 1452 | ||
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1453 | lemma bind_return_distr': | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1454 |   "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (\<lambda>x. return N (f x)) = distr M N f"
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1455 | using bind_return_distr[of M f N] by (simp add: comp_def) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1456 | |
| 58606 | 1457 | lemma bind_assoc: | 
| 1458 | fixes f :: "'a \<Rightarrow> 'b measure" and g :: "'b \<Rightarrow> 'c measure" | |
| 1459 | assumes M1: "f \<in> measurable M (subprob_algebra N)" and M2: "g \<in> measurable N (subprob_algebra R)" | |
| 1460 | shows "bind (bind M f) g = bind M (\<lambda>x. bind (f x) g)" | |
| 1461 | proof (cases "space M = {}")
 | |
| 1462 |   assume [simp]: "space M \<noteq> {}"
 | |
| 1463 |   from assms have [simp]: "space N \<noteq> {}" "space R \<noteq> {}"
 | |
| 1464 | by (auto simp: measurable_empty_iff space_subprob_algebra_empty_iff) | |
| 1465 | from assms have sets_fx[simp]: "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N" | |
| 1466 | by (simp add: sets_kernel) | |
| 1467 |   have ex_in: "\<And>A. A \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A" by blast
 | |
| 61808 | 1468 |   note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF \<open>space M \<noteq> {}\<close>]]]
 | 
| 1469 |                          sets_kernel[OF M2 someI_ex[OF ex_in[OF \<open>space N \<noteq> {}\<close>]]]
 | |
| 58606 | 1470 | note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)] | 
| 1471 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1472 | have "bind M (\<lambda>x. bind (f x) g) = | 
| 58606 | 1473 | join (distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f))" | 
| 1474 | by (simp add: sets_eq_imp_space_eq[OF sets_fx] bind_nonempty o_def | |
| 1475 | cong: subprob_algebra_cong distr_cong) | |
| 1476 | also have "distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f) = | |
| 1477 | distr (distr (distr M (subprob_algebra N) f) | |
| 1478 | (subprob_algebra (subprob_algebra R)) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1479 | (\<lambda>x. distr x (subprob_algebra R) g)) | 
| 58606 | 1480 | (subprob_algebra R) join" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1481 | apply (subst distr_distr, | 
| 58606 | 1482 | (blast intro: measurable_comp measurable_distr measurable_join M1 M2)+)+ | 
| 1483 | apply (simp add: o_assoc) | |
| 1484 | done | |
| 1485 | also have "join ... = bind (bind M f) g" | |
| 1486 | by (simp add: join_assoc join_distr_distr M2 bind_nonempty cong: subprob_algebra_cong) | |
| 1487 | finally show ?thesis .. | |
| 1488 | qed (simp add: bind_empty) | |
| 1489 | ||
| 1490 | lemma double_bind_assoc: | |
| 1491 | assumes Mg: "g \<in> measurable N (subprob_algebra N')" | |
| 1492 | assumes Mf: "f \<in> measurable M (subprob_algebra M')" | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61359diff
changeset | 1493 | assumes Mh: "case_prod h \<in> measurable (M \<Otimes>\<^sub>M M') N" | 
| 62026 | 1494 |   shows "do {x \<leftarrow> M; y \<leftarrow> f x; g (h x y)} = do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<bind> g"
 | 
| 58606 | 1495 | proof- | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1496 |   have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<bind> g =
 | 
| 62026 | 1497 |             do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<bind> g}"
 | 
| 58606 | 1498 | using Mh by (auto intro!: bind_assoc measurable_bind'[OF Mf] Mf Mg | 
| 1499 | measurable_compose[OF _ return_measurable] simp: measurable_split_conv) | |
| 1500 | also from Mh have "\<And>x. x \<in> space M \<Longrightarrow> h x \<in> measurable M' N" by measurable | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1501 |   hence "do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<bind> g} =
 | 
| 62026 | 1502 |             do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<bind> g}"
 | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1503 | apply (intro ballI bind_cong refl bind_assoc) | 
| 58606 | 1504 | apply (subst measurable_cong_sets[OF sets_kernel[OF Mf] refl], simp) | 
| 1505 | apply (rule measurable_compose[OF _ return_measurable], auto intro: Mg) | |
| 1506 | done | |
| 1507 | also have "\<And>x. x \<in> space M \<Longrightarrow> space (f x) = space M'" | |
| 1508 | by (intro sets_eq_imp_space_eq sets_kernel[OF Mf]) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1509 | with measurable_space[OF Mh] | 
| 62026 | 1510 |     have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<bind> g} = do {x \<leftarrow> M; y \<leftarrow> f x; g (h x y)}"
 | 
| 58606 | 1511 | by (intro ballI bind_cong bind_return[OF Mg]) (auto simp: space_pair_measure) | 
| 1512 | finally show ?thesis .. | |
| 1513 | qed | |
| 1514 | ||
| 59048 | 1515 | lemma (in prob_space) M_in_subprob[measurable (raw)]: "M \<in> space (subprob_algebra M)" | 
| 1516 | by (simp add: space_subprob_algebra) unfold_locales | |
| 1517 | ||
| 59000 | 1518 | lemma (in pair_prob_space) pair_measure_eq_bind: | 
| 62026 | 1519 | "(M1 \<Otimes>\<^sub>M M2) = (M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. return (M1 \<Otimes>\<^sub>M M2) (x, y))))" | 
| 59000 | 1520 | proof (rule measure_eqI) | 
| 1521 | have ps_M2: "prob_space M2" by unfold_locales | |
| 1522 | note return_measurable[measurable] | |
| 62026 | 1523 | show "sets (M1 \<Otimes>\<^sub>M M2) = sets (M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. return (M1 \<Otimes>\<^sub>M M2) (x, y))))" | 
| 59048 | 1524 | by (simp_all add: M1.not_empty M2.not_empty) | 
| 59000 | 1525 | fix A assume [measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" | 
| 62026 | 1526 | show "emeasure (M1 \<Otimes>\<^sub>M M2) A = emeasure (M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. return (M1 \<Otimes>\<^sub>M M2) (x, y)))) A" | 
| 59048 | 1527 | by (auto simp: M2.emeasure_pair_measure M1.not_empty M2.not_empty emeasure_bind[where N="M1 \<Otimes>\<^sub>M M2"] | 
| 59000 | 1528 | intro!: nn_integral_cong) | 
| 1529 | qed | |
| 1530 | ||
| 1531 | lemma (in pair_prob_space) bind_rotate: | |
| 1532 | assumes C[measurable]: "(\<lambda>(x, y). C x y) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (subprob_algebra N)" | |
| 62026 | 1533 | shows "(M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. C x y))) = (M2 \<bind> (\<lambda>y. M1 \<bind> (\<lambda>x. C x y)))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1534 | proof - | 
| 59000 | 1535 | interpret swap: pair_prob_space M2 M1 by unfold_locales | 
| 1536 | note measurable_bind[where N="M2", measurable] | |
| 1537 | note measurable_bind[where N="M1", measurable] | |
| 1538 | have [simp]: "M1 \<in> space (subprob_algebra M1)" "M2 \<in> space (subprob_algebra M2)" | |
| 1539 | by (auto simp: space_subprob_algebra) unfold_locales | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1540 | have "(M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. C x y))) = | 
| 62026 | 1541 | (M1 \<bind> (\<lambda>x. M2 \<bind> (\<lambda>y. return (M1 \<Otimes>\<^sub>M M2) (x, y)))) \<bind> (\<lambda>(x, y). C x y)" | 
| 59000 | 1542 | by (auto intro!: bind_cong simp: bind_return[where N=N] space_pair_measure bind_assoc[where N="M1 \<Otimes>\<^sub>M M2" and R=N]) | 
| 62026 | 1543 | also have "\<dots> = (distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))) \<bind> (\<lambda>(x, y). C x y)" | 
| 59000 | 1544 | unfolding pair_measure_eq_bind[symmetric] distr_pair_swap[symmetric] .. | 
| 62026 | 1545 | also have "\<dots> = (M2 \<bind> (\<lambda>x. M1 \<bind> (\<lambda>y. return (M2 \<Otimes>\<^sub>M M1) (x, y)))) \<bind> (\<lambda>(y, x). C x y)" | 
| 59000 | 1546 | unfolding swap.pair_measure_eq_bind[symmetric] | 
| 1547 | by (auto simp add: space_pair_measure M1.not_empty M2.not_empty bind_distr[OF _ C] intro!: bind_cong) | |
| 62026 | 1548 | also have "\<dots> = (M2 \<bind> (\<lambda>y. M1 \<bind> (\<lambda>x. C x y)))" | 
| 59000 | 1549 | by (auto intro!: bind_cong simp: bind_return[where N=N] space_pair_measure bind_assoc[where N="M2 \<Otimes>\<^sub>M M1" and R=N]) | 
| 1550 | finally show ?thesis . | |
| 1551 | qed | |
| 1552 | ||
| 62026 | 1553 | lemma bind_return'': "sets M = sets N \<Longrightarrow> M \<bind> return N = M" | 
| 59425 | 1554 |    by (cases "space M = {}")
 | 
| 1555 | (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return' | |
| 1556 | cong: subprob_algebra_cong) | |
| 1557 | ||
| 1558 | lemma (in prob_space) distr_const[simp]: | |
| 1559 | "c \<in> space N \<Longrightarrow> distr M N (\<lambda>x. c) = return N c" | |
| 1560 | by (rule measure_eqI) (auto simp: emeasure_distr emeasure_space_1) | |
| 1561 | ||
| 1562 | lemma return_count_space_eq_density: | |
| 1563 |     "return (count_space M) x = density (count_space M) (indicator {x})"
 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1564 | by (rule measure_eqI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62026diff
changeset | 1565 | (auto simp: indicator_inter_arith[symmetric] emeasure_density split: split_indicator) | 
| 59425 | 1566 | |
| 61634 | 1567 | lemma null_measure_in_space_subprob_algebra [simp]: | 
| 1568 |   "null_measure M \<in> space (subprob_algebra M) \<longleftrightarrow> space M \<noteq> {}"
 | |
| 1569 | by(simp add: space_subprob_algebra subprob_space_null_measure_iff) | |
| 1570 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1571 | subsection \<open>Giry monad on probability spaces\<close> | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1572 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1573 | definition prob_algebra :: "'a measure \<Rightarrow> 'a measure measure" where | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1574 |   "prob_algebra K = restrict_space (subprob_algebra K) {M. prob_space M}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1575 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1576 | lemma space_prob_algebra: "space (prob_algebra M) = {N. sets N = sets M \<and> prob_space N}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1577 | unfolding prob_algebra_def by (auto simp: space_subprob_algebra space_restrict_space prob_space_imp_subprob_space) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1578 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1579 | lemma measurable_measure_prob_algebra[measurable]: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1580 | "a \<in> sets A \<Longrightarrow> (\<lambda>M. Sigma_Algebra.measure M a) \<in> prob_algebra A \<rightarrow>\<^sub>M borel" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1581 | unfolding prob_algebra_def by (intro measurable_restrict_space1 measurable_measure_subprob_algebra) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1582 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1583 | lemma measurable_prob_algebraD: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1584 | "f \<in> N \<rightarrow>\<^sub>M prob_algebra M \<Longrightarrow> f \<in> N \<rightarrow>\<^sub>M subprob_algebra M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1585 | unfolding prob_algebra_def measurable_restrict_space2_iff by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1586 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1587 | lemma measure_measurable_prob_algebra2: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1588 | "Sigma (space M) A \<in> sets (M \<Otimes>\<^sub>M N) \<Longrightarrow> L \<in> M \<rightarrow>\<^sub>M prob_algebra N \<Longrightarrow> | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1589 | (\<lambda>x. Sigma_Algebra.measure (L x) (A x)) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1590 | using measure_measurable_subprob_algebra2[of M A N L] by (auto intro: measurable_prob_algebraD) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1591 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1592 | lemma measurable_prob_algebraI: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1593 | "(\<And>x. x \<in> space N \<Longrightarrow> prob_space (f x)) \<Longrightarrow> f \<in> N \<rightarrow>\<^sub>M subprob_algebra M \<Longrightarrow> f \<in> N \<rightarrow>\<^sub>M prob_algebra M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1594 | unfolding prob_algebra_def by (intro measurable_restrict_space2) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1595 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1596 | lemma measurable_distr_prob_space: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1597 | assumes f: "f \<in> M \<rightarrow>\<^sub>M N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1598 | shows "(\<lambda>M'. distr M' N f) \<in> prob_algebra M \<rightarrow>\<^sub>M prob_algebra N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1599 | unfolding prob_algebra_def measurable_restrict_space2_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1600 | proof (intro conjI measurable_restrict_space1 measurable_distr f) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1601 | show "(\<lambda>M'. distr M' N f) \<in> space (restrict_space (subprob_algebra M) (Collect prob_space)) \<rightarrow> Collect prob_space" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1602 | using f by (auto simp: space_restrict_space space_subprob_algebra intro!: prob_space.prob_space_distr) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1603 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1604 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1605 | lemma measurable_return_prob_space[measurable]: "return N \<in> N \<rightarrow>\<^sub>M prob_algebra N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1606 | by (rule measurable_prob_algebraI) (auto simp: prob_space_return) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1607 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1608 | lemma measurable_distr_prob_space2[measurable (raw)]: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1609 | assumes f: "g \<in> L \<rightarrow>\<^sub>M prob_algebra M" "(\<lambda>(x, y). f x y) \<in> L \<Otimes>\<^sub>M M \<rightarrow>\<^sub>M N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1610 | shows "(\<lambda>x. distr (g x) N (f x)) \<in> L \<rightarrow>\<^sub>M prob_algebra N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1611 | unfolding prob_algebra_def measurable_restrict_space2_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1612 | proof (intro conjI measurable_restrict_space1 measurable_distr2[where M=M] f measurable_prob_algebraD) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1613 | show "(\<lambda>x. distr (g x) N (f x)) \<in> space L \<rightarrow> Collect prob_space" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1614 | using f subprob_measurableD[OF measurable_prob_algebraD[OF f(1)]] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1615 | by (auto simp: measurable_restrict_space2_iff prob_algebra_def | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1616 | intro!: prob_space.prob_space_distr) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1617 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1618 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1619 | lemma measurable_bind_prob_space: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1620 | assumes f: "f \<in> M \<rightarrow>\<^sub>M prob_algebra N" and g: "g \<in> N \<rightarrow>\<^sub>M prob_algebra R" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1621 | shows "(\<lambda>x. bind (f x) g) \<in> M \<rightarrow>\<^sub>M prob_algebra R" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1622 | unfolding prob_algebra_def measurable_restrict_space2_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1623 | proof (intro conjI measurable_restrict_space1 measurable_bind2[where N=N] f g measurable_prob_algebraD) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1624 | show "(\<lambda>x. f x \<bind> g) \<in> space M \<rightarrow> Collect prob_space" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1625 | using g f subprob_measurableD[OF measurable_prob_algebraD[OF f]] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1626 | by (auto simp: measurable_restrict_space2_iff prob_algebra_def | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1627 | intro!: prob_space.prob_space_bind[where S=R] AE_I2) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1628 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1629 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1630 | lemma measurable_bind_prob_space2[measurable (raw)]: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1631 | assumes f: "f \<in> M \<rightarrow>\<^sub>M prob_algebra N" and g: "(\<lambda>(x, y). g x y) \<in> (M \<Otimes>\<^sub>M N) \<rightarrow>\<^sub>M prob_algebra R" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1632 | shows "(\<lambda>x. bind (f x) (g x)) \<in> M \<rightarrow>\<^sub>M prob_algebra R" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1633 | unfolding prob_algebra_def measurable_restrict_space2_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1634 | proof (intro conjI measurable_restrict_space1 measurable_bind[where N=N] f g measurable_prob_algebraD) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1635 | show "(\<lambda>x. f x \<bind> g x) \<in> space M \<rightarrow> Collect prob_space" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1636 | using g f subprob_measurableD[OF measurable_prob_algebraD[OF f]] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1637 | using measurable_space[OF g] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1638 | by (auto simp: measurable_restrict_space2_iff prob_algebra_def space_pair_measure Pi_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1639 | intro!: prob_space.prob_space_bind[where S=R] AE_I2) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1640 | qed (insert g, simp) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1641 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1642 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1643 | lemma measurable_prob_algebra_generated: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1644 | assumes eq: "sets N = sigma_sets \<Omega> G" and "Int_stable G" "G \<subseteq> Pow \<Omega>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1645 | assumes subsp: "\<And>a. a \<in> space M \<Longrightarrow> prob_space (K a)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1646 | assumes sets: "\<And>a. a \<in> space M \<Longrightarrow> sets (K a) = sets N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1647 | assumes "\<And>A. A \<in> G \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1648 | shows "K \<in> measurable M (prob_algebra N)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1649 | unfolding measurable_restrict_space2_iff prob_algebra_def | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1650 | proof | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1651 | show "K \<in> M \<rightarrow>\<^sub>M subprob_algebra N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1652 | proof (rule measurable_subprob_algebra_generated[OF assms(1,2,3) _ assms(5,6)]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1653 | fix a assume "a \<in> space M" then show "subprob_space (K a)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1654 | using subsp[of a] by (intro prob_space_imp_subprob_space) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1655 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1656 | have "(\<lambda>a. emeasure (K a) \<Omega>) \<in> borel_measurable M \<longleftrightarrow> (\<lambda>a. 1::ennreal) \<in> borel_measurable M" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1657 | using sets_eq_imp_space_eq[of "sigma \<Omega> G" N] \<open>G \<subseteq> Pow \<Omega>\<close> eq sets_eq_imp_space_eq[OF sets] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1658 | prob_space.emeasure_space_1[OF subsp] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1659 | by (intro measurable_cong) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1660 | then show "(\<lambda>a. emeasure (K a) \<Omega>) \<in> borel_measurable M" by simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1661 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1662 | qed (insert subsp, auto) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1663 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1664 | lemma in_space_prob_algebra: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1665 | "x \<in> space (prob_algebra M) \<Longrightarrow> emeasure x (space M) = 1" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1666 | unfolding prob_algebra_def space_restrict_space space_subprob_algebra | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1667 | by (auto dest!: prob_space.emeasure_space_1 sets_eq_imp_space_eq) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1668 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1669 | lemma prob_space_pair: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1670 | assumes "prob_space M" "prob_space N" shows "prob_space (M \<Otimes>\<^sub>M N)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1671 | proof - | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1672 | interpret M: prob_space M by fact | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1673 | interpret N: prob_space N by fact | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1674 | interpret P: pair_prob_space M N proof qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1675 | show ?thesis | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1676 | by unfold_locales | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1677 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1678 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1679 | lemma measurable_pair_prob[measurable]: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1680 | "f \<in> M \<rightarrow>\<^sub>M prob_algebra N \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M prob_algebra L \<Longrightarrow> (\<lambda>x. f x \<Otimes>\<^sub>M g x) \<in> M \<rightarrow>\<^sub>M prob_algebra (N \<Otimes>\<^sub>M L)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1681 | unfolding prob_algebra_def measurable_restrict_space2_iff | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1682 | by (auto intro!: measurable_pair_measure prob_space_pair) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1683 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1684 | lemma emeasure_bind_prob_algebra: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1685 | assumes A: "A \<in> space (prob_algebra N)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1686 | assumes B: "B \<in> N \<rightarrow>\<^sub>M prob_algebra L" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1687 | assumes X: "X \<in> sets L" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1688 | shows "emeasure (bind A B) X = (\<integral>\<^sup>+x. emeasure (B x) X \<partial>A)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1689 | using A B | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1690 | by (intro emeasure_bind[OF _ _ X]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1691 | (auto simp: space_prob_algebra measurable_prob_algebraD cong: measurable_cong_sets intro!: prob_space.not_empty) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1692 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1693 | lemma prob_space_bind': | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1694 | assumes A: "A \<in> space (prob_algebra M)" and B: "B \<in> M \<rightarrow>\<^sub>M prob_algebra N" shows "prob_space (A \<bind> B)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1695 | using measurable_bind_prob_space[OF measurable_const, OF A B, THEN measurable_space, of undefined "count_space UNIV"] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1696 | by (simp add: space_prob_algebra) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1697 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1698 | lemma sets_bind': | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1699 | assumes A: "A \<in> space (prob_algebra M)" and B: "B \<in> M \<rightarrow>\<^sub>M prob_algebra N" shows "sets (A \<bind> B) = sets N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1700 | using measurable_bind_prob_space[OF measurable_const, OF A B, THEN measurable_space, of undefined "count_space UNIV"] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1701 | by (simp add: space_prob_algebra) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1702 | |
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1703 | lemma bind_cong_AE': | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1704 | assumes M: "M \<in> space (prob_algebra L)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1705 | and f: "f \<in> L \<rightarrow>\<^sub>M prob_algebra N" and g: "g \<in> L \<rightarrow>\<^sub>M prob_algebra N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1706 | and ae: "AE x in M. f x = g x" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1707 | shows "bind M f = bind M g" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1708 | proof (rule measure_eqI) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1709 | show "sets (M \<bind> f) = sets (M \<bind> g)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1710 | unfolding sets_bind'[OF M f] sets_bind'[OF M g] .. | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1711 | show "A \<in> sets (M \<bind> f) \<Longrightarrow> emeasure (M \<bind> f) A = emeasure (M \<bind> g) A" for A | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1712 | unfolding sets_bind'[OF M f] | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1713 | using emeasure_bind_prob_algebra[OF M f, of A] emeasure_bind_prob_algebra[OF M g, of A] ae | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1714 | by (auto intro: nn_integral_cong_AE) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1715 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63886diff
changeset | 1716 | |
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1717 | lemma density_discrete: | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1718 |   "countable A \<Longrightarrow> sets N = Set.Pow A \<Longrightarrow> (\<And>x. f x \<ge> 0) \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x = emeasure N {x}) \<Longrightarrow>
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1719 | density (count_space A) f = N" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1720 | by (rule measure_eqI_countable[of _ A]) (auto simp: emeasure_density) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1721 | |
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1722 | lemma distr_density_discrete: | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1723 | fixes f' | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1724 | assumes "countable A" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1725 | assumes "f' \<in> borel_measurable M" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1726 | assumes "g \<in> measurable M (count_space A)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1727 | defines "f \<equiv> \<lambda>x. \<integral>\<^sup>+t. (if g t = x then 1 else 0) * f' t \<partial>M" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1728 | assumes "\<And>x. x \<in> space M \<Longrightarrow> g x \<in> A" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1729 | shows "density (count_space A) (\<lambda>x. f x) = distr (density M f') (count_space A) g" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1730 | proof (rule density_discrete) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1731 | fix x assume x: "x \<in> A" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1732 |   have "f x = \<integral>\<^sup>+t. indicator (g -` {x} \<inter> space M) t * f' t \<partial>M" (is "_ = ?I") unfolding f_def
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1733 | by (intro nn_integral_cong) (simp split: split_indicator) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1734 |   also from x have in_sets: "g -` {x} \<inter> space M \<in> sets M"
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1735 | by (intro measurable_sets[OF assms(3)]) simp | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1736 |   have "?I = emeasure (density M f') (g -` {x} \<inter> space M)" unfolding f_def
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1737 | by (subst emeasure_density[OF assms(2) in_sets], subst mult.commute) (rule refl) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1738 |   also from assms(3) x have "... = emeasure (distr (density M f') (count_space A) g) {x}"
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1739 | by (subst emeasure_distr) simp_all | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1740 |   finally show "f x = emeasure (distr (density M f') (count_space A) g) {x}" .
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1741 | qed (insert assms, auto) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1742 | |
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1743 | lemma bind_cong_AE: | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1744 | assumes "M = N" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1745 | assumes f: "f \<in> measurable N (subprob_algebra B)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1746 | assumes g: "g \<in> measurable N (subprob_algebra B)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1747 | assumes ae: "AE x in N. f x = g x" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1748 | shows "bind M f = bind N g" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1749 | proof cases | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1750 |   assume "space N = {}" then show ?thesis
 | 
| 67226 | 1751 | using \<open>M = N\<close> by (simp add: bind_empty) | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1752 | next | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1753 |   assume "space N \<noteq> {}"
 | 
| 67226 | 1754 | show ?thesis unfolding \<open>M = N\<close> | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1755 | proof (rule measure_eqI) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1756 | have *: "sets (N \<bind> f) = sets B" | 
| 67226 | 1757 |       using sets_bind[OF sets_kernel[OF f] \<open>space N \<noteq> {}\<close>] by simp
 | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1758 | then show "sets (N \<bind> f) = sets (N \<bind> g)" | 
| 67226 | 1759 |       using sets_bind[OF sets_kernel[OF g] \<open>space N \<noteq> {}\<close>] by auto
 | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1760 | fix A assume "A \<in> sets (N \<bind> f)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1761 | then have "A \<in> sets B" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1762 | unfolding * . | 
| 67226 | 1763 |     with ae f g \<open>space N \<noteq> {}\<close> show "emeasure (N \<bind> f) A = emeasure (N \<bind> g) A"
 | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1764 | by (subst (1 2) emeasure_bind[where N=B]) (auto intro!: nn_integral_cong_AE) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1765 | qed | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1766 | qed | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1767 | |
| 69546 
27dae626822b
prefer naming convention from datatype package for strong congruence rules
 haftmann parents: 
69260diff
changeset | 1768 | lemma bind_cong_simp: "M = N \<Longrightarrow> (\<And>x. x\<in>space M =simp=> f x = g x) \<Longrightarrow> bind M f = bind N g" | 
| 64010 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1769 | by (auto simp: simp_implies_def intro!: bind_cong) | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1770 | |
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1771 | lemma sets_bind_measurable: | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1772 | assumes f: "f \<in> measurable M (subprob_algebra B)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1773 |   assumes M: "space M \<noteq> {}"
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1774 | shows "sets (M \<bind> f) = sets B" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1775 | using M by (intro sets_bind[OF sets_kernel[OF f]]) auto | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1776 | |
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1777 | lemma space_bind_measurable: | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1778 | assumes f: "f \<in> measurable M (subprob_algebra B)" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1779 |   assumes M: "space M \<noteq> {}"
 | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1780 | shows "space (M \<bind> f) = space B" | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1781 | using M by (intro space_bind[OF sets_kernel[OF f]]) auto | 
| 
9c99fccce3cf
Probability: move some theorems from AFP/Density_Compiler
 hoelzl parents: 
64008diff
changeset | 1782 | |
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1783 | lemma bind_distr_return: | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1784 |   "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> g \<in> N \<rightarrow>\<^sub>M L \<Longrightarrow> space M \<noteq> {} \<Longrightarrow>
 | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1785 | distr M N f \<bind> (\<lambda>x. return L (g x)) = distr M L (\<lambda>x. g (f x))" | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1786 | by (subst bind_distr[OF _ measurable_compose[OF _ return_measurable]]) | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1787 | (auto intro!: bind_return_distr') | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64010diff
changeset | 1788 | |
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1789 | lemma (in prob_space) AE_eq_constD: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1790 | assumes "AE x in M. x = y" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1791 | shows "M = return M y" "y \<in> space M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1792 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1793 | have "AE x in M. x \<in> space M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1794 | by auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1795 | with assms have "AE x in M. y \<in> space M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1796 | by eventually_elim auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1797 | thus "y \<in> space M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1798 | by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1799 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1800 | show "M = return M y" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1801 | proof (rule measure_eqI) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1802 | fix X assume X: "X \<in> sets M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1803 |     have "AE x in M. (x \<in> X) = (x \<in> (if y \<in> X then space M else {}))"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1804 | using assms by eventually_elim (use X \<open>y \<in> space M\<close> in auto) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1805 |     hence "emeasure M X = emeasure M (if y \<in> X then space M else {})"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1806 | using X by (intro emeasure_eq_AE) auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1807 | also have "\<dots> = emeasure (return M y) X" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1808 | using X by (auto simp: emeasure_space_1) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1809 | finally show "emeasure M X = \<dots>" . | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1810 | qed auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1811 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
69861diff
changeset | 1812 | |
| 58606 | 1813 | end |