| 58606 |      1 | (*  Title:      HOL/Probability/Giry_Monad.thy
 | 
|  |      2 |     Author:     Johannes Hölzl, TU München
 | 
|  |      3 |     Author:     Manuel Eberl, TU München
 | 
|  |      4 | 
 | 
|  |      5 | Defines the subprobability spaces, the subprobability functor and the Giry monad on subprobability
 | 
|  |      6 | spaces.
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | theory Giry_Monad
 | 
|  |     10 |   imports Probability_Measure "~~/src/HOL/Library/Monad_Syntax"
 | 
|  |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | section {* Sub-probability spaces *}
 | 
|  |     14 | 
 | 
|  |     15 | locale subprob_space = finite_measure +
 | 
|  |     16 |   assumes emeasure_space_le_1: "emeasure M (space M) \<le> 1"
 | 
|  |     17 |   assumes subprob_not_empty: "space M \<noteq> {}"
 | 
|  |     18 | 
 | 
|  |     19 | lemma subprob_spaceI[Pure.intro!]:
 | 
|  |     20 |   assumes *: "emeasure M (space M) \<le> 1"
 | 
|  |     21 |   assumes "space M \<noteq> {}"
 | 
|  |     22 |   shows "subprob_space M"
 | 
|  |     23 | proof -
 | 
|  |     24 |   interpret finite_measure M
 | 
|  |     25 |   proof
 | 
|  |     26 |     show "emeasure M (space M) \<noteq> \<infinity>" using * by auto
 | 
|  |     27 |   qed
 | 
|  |     28 |   show "subprob_space M" by default fact+
 | 
|  |     29 | qed
 | 
|  |     30 | 
 | 
|  |     31 | lemma prob_space_imp_subprob_space:
 | 
|  |     32 |   "prob_space M \<Longrightarrow> subprob_space M"
 | 
|  |     33 |   by (rule subprob_spaceI) (simp_all add: prob_space.emeasure_space_1 prob_space.not_empty)
 | 
|  |     34 | 
 | 
|  |     35 | sublocale prob_space \<subseteq> subprob_space
 | 
|  |     36 |   by (rule subprob_spaceI) (simp_all add: emeasure_space_1 not_empty)
 | 
|  |     37 | 
 | 
|  |     38 | lemma (in subprob_space) subprob_space_distr:
 | 
|  |     39 |   assumes f: "f \<in> measurable M M'" and "space M' \<noteq> {}" shows "subprob_space (distr M M' f)"
 | 
|  |     40 | proof (rule subprob_spaceI)
 | 
|  |     41 |   have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
 | 
|  |     42 |   with f show "emeasure (distr M M' f) (space (distr M M' f)) \<le> 1"
 | 
|  |     43 |     by (auto simp: emeasure_distr emeasure_space_le_1)
 | 
|  |     44 |   show "space (distr M M' f) \<noteq> {}" by (simp add: assms)
 | 
|  |     45 | qed
 | 
|  |     46 | 
 | 
|  |     47 | lemma (in subprob_space) subprob_measure_le_1: "emeasure M X \<le> 1"
 | 
|  |     48 |   by (rule order.trans[OF emeasure_space emeasure_space_le_1])
 | 
|  |     49 | 
 | 
|  |     50 | locale pair_subprob_space = 
 | 
|  |     51 |   pair_sigma_finite M1 M2 + M1: subprob_space M1 + M2: subprob_space M2 for M1 M2
 | 
|  |     52 | 
 | 
|  |     53 | sublocale pair_subprob_space \<subseteq> P: subprob_space "M1 \<Otimes>\<^sub>M M2"
 | 
|  |     54 | proof
 | 
|  |     55 |   have "\<And>a b. \<lbrakk>a \<ge> 0; b \<ge> 0; a \<le> 1; b \<le> 1\<rbrakk> \<Longrightarrow> a * b \<le> (1::ereal)"
 | 
|  |     56 |     by (metis comm_monoid_mult_class.mult.left_neutral dual_order.trans ereal_mult_right_mono)
 | 
|  |     57 |   from this[OF _ _ M1.emeasure_space_le_1 M2.emeasure_space_le_1]
 | 
|  |     58 |     show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) \<le> 1"
 | 
|  |     59 |     by (simp add: M2.emeasure_pair_measure_Times space_pair_measure emeasure_nonneg)
 | 
|  |     60 |   from M1.subprob_not_empty and M2.subprob_not_empty show "space (M1 \<Otimes>\<^sub>M M2) \<noteq> {}"
 | 
|  |     61 |     by (simp add: space_pair_measure)
 | 
|  |     62 | qed
 | 
|  |     63 | 
 | 
|  |     64 | definition subprob_algebra :: "'a measure \<Rightarrow> 'a measure measure" where
 | 
|  |     65 |   "subprob_algebra K =
 | 
|  |     66 |     (\<Squnion>\<^sub>\<sigma> A\<in>sets K. vimage_algebra {M. subprob_space M \<and> sets M = sets K} (\<lambda>M. emeasure M A) borel)"
 | 
|  |     67 | 
 | 
|  |     68 | lemma space_subprob_algebra: "space (subprob_algebra A) = {M. subprob_space M \<and> sets M = sets A}"
 | 
|  |     69 |   by (auto simp add: subprob_algebra_def space_Sup_sigma)
 | 
|  |     70 | 
 | 
|  |     71 | lemma subprob_algebra_cong: "sets M = sets N \<Longrightarrow> subprob_algebra M = subprob_algebra N"
 | 
|  |     72 |   by (simp add: subprob_algebra_def)
 | 
|  |     73 | 
 | 
|  |     74 | lemma measurable_emeasure_subprob_algebra[measurable]: 
 | 
|  |     75 |   "a \<in> sets A \<Longrightarrow> (\<lambda>M. emeasure M a) \<in> borel_measurable (subprob_algebra A)"
 | 
|  |     76 |   by (auto intro!: measurable_Sup_sigma1 measurable_vimage_algebra1 simp: subprob_algebra_def)
 | 
|  |     77 | 
 | 
|  |     78 | context
 | 
|  |     79 |   fixes K M N assumes K: "K \<in> measurable M (subprob_algebra N)"
 | 
|  |     80 | begin
 | 
|  |     81 | 
 | 
|  |     82 | lemma subprob_space_kernel: "a \<in> space M \<Longrightarrow> subprob_space (K a)"
 | 
|  |     83 |   using measurable_space[OF K] by (simp add: space_subprob_algebra)
 | 
|  |     84 | 
 | 
|  |     85 | lemma sets_kernel: "a \<in> space M \<Longrightarrow> sets (K a) = sets N"
 | 
|  |     86 |   using measurable_space[OF K] by (simp add: space_subprob_algebra)
 | 
|  |     87 | 
 | 
|  |     88 | lemma measurable_emeasure_kernel[measurable]: 
 | 
|  |     89 |     "A \<in> sets N \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M"
 | 
|  |     90 |   using measurable_compose[OF K measurable_emeasure_subprob_algebra] .
 | 
|  |     91 | 
 | 
|  |     92 | end
 | 
|  |     93 | 
 | 
|  |     94 | lemma measurable_subprob_algebra:
 | 
|  |     95 |   "(\<And>a. a \<in> space M \<Longrightarrow> subprob_space (K a)) \<Longrightarrow>
 | 
|  |     96 |   (\<And>a. a \<in> space M \<Longrightarrow> sets (K a) = sets N) \<Longrightarrow>
 | 
|  |     97 |   (\<And>A. A \<in> sets N \<Longrightarrow> (\<lambda>a. emeasure (K a) A) \<in> borel_measurable M) \<Longrightarrow>
 | 
|  |     98 |   K \<in> measurable M (subprob_algebra N)"
 | 
|  |     99 |   by (auto intro!: measurable_Sup_sigma2 measurable_vimage_algebra2 simp: subprob_algebra_def)
 | 
|  |    100 | 
 | 
|  |    101 | lemma space_subprob_algebra_empty_iff:
 | 
|  |    102 |   "space (subprob_algebra N) = {} \<longleftrightarrow> space N = {}"
 | 
|  |    103 | proof
 | 
|  |    104 |   have "\<And>x. x \<in> space N \<Longrightarrow> density N (\<lambda>_. 0) \<in> space (subprob_algebra N)"
 | 
|  |    105 |     by (auto simp: space_subprob_algebra emeasure_density intro!: subprob_spaceI)
 | 
|  |    106 |   then show "space (subprob_algebra N) = {} \<Longrightarrow> space N = {}"
 | 
|  |    107 |     by auto
 | 
|  |    108 | next
 | 
|  |    109 |   assume "space N = {}"
 | 
|  |    110 |   hence "sets N = {{}}" by (simp add: space_empty_iff)
 | 
|  |    111 |   moreover have "\<And>M. subprob_space M \<Longrightarrow> sets M \<noteq> {{}}"
 | 
|  |    112 |     by (simp add: subprob_space.subprob_not_empty space_empty_iff[symmetric])
 | 
|  |    113 |   ultimately show "space (subprob_algebra N) = {}" by (auto simp: space_subprob_algebra)
 | 
|  |    114 | qed
 | 
|  |    115 | 
 | 
|  |    116 | lemma measurable_distr:
 | 
|  |    117 |   assumes [measurable]: "f \<in> measurable M N"
 | 
|  |    118 |   shows "(\<lambda>M'. distr M' N f) \<in> measurable (subprob_algebra M) (subprob_algebra N)"
 | 
|  |    119 | proof (cases "space N = {}")
 | 
|  |    120 |   assume not_empty: "space N \<noteq> {}"
 | 
|  |    121 |   show ?thesis
 | 
|  |    122 |   proof (rule measurable_subprob_algebra)
 | 
|  |    123 |     fix A assume A: "A \<in> sets N"
 | 
|  |    124 |     then have "(\<lambda>M'. emeasure (distr M' N f) A) \<in> borel_measurable (subprob_algebra M) \<longleftrightarrow>
 | 
|  |    125 |       (\<lambda>M'. emeasure M' (f -` A \<inter> space M)) \<in> borel_measurable (subprob_algebra M)"
 | 
|  |    126 |       by (intro measurable_cong)
 | 
|  |    127 |          (auto simp: emeasure_distr space_subprob_algebra dest: sets_eq_imp_space_eq cong: measurable_cong)
 | 
|  |    128 |     also have "\<dots>"
 | 
|  |    129 |       using A by (intro measurable_emeasure_subprob_algebra) simp
 | 
|  |    130 |     finally show "(\<lambda>M'. emeasure (distr M' N f) A) \<in> borel_measurable (subprob_algebra M)" .
 | 
|  |    131 |   qed (auto intro!: subprob_space.subprob_space_distr simp: space_subprob_algebra not_empty)
 | 
|  |    132 | qed (insert assms, auto simp: measurable_empty_iff space_subprob_algebra_empty_iff)
 | 
|  |    133 | 
 | 
|  |    134 | section {* Properties of return *}
 | 
|  |    135 | 
 | 
|  |    136 | definition return :: "'a measure \<Rightarrow> 'a \<Rightarrow> 'a measure" where
 | 
|  |    137 |   "return R x = measure_of (space R) (sets R) (\<lambda>A. indicator A x)"
 | 
|  |    138 | 
 | 
|  |    139 | lemma space_return[simp]: "space (return M x) = space M"
 | 
|  |    140 |   by (simp add: return_def)
 | 
|  |    141 | 
 | 
|  |    142 | lemma sets_return[simp]: "sets (return M x) = sets M"
 | 
|  |    143 |   by (simp add: return_def)
 | 
|  |    144 | 
 | 
|  |    145 | lemma measurable_return1[simp]: "measurable (return N x) L = measurable N L"
 | 
|  |    146 |   by (simp cong: measurable_cong_sets) 
 | 
|  |    147 | 
 | 
|  |    148 | lemma measurable_return2[simp]: "measurable L (return N x) = measurable L N"
 | 
|  |    149 |   by (simp cong: measurable_cong_sets) 
 | 
|  |    150 | 
 | 
|  |    151 | lemma emeasure_return[simp]:
 | 
|  |    152 |   assumes "A \<in> sets M"
 | 
|  |    153 |   shows "emeasure (return M x) A = indicator A x"
 | 
|  |    154 | proof (rule emeasure_measure_of[OF return_def])
 | 
|  |    155 |   show "sets M \<subseteq> Pow (space M)" by (rule sets.space_closed)
 | 
|  |    156 |   show "positive (sets (return M x)) (\<lambda>A. indicator A x)" by (simp add: positive_def)
 | 
|  |    157 |   from assms show "A \<in> sets (return M x)" unfolding return_def by simp
 | 
|  |    158 |   show "countably_additive (sets (return M x)) (\<lambda>A. indicator A x)"
 | 
|  |    159 |     by (auto intro: countably_additiveI simp: suminf_indicator)
 | 
|  |    160 | qed
 | 
|  |    161 | 
 | 
|  |    162 | lemma prob_space_return: "x \<in> space M \<Longrightarrow> prob_space (return M x)"
 | 
|  |    163 |   by rule simp
 | 
|  |    164 | 
 | 
|  |    165 | lemma subprob_space_return: "x \<in> space M \<Longrightarrow> subprob_space (return M x)"
 | 
|  |    166 |   by (intro prob_space_return prob_space_imp_subprob_space)
 | 
|  |    167 | 
 | 
|  |    168 | lemma AE_return:
 | 
|  |    169 |   assumes [simp]: "x \<in> space M" and [measurable]: "Measurable.pred M P"
 | 
|  |    170 |   shows "(AE y in return M x. P y) \<longleftrightarrow> P x"
 | 
|  |    171 | proof -
 | 
|  |    172 |   have "(AE y in return M x. y \<notin> {x\<in>space M. \<not> P x}) \<longleftrightarrow> P x"
 | 
|  |    173 |     by (subst AE_iff_null_sets[symmetric]) (simp_all add: null_sets_def split: split_indicator)
 | 
|  |    174 |   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)"
 | 
|  |    175 |     by (rule AE_cong) auto
 | 
|  |    176 |   finally show ?thesis .
 | 
|  |    177 | qed
 | 
|  |    178 |   
 | 
|  |    179 | lemma nn_integral_return:
 | 
|  |    180 |   assumes "g x \<ge> 0" "x \<in> space M" "g \<in> borel_measurable M"
 | 
|  |    181 |   shows "(\<integral>\<^sup>+ a. g a \<partial>return M x) = g x"
 | 
|  |    182 | proof-
 | 
|  |    183 |   interpret prob_space "return M x" by (rule prob_space_return[OF `x \<in> space M`])
 | 
|  |    184 |   have "(\<integral>\<^sup>+ a. g a \<partial>return M x) = (\<integral>\<^sup>+ a. g x \<partial>return M x)" using assms
 | 
|  |    185 |     by (intro nn_integral_cong_AE) (auto simp: AE_return)
 | 
|  |    186 |   also have "... = g x"
 | 
|  |    187 |     using nn_integral_const[OF `g x \<ge> 0`, of "return M x"] emeasure_space_1 by simp
 | 
|  |    188 |   finally show ?thesis .
 | 
|  |    189 | qed
 | 
|  |    190 | 
 | 
|  |    191 | lemma return_measurable: "return N \<in> measurable N (subprob_algebra N)"
 | 
|  |    192 |   by (rule measurable_subprob_algebra) (auto simp: subprob_space_return)
 | 
|  |    193 | 
 | 
|  |    194 | lemma distr_return:
 | 
|  |    195 |   assumes "f \<in> measurable M N" and "x \<in> space M"
 | 
|  |    196 |   shows "distr (return M x) N f = return N (f x)"
 | 
|  |    197 |   using assms by (intro measure_eqI) (simp_all add: indicator_def emeasure_distr)
 | 
|  |    198 | 
 | 
|  |    199 | definition "select_sets M = (SOME N. sets M = sets (subprob_algebra N))"
 | 
|  |    200 | 
 | 
|  |    201 | lemma select_sets1:
 | 
|  |    202 |   "sets M = sets (subprob_algebra N) \<Longrightarrow> sets M = sets (subprob_algebra (select_sets M))"
 | 
|  |    203 |   unfolding select_sets_def by (rule someI)
 | 
|  |    204 | 
 | 
|  |    205 | lemma sets_select_sets[simp]:
 | 
|  |    206 |   assumes sets: "sets M = sets (subprob_algebra N)"
 | 
|  |    207 |   shows "sets (select_sets M) = sets N"
 | 
|  |    208 |   unfolding select_sets_def
 | 
|  |    209 | proof (rule someI2)
 | 
|  |    210 |   show "sets M = sets (subprob_algebra N)"
 | 
|  |    211 |     by fact
 | 
|  |    212 | next
 | 
|  |    213 |   fix L assume "sets M = sets (subprob_algebra L)"
 | 
|  |    214 |   with sets have eq: "space (subprob_algebra N) = space (subprob_algebra L)"
 | 
|  |    215 |     by (intro sets_eq_imp_space_eq) simp
 | 
|  |    216 |   show "sets L = sets N"
 | 
|  |    217 |   proof cases
 | 
|  |    218 |     assume "space (subprob_algebra N) = {}"
 | 
|  |    219 |     with space_subprob_algebra_empty_iff[of N] space_subprob_algebra_empty_iff[of L]
 | 
|  |    220 |     show ?thesis
 | 
|  |    221 |       by (simp add: eq space_empty_iff)
 | 
|  |    222 |   next
 | 
|  |    223 |     assume "space (subprob_algebra N) \<noteq> {}"
 | 
|  |    224 |     with eq show ?thesis
 | 
|  |    225 |       by (fastforce simp add: space_subprob_algebra)
 | 
|  |    226 |   qed
 | 
|  |    227 | qed
 | 
|  |    228 | 
 | 
|  |    229 | lemma space_select_sets[simp]:
 | 
|  |    230 |   "sets M = sets (subprob_algebra N) \<Longrightarrow> space (select_sets M) = space N"
 | 
|  |    231 |   by (intro sets_eq_imp_space_eq sets_select_sets)
 | 
|  |    232 | 
 | 
|  |    233 | section {* Join *}
 | 
|  |    234 | 
 | 
|  |    235 | definition join :: "'a measure measure \<Rightarrow> 'a measure" where
 | 
|  |    236 |   "join M = measure_of (space (select_sets M)) (sets (select_sets M)) (\<lambda>B. \<integral>\<^sup>+ M'. emeasure M' B \<partial>M)"
 | 
|  |    237 | 
 | 
|  |    238 | lemma
 | 
|  |    239 |   shows space_join[simp]: "space (join M) = space (select_sets M)"
 | 
|  |    240 |     and sets_join[simp]: "sets (join M) = sets (select_sets M)"
 | 
|  |    241 |   by (simp_all add: join_def)
 | 
|  |    242 | 
 | 
|  |    243 | lemma emeasure_join:
 | 
|  |    244 |   assumes M[simp]: "sets M = sets (subprob_algebra N)" and A: "A \<in> sets N"
 | 
|  |    245 |   shows "emeasure (join M) A = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)"
 | 
|  |    246 | proof (rule emeasure_measure_of[OF join_def])
 | 
|  |    247 |   have eq: "borel_measurable M = borel_measurable (subprob_algebra N)"
 | 
|  |    248 |     by auto
 | 
|  |    249 |   show "countably_additive (sets (join M)) (\<lambda>B. \<integral>\<^sup>+ M'. emeasure M' B \<partial>M)"
 | 
|  |    250 |   proof (rule countably_additiveI)
 | 
|  |    251 |     fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (join M)" "disjoint_family A"
 | 
|  |    252 |     have "(\<Sum>i. \<integral>\<^sup>+ M'. emeasure M' (A i) \<partial>M) = (\<integral>\<^sup>+M'. (\<Sum>i. emeasure M' (A i)) \<partial>M)"
 | 
|  |    253 |       using A by (subst nn_integral_suminf) (auto simp: measurable_emeasure_subprob_algebra eq)
 | 
|  |    254 |     also have "\<dots> = (\<integral>\<^sup>+M'. emeasure M' (\<Union>i. A i) \<partial>M)"
 | 
|  |    255 |     proof (rule nn_integral_cong)
 | 
|  |    256 |       fix M' assume "M' \<in> space M"
 | 
|  |    257 |       then show "(\<Sum>i. emeasure M' (A i)) = emeasure M' (\<Union>i. A i)"
 | 
|  |    258 |         using A sets_eq_imp_space_eq[OF M] by (simp add: suminf_emeasure space_subprob_algebra)
 | 
|  |    259 |     qed
 | 
|  |    260 |     finally show "(\<Sum>i. \<integral>\<^sup>+M'. emeasure M' (A i) \<partial>M) = (\<integral>\<^sup>+M'. emeasure M' (\<Union>i. A i) \<partial>M)" .
 | 
|  |    261 |   qed
 | 
|  |    262 | qed (auto simp: A sets.space_closed positive_def nn_integral_nonneg)
 | 
|  |    263 | 
 | 
|  |    264 | lemma nn_integral_measurable_subprob_algebra:
 | 
|  |    265 |   assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
 | 
|  |    266 |   shows "(\<lambda>M. integral\<^sup>N M f) \<in> borel_measurable (subprob_algebra N)" (is "_ \<in> ?B")
 | 
|  |    267 |   using f
 | 
|  |    268 | proof induct
 | 
|  |    269 |   case (cong f g)
 | 
|  |    270 |   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"
 | 
|  |    271 |     by (intro measurable_cong nn_integral_cong cong)
 | 
|  |    272 |        (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq)
 | 
|  |    273 |   ultimately show ?case by simp
 | 
|  |    274 | next
 | 
|  |    275 |   case (set B)
 | 
|  |    276 |   moreover then have "(\<lambda>M'. \<integral>\<^sup>+M''. indicator B M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. emeasure M' B) \<in> ?B"
 | 
|  |    277 |     by (intro measurable_cong nn_integral_indicator) (simp add: space_subprob_algebra)
 | 
|  |    278 |   ultimately show ?case
 | 
|  |    279 |     by (simp add: measurable_emeasure_subprob_algebra)
 | 
|  |    280 | next
 | 
|  |    281 |   case (mult f c)
 | 
|  |    282 |   moreover 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"
 | 
|  |    283 |     by (intro measurable_cong nn_integral_cmult) (simp add: space_subprob_algebra)
 | 
|  |    284 |   ultimately show ?case
 | 
|  |    285 |     by simp
 | 
|  |    286 | next
 | 
|  |    287 |   case (add f g)
 | 
|  |    288 |   moreover 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"
 | 
|  |    289 |     by (intro measurable_cong nn_integral_add) (simp_all add: space_subprob_algebra)
 | 
|  |    290 |   ultimately show ?case
 | 
|  |    291 |     by (simp add: ac_simps)
 | 
|  |    292 | next
 | 
|  |    293 |   case (seq F)
 | 
|  |    294 |   moreover 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"
 | 
|  |    295 |     unfolding SUP_apply
 | 
|  |    296 |     by (intro measurable_cong nn_integral_monotone_convergence_SUP) (simp_all add: space_subprob_algebra)
 | 
|  |    297 |   ultimately show ?case
 | 
|  |    298 |     by (simp add: ac_simps)
 | 
|  |    299 | qed
 | 
|  |    300 | 
 | 
|  |    301 | 
 | 
|  |    302 | lemma measurable_join:
 | 
|  |    303 |   "join \<in> measurable (subprob_algebra (subprob_algebra N)) (subprob_algebra N)"
 | 
|  |    304 | proof (cases "space N \<noteq> {}", rule measurable_subprob_algebra)
 | 
|  |    305 |   fix A assume "A \<in> sets N"
 | 
|  |    306 |   let ?B = "borel_measurable (subprob_algebra (subprob_algebra N))"
 | 
|  |    307 |   have "(\<lambda>M'. emeasure (join M') A) \<in> ?B \<longleftrightarrow> (\<lambda>M'. (\<integral>\<^sup>+ M''. emeasure M'' A \<partial>M')) \<in> ?B"
 | 
|  |    308 |   proof (rule measurable_cong)
 | 
|  |    309 |     fix M' assume "M' \<in> space (subprob_algebra (subprob_algebra N))"
 | 
|  |    310 |     then show "emeasure (join M') A = (\<integral>\<^sup>+ M''. emeasure M'' A \<partial>M')"
 | 
|  |    311 |       by (intro emeasure_join) (auto simp: space_subprob_algebra `A\<in>sets N`)
 | 
|  |    312 |   qed
 | 
|  |    313 |   also have "(\<lambda>M'. \<integral>\<^sup>+M''. emeasure M'' A \<partial>M') \<in> ?B"
 | 
|  |    314 |     using measurable_emeasure_subprob_algebra[OF `A\<in>sets N`] emeasure_nonneg[of _ A]
 | 
|  |    315 |     by (rule nn_integral_measurable_subprob_algebra)
 | 
|  |    316 |   finally show "(\<lambda>M'. emeasure (join M') A) \<in> borel_measurable (subprob_algebra (subprob_algebra N))" .
 | 
|  |    317 | next
 | 
|  |    318 |   assume [simp]: "space N \<noteq> {}"
 | 
|  |    319 |   fix M assume M: "M \<in> space (subprob_algebra (subprob_algebra N))"
 | 
|  |    320 |   then have "(\<integral>\<^sup>+M'. emeasure M' (space N) \<partial>M) \<le> (\<integral>\<^sup>+M'. 1 \<partial>M)"
 | 
|  |    321 |     apply (intro nn_integral_mono)
 | 
|  |    322 |     apply (auto simp: space_subprob_algebra 
 | 
|  |    323 |                  dest!: sets_eq_imp_space_eq subprob_space.emeasure_space_le_1)
 | 
|  |    324 |     done
 | 
|  |    325 |   with M show "subprob_space (join M)"
 | 
|  |    326 |     by (intro subprob_spaceI)
 | 
|  |    327 |        (auto simp: emeasure_join space_subprob_algebra M assms dest: subprob_space.emeasure_space_le_1)
 | 
|  |    328 | next
 | 
|  |    329 |   assume "\<not>(space N \<noteq> {})"
 | 
|  |    330 |   thus ?thesis by (simp add: measurable_empty_iff space_subprob_algebra_empty_iff)
 | 
|  |    331 | qed (auto simp: space_subprob_algebra)
 | 
|  |    332 | 
 | 
|  |    333 | lemma nn_integral_join:
 | 
|  |    334 |   assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x" and M: "sets M = sets (subprob_algebra N)"
 | 
|  |    335 |   shows "(\<integral>\<^sup>+x. f x \<partial>join M) = (\<integral>\<^sup>+M'. \<integral>\<^sup>+x. f x \<partial>M' \<partial>M)"
 | 
|  |    336 |   using f
 | 
|  |    337 | proof induct
 | 
|  |    338 |   case (cong f g)
 | 
|  |    339 |   moreover have "integral\<^sup>N (join M) f = integral\<^sup>N (join M) g"
 | 
|  |    340 |     by (intro nn_integral_cong cong) (simp add: M)
 | 
|  |    341 |   moreover from M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' f \<partial>M) = (\<integral>\<^sup>+ M'. integral\<^sup>N M' g \<partial>M)"
 | 
|  |    342 |     by (intro nn_integral_cong cong)
 | 
|  |    343 |        (auto simp add: space_subprob_algebra dest!: sets_eq_imp_space_eq)
 | 
|  |    344 |   ultimately show ?case
 | 
|  |    345 |     by simp
 | 
|  |    346 | next
 | 
|  |    347 |   case (set A)
 | 
|  |    348 |   moreover with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)" 
 | 
|  |    349 |     by (intro nn_integral_cong nn_integral_indicator)
 | 
|  |    350 |        (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq)
 | 
|  |    351 |   ultimately show ?case
 | 
|  |    352 |     using M by (simp add: emeasure_join)
 | 
|  |    353 | next
 | 
|  |    354 |   case (mult f c)
 | 
|  |    355 |   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)"
 | 
|  |    356 |     using mult M
 | 
|  |    357 |     by (intro nn_integral_cong nn_integral_cmult)
 | 
|  |    358 |        (auto simp add: space_subprob_algebra cong: measurable_cong dest!: sets_eq_imp_space_eq)
 | 
|  |    359 |   also have "\<dots> = c * (\<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. f x \<partial>M' \<partial>M)"
 | 
|  |    360 |     using nn_integral_measurable_subprob_algebra[OF mult(3,4)]
 | 
|  |    361 |     by (intro nn_integral_cmult mult) (simp add: M)
 | 
|  |    362 |   also have "\<dots> = c * (integral\<^sup>N (join M) f)"
 | 
|  |    363 |     by (simp add: mult)
 | 
|  |    364 |   also have "\<dots> = (\<integral>\<^sup>+ x. c * f x \<partial>join M)"
 | 
|  |    365 |     using mult(2,3) by (intro nn_integral_cmult[symmetric] mult) (simp add: M)
 | 
|  |    366 |   finally show ?case by simp
 | 
|  |    367 | next
 | 
|  |    368 |   case (add f g)
 | 
|  |    369 |   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)"
 | 
|  |    370 |     using add M
 | 
|  |    371 |     by (intro nn_integral_cong nn_integral_add)
 | 
|  |    372 |        (auto simp add: space_subprob_algebra cong: measurable_cong dest!: sets_eq_imp_space_eq)
 | 
|  |    373 |   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)"
 | 
|  |    374 |     using nn_integral_measurable_subprob_algebra[OF add(1,2)]
 | 
|  |    375 |     using nn_integral_measurable_subprob_algebra[OF add(5,6)]
 | 
|  |    376 |     by (intro nn_integral_add add) (simp_all add: M nn_integral_nonneg)
 | 
|  |    377 |   also have "\<dots> = (integral\<^sup>N (join M) f) + (integral\<^sup>N (join M) g)"
 | 
|  |    378 |     by (simp add: add)
 | 
|  |    379 |   also have "\<dots> = (\<integral>\<^sup>+ x. f x + g x \<partial>join M)"
 | 
|  |    380 |     using add by (intro nn_integral_add[symmetric] add) (simp_all add: M)
 | 
|  |    381 |   finally show ?case by (simp add: ac_simps)
 | 
|  |    382 | next
 | 
|  |    383 |   case (seq F)
 | 
|  |    384 |   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)"
 | 
|  |    385 |     using seq M unfolding SUP_apply
 | 
|  |    386 |     by (intro nn_integral_cong nn_integral_monotone_convergence_SUP)
 | 
|  |    387 |        (auto simp add: space_subprob_algebra cong: measurable_cong dest!: sets_eq_imp_space_eq)
 | 
|  |    388 |   also have "\<dots> = (SUP i. \<integral>\<^sup>+ M'. \<integral>\<^sup>+ x. F i x \<partial>M' \<partial>M)"
 | 
|  |    389 |     using nn_integral_measurable_subprob_algebra[OF seq(1,2)] seq
 | 
|  |    390 |     by (intro nn_integral_monotone_convergence_SUP)
 | 
|  |    391 |        (simp_all add: M nn_integral_nonneg incseq_nn_integral incseq_def le_fun_def nn_integral_mono )
 | 
|  |    392 |   also have "\<dots> = (SUP i. integral\<^sup>N (join M) (F i))"
 | 
|  |    393 |     by (simp add: seq)
 | 
|  |    394 |   also have "\<dots> = (\<integral>\<^sup>+ x. (SUP i. F i x) \<partial>join M)"
 | 
|  |    395 |     using seq by (intro nn_integral_monotone_convergence_SUP[symmetric] seq) (simp_all add: M)
 | 
|  |    396 |   finally show ?case by (simp add: ac_simps)
 | 
|  |    397 | qed
 | 
|  |    398 | 
 | 
|  |    399 | lemma join_assoc:
 | 
|  |    400 |   assumes M: "sets M = sets (subprob_algebra (subprob_algebra N))"
 | 
|  |    401 |   shows "join (distr M (subprob_algebra N) join) = join (join M)"
 | 
|  |    402 | proof (rule measure_eqI)
 | 
|  |    403 |   fix A assume "A \<in> sets (join (distr M (subprob_algebra N) join))"
 | 
|  |    404 |   then have A: "A \<in> sets N" by simp
 | 
|  |    405 |   show "emeasure (join (distr M (subprob_algebra N) join)) A = emeasure (join (join M)) A"
 | 
|  |    406 |     using measurable_join[of N]
 | 
|  |    407 |     by (auto simp: M A nn_integral_distr emeasure_join measurable_emeasure_subprob_algebra emeasure_nonneg
 | 
|  |    408 |                    sets_eq_imp_space_eq[OF M] space_subprob_algebra nn_integral_join[OF _ _ M]
 | 
|  |    409 |              intro!: nn_integral_cong emeasure_join cong: measurable_cong)
 | 
|  |    410 | qed (simp add: M)
 | 
|  |    411 | 
 | 
|  |    412 | lemma join_return: 
 | 
|  |    413 |   assumes "sets M = sets N" and "subprob_space M"
 | 
|  |    414 |   shows "join (return (subprob_algebra N) M) = M"
 | 
|  |    415 |   by (rule measure_eqI)
 | 
|  |    416 |      (simp_all add: emeasure_join emeasure_nonneg space_subprob_algebra  
 | 
|  |    417 |                     measurable_emeasure_subprob_algebra nn_integral_return assms)
 | 
|  |    418 | 
 | 
|  |    419 | lemma join_return':
 | 
|  |    420 |   assumes "sets N = sets M"
 | 
|  |    421 |   shows "join (distr M (subprob_algebra N) (return N)) = M"
 | 
|  |    422 | apply (rule measure_eqI)
 | 
|  |    423 | apply (simp add: assms)
 | 
|  |    424 | apply (subgoal_tac "return N \<in> measurable M (subprob_algebra N)")
 | 
|  |    425 | apply (simp add: emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra assms)
 | 
|  |    426 | apply (subst measurable_cong_sets, rule assms[symmetric], rule refl, rule return_measurable)
 | 
|  |    427 | done
 | 
|  |    428 | 
 | 
|  |    429 | lemma join_distr_distr:
 | 
|  |    430 |   fixes f :: "'a \<Rightarrow> 'b" and M :: "'a measure measure" and N :: "'b measure"
 | 
|  |    431 |   assumes "sets M = sets (subprob_algebra R)" and "f \<in> measurable R N"
 | 
|  |    432 |   shows "join (distr M (subprob_algebra N) (\<lambda>M. distr M N f)) = distr (join M) N f" (is "?r = ?l")
 | 
|  |    433 | proof (rule measure_eqI)
 | 
|  |    434 |   fix A assume "A \<in> sets ?r"
 | 
|  |    435 |   hence A_in_N: "A \<in> sets N" by simp
 | 
|  |    436 | 
 | 
|  |    437 |   from assms have "f \<in> measurable (join M) N" 
 | 
|  |    438 |       by (simp cong: measurable_cong_sets)
 | 
|  |    439 |   moreover from assms and A_in_N have "f-`A \<inter> space R \<in> sets R" 
 | 
|  |    440 |       by (intro measurable_sets) simp_all
 | 
|  |    441 |   ultimately have "emeasure (distr (join M) N f) A = \<integral>\<^sup>+M'. emeasure M' (f-`A \<inter> space R) \<partial>M"
 | 
|  |    442 |       by (simp_all add: A_in_N emeasure_distr emeasure_join assms)
 | 
|  |    443 |   
 | 
|  |    444 |   also have "... = \<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M" using A_in_N
 | 
|  |    445 |   proof (intro nn_integral_cong, subst emeasure_distr)
 | 
|  |    446 |     fix M' assume "M' \<in> space M"
 | 
|  |    447 |     from assms have "space M = space (subprob_algebra R)"
 | 
|  |    448 |         using sets_eq_imp_space_eq by blast
 | 
|  |    449 |     with `M' \<in> space M` have [simp]: "sets M' = sets R" using space_subprob_algebra by blast
 | 
|  |    450 |     show "f \<in> measurable M' N" by (simp cong: measurable_cong_sets add: assms)
 | 
|  |    451 |     have "space M' = space R" by (rule sets_eq_imp_space_eq) simp
 | 
|  |    452 |     thus "emeasure M' (f -` A \<inter> space R) = emeasure M' (f -` A \<inter> space M')" by simp
 | 
|  |    453 |   qed
 | 
|  |    454 | 
 | 
|  |    455 |   also have "(\<lambda>M. distr M N f) \<in> measurable M (subprob_algebra N)"
 | 
|  |    456 |       by (simp cong: measurable_cong_sets add: assms measurable_distr)
 | 
|  |    457 |   hence "(\<integral>\<^sup>+ x. emeasure (distr x N f) A \<partial>M) = 
 | 
|  |    458 |              emeasure (join (distr M (subprob_algebra N) (\<lambda>M. distr M N f))) A"
 | 
|  |    459 |       by (simp_all add: emeasure_join assms A_in_N nn_integral_distr measurable_emeasure_subprob_algebra)
 | 
|  |    460 |   finally show "emeasure ?r A = emeasure ?l A" ..
 | 
|  |    461 | qed simp
 | 
|  |    462 | 
 | 
|  |    463 | definition bind :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b measure) \<Rightarrow> 'b measure" where
 | 
|  |    464 |   "bind M f = (if space M = {} then count_space {} else
 | 
|  |    465 |     join (distr M (subprob_algebra (f (SOME x. x \<in> space M))) f))"
 | 
|  |    466 | 
 | 
|  |    467 | adhoc_overloading Monad_Syntax.bind bind
 | 
|  |    468 | 
 | 
|  |    469 | lemma bind_empty: 
 | 
|  |    470 |   "space M = {} \<Longrightarrow> bind M f = count_space {}"
 | 
|  |    471 |   by (simp add: bind_def)
 | 
|  |    472 | 
 | 
|  |    473 | lemma bind_nonempty:
 | 
|  |    474 |   "space M \<noteq> {} \<Longrightarrow> bind M f = join (distr M (subprob_algebra (f (SOME x. x \<in> space M))) f)"
 | 
|  |    475 |   by (simp add: bind_def)
 | 
|  |    476 | 
 | 
|  |    477 | lemma sets_bind_empty: "sets M = {} \<Longrightarrow> sets (bind M f) = {{}}"
 | 
|  |    478 |   by (auto simp: bind_def)
 | 
|  |    479 | 
 | 
|  |    480 | lemma space_bind_empty: "space M = {} \<Longrightarrow> space (bind M f) = {}"
 | 
|  |    481 |   by (simp add: bind_def)
 | 
|  |    482 | 
 | 
|  |    483 | lemma sets_bind[simp]: 
 | 
|  |    484 |   assumes "f \<in> measurable M (subprob_algebra N)" and "space M \<noteq> {}"
 | 
|  |    485 |   shows "sets (bind M f) = sets N"
 | 
|  |    486 |     using assms(2) by (force simp: bind_nonempty intro!: sets_kernel[OF assms(1) someI_ex])
 | 
|  |    487 | 
 | 
|  |    488 | lemma space_bind[simp]: 
 | 
|  |    489 |   assumes "f \<in> measurable M (subprob_algebra N)" and "space M \<noteq> {}"
 | 
|  |    490 |   shows "space (bind M f) = space N"
 | 
|  |    491 |     using assms by (intro sets_eq_imp_space_eq sets_bind)
 | 
|  |    492 | 
 | 
|  |    493 | lemma bind_cong: 
 | 
|  |    494 |   assumes "\<forall>x \<in> space M. f x = g x"
 | 
|  |    495 |   shows "bind M f = bind M g"
 | 
|  |    496 | proof (cases "space M = {}")
 | 
|  |    497 |   assume "space M \<noteq> {}"
 | 
|  |    498 |   hence "(SOME x. x \<in> space M) \<in> space M" by (rule_tac someI_ex) blast
 | 
|  |    499 |   with assms have "f (SOME x. x \<in> space M) = g (SOME x. x \<in> space M)" by blast
 | 
|  |    500 |   with `space M \<noteq> {}` and assms show ?thesis by (simp add: bind_nonempty cong: distr_cong)
 | 
|  |    501 | qed (simp add: bind_empty)
 | 
|  |    502 | 
 | 
|  |    503 | lemma bind_nonempty':
 | 
|  |    504 |   assumes "f \<in> measurable M (subprob_algebra N)" "x \<in> space M"
 | 
|  |    505 |   shows "bind M f = join (distr M (subprob_algebra N) f)"
 | 
|  |    506 |   using assms
 | 
|  |    507 |   apply (subst bind_nonempty, blast)
 | 
|  |    508 |   apply (subst subprob_algebra_cong[OF sets_kernel[OF assms(1) someI_ex]], blast)
 | 
|  |    509 |   apply (simp add: subprob_algebra_cong[OF sets_kernel[OF assms]])
 | 
|  |    510 |   done
 | 
|  |    511 | 
 | 
|  |    512 | lemma bind_nonempty'':
 | 
|  |    513 |   assumes "f \<in> measurable M (subprob_algebra N)" "space M \<noteq> {}"
 | 
|  |    514 |   shows "bind M f = join (distr M (subprob_algebra N) f)"
 | 
|  |    515 |   using assms by (auto intro: bind_nonempty')
 | 
|  |    516 | 
 | 
|  |    517 | lemma emeasure_bind:
 | 
|  |    518 |     "\<lbrakk>space M \<noteq> {}; f \<in> measurable M (subprob_algebra N);X \<in> sets N\<rbrakk>
 | 
|  |    519 |       \<Longrightarrow> emeasure (M \<guillemotright>= f) X = \<integral>\<^sup>+x. emeasure (f x) X \<partial>M"
 | 
|  |    520 |   by (simp add: bind_nonempty'' emeasure_join nn_integral_distr measurable_emeasure_subprob_algebra)
 | 
|  |    521 | 
 | 
|  |    522 | lemma bind_return: 
 | 
|  |    523 |   assumes "f \<in> measurable M (subprob_algebra N)" and "x \<in> space M"
 | 
|  |    524 |   shows "bind (return M x) f = f x"
 | 
|  |    525 |   using sets_kernel[OF assms] assms
 | 
|  |    526 |   by (simp_all add: distr_return join_return subprob_space_kernel bind_nonempty'
 | 
|  |    527 |                cong: subprob_algebra_cong)
 | 
|  |    528 | 
 | 
|  |    529 | lemma bind_return':
 | 
|  |    530 |   shows "bind M (return M) = M"
 | 
|  |    531 |   by (cases "space M = {}")
 | 
|  |    532 |      (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return' 
 | 
|  |    533 |                cong: subprob_algebra_cong)
 | 
|  |    534 | 
 | 
|  |    535 | lemma bind_count_space_singleton:
 | 
|  |    536 |   assumes "subprob_space (f x)"
 | 
|  |    537 |   shows "count_space {x} \<guillemotright>= f = f x"
 | 
|  |    538 | proof-
 | 
|  |    539 |   have A: "\<And>A. A \<subseteq> {x} \<Longrightarrow> A = {} \<or> A = {x}" by auto
 | 
|  |    540 |   have "count_space {x} = return (count_space {x}) x"
 | 
|  |    541 |     by (intro measure_eqI) (auto dest: A)
 | 
|  |    542 |   also have "... \<guillemotright>= f = f x"
 | 
|  |    543 |     by (subst bind_return[of _ _ "f x"]) (auto simp: space_subprob_algebra assms)
 | 
|  |    544 |   finally show ?thesis .
 | 
|  |    545 | qed
 | 
|  |    546 | 
 | 
|  |    547 | lemma emeasure_bind_const: 
 | 
|  |    548 |     "space M \<noteq> {} \<Longrightarrow> X \<in> sets N \<Longrightarrow> subprob_space N \<Longrightarrow> 
 | 
|  |    549 |          emeasure (M \<guillemotright>= (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)"
 | 
|  |    550 |   by (simp add: bind_nonempty emeasure_join nn_integral_distr 
 | 
|  |    551 |                 space_subprob_algebra measurable_emeasure_subprob_algebra emeasure_nonneg)
 | 
|  |    552 | 
 | 
|  |    553 | lemma emeasure_bind_const':
 | 
|  |    554 |   assumes "subprob_space M" "subprob_space N"
 | 
|  |    555 |   shows "emeasure (M \<guillemotright>= (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)"
 | 
|  |    556 | using assms
 | 
|  |    557 | proof (case_tac "X \<in> sets N")
 | 
|  |    558 |   fix X assume "X \<in> sets N"
 | 
|  |    559 |   thus "emeasure (M \<guillemotright>= (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)" using assms
 | 
|  |    560 |       by (subst emeasure_bind_const) 
 | 
|  |    561 |          (simp_all add: subprob_space.subprob_not_empty subprob_space.emeasure_space_le_1)
 | 
|  |    562 | next
 | 
|  |    563 |   fix X assume "X \<notin> sets N"
 | 
|  |    564 |   with assms show "emeasure (M \<guillemotright>= (\<lambda>x. N)) X = emeasure N X * emeasure M (space M)"
 | 
|  |    565 |       by (simp add: sets_bind[of _ _ N] subprob_space.subprob_not_empty
 | 
|  |    566 |                     space_subprob_algebra emeasure_notin_sets)
 | 
|  |    567 | qed
 | 
|  |    568 | 
 | 
|  |    569 | lemma emeasure_bind_const_prob_space:
 | 
|  |    570 |   assumes "prob_space M" "subprob_space N"
 | 
|  |    571 |   shows "emeasure (M \<guillemotright>= (\<lambda>x. N)) X = emeasure N X"
 | 
|  |    572 |   using assms by (simp add: emeasure_bind_const' prob_space_imp_subprob_space 
 | 
|  |    573 |                             prob_space.emeasure_space_1)
 | 
|  |    574 | 
 | 
|  |    575 | lemma bind_const': "\<lbrakk>prob_space M; subprob_space N\<rbrakk> \<Longrightarrow> M \<guillemotright>= (\<lambda>x. N) = N"
 | 
|  |    576 |   by (intro measure_eqI) 
 | 
|  |    577 |      (simp_all add: space_subprob_algebra prob_space.not_empty emeasure_bind_const_prob_space)
 | 
|  |    578 | 
 | 
|  |    579 | lemma bind_return_distr: 
 | 
|  |    580 |     "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (return N \<circ> f) = distr M N f"
 | 
|  |    581 |   apply (simp add: bind_nonempty)
 | 
|  |    582 |   apply (subst subprob_algebra_cong)
 | 
|  |    583 |   apply (rule sets_return)
 | 
|  |    584 |   apply (subst distr_distr[symmetric])
 | 
|  |    585 |   apply (auto intro!: return_measurable simp: distr_distr[symmetric] join_return')
 | 
|  |    586 |   done
 | 
|  |    587 | 
 | 
|  |    588 | lemma bind_assoc:
 | 
|  |    589 |   fixes f :: "'a \<Rightarrow> 'b measure" and g :: "'b \<Rightarrow> 'c measure"
 | 
|  |    590 |   assumes M1: "f \<in> measurable M (subprob_algebra N)" and M2: "g \<in> measurable N (subprob_algebra R)"
 | 
|  |    591 |   shows "bind (bind M f) g = bind M (\<lambda>x. bind (f x) g)"
 | 
|  |    592 | proof (cases "space M = {}")
 | 
|  |    593 |   assume [simp]: "space M \<noteq> {}"
 | 
|  |    594 |   from assms have [simp]: "space N \<noteq> {}" "space R \<noteq> {}"
 | 
|  |    595 |       by (auto simp: measurable_empty_iff space_subprob_algebra_empty_iff)
 | 
|  |    596 |   from assms have sets_fx[simp]: "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N"
 | 
|  |    597 |       by (simp add: sets_kernel)
 | 
|  |    598 |   have ex_in: "\<And>A. A \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A" by blast
 | 
|  |    599 |   note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF `space M \<noteq> {}`]]]
 | 
|  |    600 |                          sets_kernel[OF M2 someI_ex[OF ex_in[OF `space N \<noteq> {}`]]]
 | 
|  |    601 |   note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)]
 | 
|  |    602 | 
 | 
|  |    603 |   have "bind M (\<lambda>x. bind (f x) g) = 
 | 
|  |    604 |         join (distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f))"
 | 
|  |    605 |     by (simp add: sets_eq_imp_space_eq[OF sets_fx] bind_nonempty o_def
 | 
|  |    606 |              cong: subprob_algebra_cong distr_cong)
 | 
|  |    607 |   also have "distr M (subprob_algebra R) (join \<circ> (\<lambda>x. (distr x (subprob_algebra R) g)) \<circ> f) =
 | 
|  |    608 |              distr (distr (distr M (subprob_algebra N) f)
 | 
|  |    609 |                           (subprob_algebra (subprob_algebra R))
 | 
|  |    610 |                           (\<lambda>x. distr x (subprob_algebra R) g)) 
 | 
|  |    611 |                    (subprob_algebra R) join"
 | 
|  |    612 |       apply (subst distr_distr, 
 | 
|  |    613 |              (blast intro: measurable_comp measurable_distr measurable_join M1 M2)+)+
 | 
|  |    614 |       apply (simp add: o_assoc)
 | 
|  |    615 |       done
 | 
|  |    616 |   also have "join ... = bind (bind M f) g"
 | 
|  |    617 |       by (simp add: join_assoc join_distr_distr M2 bind_nonempty cong: subprob_algebra_cong)
 | 
|  |    618 |   finally show ?thesis ..
 | 
|  |    619 | qed (simp add: bind_empty)
 | 
|  |    620 | 
 | 
|  |    621 | lemma emeasure_space_subprob_algebra[measurable]:
 | 
|  |    622 |   "(\<lambda>a. emeasure a (space a)) \<in> borel_measurable (subprob_algebra N)"
 | 
|  |    623 | proof-
 | 
|  |    624 |   have "(\<lambda>a. emeasure a (space N)) \<in> borel_measurable (subprob_algebra N)" (is "?f \<in> ?M")
 | 
|  |    625 |     by (rule measurable_emeasure_subprob_algebra) simp
 | 
|  |    626 |   also have "?f \<in> ?M \<longleftrightarrow> (\<lambda>a. emeasure a (space a)) \<in> ?M"
 | 
|  |    627 |     by (rule measurable_cong) (auto simp: space_subprob_algebra dest: sets_eq_imp_space_eq)
 | 
|  |    628 |   finally show ?thesis .
 | 
|  |    629 | qed
 | 
|  |    630 | 
 | 
|  |    631 | (* TODO: Rename. This name is too general – Manuel *)
 | 
|  |    632 | lemma measurable_pair_measure:
 | 
|  |    633 |   assumes f: "f \<in> measurable M (subprob_algebra N)"
 | 
|  |    634 |   assumes g: "g \<in> measurable M (subprob_algebra L)"
 | 
|  |    635 |   shows "(\<lambda>x. f x \<Otimes>\<^sub>M g x) \<in> measurable M (subprob_algebra (N \<Otimes>\<^sub>M L))"
 | 
|  |    636 | proof (rule measurable_subprob_algebra)
 | 
|  |    637 |   { fix x assume "x \<in> space M"
 | 
|  |    638 |     with measurable_space[OF f] measurable_space[OF g]
 | 
|  |    639 |     have fx: "f x \<in> space (subprob_algebra N)" and gx: "g x \<in> space (subprob_algebra L)"
 | 
|  |    640 |       by auto
 | 
|  |    641 |     interpret F: subprob_space "f x"
 | 
|  |    642 |       using fx by (simp add: space_subprob_algebra)
 | 
|  |    643 |     interpret G: subprob_space "g x"
 | 
|  |    644 |       using gx by (simp add: space_subprob_algebra)
 | 
|  |    645 | 
 | 
|  |    646 |     interpret pair_subprob_space "f x" "g x" ..
 | 
|  |    647 |     show "subprob_space (f x \<Otimes>\<^sub>M g x)" by unfold_locales
 | 
|  |    648 |     show sets_eq: "sets (f x \<Otimes>\<^sub>M g x) = sets (N \<Otimes>\<^sub>M L)"
 | 
|  |    649 |       using fx gx by (simp add: space_subprob_algebra)
 | 
|  |    650 | 
 | 
|  |    651 |     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"
 | 
|  |    652 |       using fx gx by (intro G.emeasure_pair_measure_Times) (auto simp: space_subprob_algebra) 
 | 
|  |    653 |     have "emeasure (f x \<Otimes>\<^sub>M g x) (space (f x \<Otimes>\<^sub>M g x)) = 
 | 
|  |    654 |               emeasure (f x) (space (f x)) * emeasure (g x) (space (g x))"
 | 
|  |    655 |       by (subst G.emeasure_pair_measure_Times[symmetric]) (simp_all add: space_pair_measure)
 | 
|  |    656 |     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) =
 | 
|  |    657 |                                              ... - emeasure (f x \<Otimes>\<^sub>M g x) A"
 | 
|  |    658 |       using emeasure_compl[OF _ P.emeasure_finite]
 | 
|  |    659 |       unfolding sets_eq
 | 
|  |    660 |       unfolding sets_eq_imp_space_eq[OF sets_eq]
 | 
|  |    661 |       by (simp add: space_pair_measure G.emeasure_pair_measure_Times)
 | 
|  |    662 |     note 1 2 sets_eq }
 | 
|  |    663 |   note Times = this(1) and Compl = this(2) and sets_eq = this(3)
 | 
|  |    664 | 
 | 
|  |    665 |   fix A assume A: "A \<in> sets (N \<Otimes>\<^sub>M L)"
 | 
|  |    666 |   show "(\<lambda>a. emeasure (f a \<Otimes>\<^sub>M g a) A) \<in> borel_measurable M"
 | 
|  |    667 |     using Int_stable_pair_measure_generator pair_measure_closed A
 | 
|  |    668 |     unfolding sets_pair_measure
 | 
|  |    669 |   proof (induct A rule: sigma_sets_induct_disjoint)
 | 
|  |    670 |     case (basic A) then show ?case
 | 
|  |    671 |       by (auto intro!: borel_measurable_ereal_times simp: Times cong: measurable_cong)
 | 
|  |    672 |          (auto intro!: measurable_emeasure_kernel f g)
 | 
|  |    673 |   next
 | 
|  |    674 |     case (compl A)
 | 
|  |    675 |     then have A: "A \<in> sets (N \<Otimes>\<^sub>M L)"
 | 
|  |    676 |       by (auto simp: sets_pair_measure)
 | 
|  |    677 |     have "(\<lambda>x. emeasure (f x) (space (f x)) * emeasure (g x) (space (g x)) - 
 | 
|  |    678 |                    emeasure (f x \<Otimes>\<^sub>M g x) A) \<in> borel_measurable M" (is "?f \<in> ?M")
 | 
|  |    679 |       using compl(2) f g by measurable
 | 
|  |    680 |     thus ?case by (simp add: Compl A cong: measurable_cong)
 | 
|  |    681 |   next
 | 
|  |    682 |     case (union A)
 | 
|  |    683 |     then have "range A \<subseteq> sets (N \<Otimes>\<^sub>M L)" "disjoint_family A"
 | 
|  |    684 |       by (auto simp: sets_pair_measure)
 | 
|  |    685 |     then have "(\<lambda>a. emeasure (f a \<Otimes>\<^sub>M g a) (\<Union>i. A i)) \<in> borel_measurable M \<longleftrightarrow>
 | 
|  |    686 |       (\<lambda>a. \<Sum>i. emeasure (f a \<Otimes>\<^sub>M g a) (A i)) \<in> borel_measurable M"
 | 
|  |    687 |       by (intro measurable_cong suminf_emeasure[symmetric])
 | 
|  |    688 |          (auto simp: sets_eq)
 | 
|  |    689 |     also have "\<dots>"
 | 
|  |    690 |       using union by auto
 | 
|  |    691 |     finally show ?case .
 | 
|  |    692 |   qed simp
 | 
|  |    693 | qed
 | 
|  |    694 | 
 | 
|  |    695 | (* TODO: Move *)
 | 
|  |    696 | lemma measurable_distr2:
 | 
|  |    697 |   assumes f[measurable]: "split f \<in> measurable (L \<Otimes>\<^sub>M M) N"
 | 
|  |    698 |   assumes g[measurable]: "g \<in> measurable L (subprob_algebra M)"
 | 
|  |    699 |   shows "(\<lambda>x. distr (g x) N (f x)) \<in> measurable L (subprob_algebra N)"
 | 
|  |    700 | proof -
 | 
|  |    701 |   have "(\<lambda>x. distr (g x) N (f x)) \<in> measurable L (subprob_algebra N)
 | 
|  |    702 |     \<longleftrightarrow> (\<lambda>x. distr (return L x \<Otimes>\<^sub>M g x) N (split f)) \<in> measurable L (subprob_algebra N)"
 | 
|  |    703 |   proof (rule measurable_cong)
 | 
|  |    704 |     fix x assume x: "x \<in> space L"
 | 
|  |    705 |     have gx: "g x \<in> space (subprob_algebra M)"
 | 
|  |    706 |       using measurable_space[OF g x] .
 | 
|  |    707 |     then have [simp]: "sets (g x) = sets M"
 | 
|  |    708 |       by (simp add: space_subprob_algebra)
 | 
|  |    709 |     then have [simp]: "space (g x) = space M"
 | 
|  |    710 |       by (rule sets_eq_imp_space_eq)
 | 
|  |    711 |     let ?R = "return L x"
 | 
|  |    712 |     from measurable_compose_Pair1[OF x f] have f_M': "f x \<in> measurable M N"
 | 
|  |    713 |       by simp
 | 
|  |    714 |     interpret subprob_space "g x"
 | 
|  |    715 |       using gx by (simp add: space_subprob_algebra)
 | 
|  |    716 |     have space_pair_M'[simp]: "\<And>X. space (X \<Otimes>\<^sub>M g x) = space (X \<Otimes>\<^sub>M M)"
 | 
|  |    717 |       by (simp add: space_pair_measure)
 | 
|  |    718 |     show "distr (g x) N (f x) = distr (?R \<Otimes>\<^sub>M g x) N (split f)" (is "?l = ?r")
 | 
|  |    719 |     proof (rule measure_eqI)
 | 
|  |    720 |       show "sets ?l = sets ?r"
 | 
|  |    721 |         by simp
 | 
|  |    722 |     next
 | 
|  |    723 |       fix A assume "A \<in> sets ?l"
 | 
|  |    724 |       then have A[measurable]: "A \<in> sets N"
 | 
|  |    725 |         by simp
 | 
|  |    726 |       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))"
 | 
|  |    727 |         by (auto simp add: emeasure_distr f_M' cong: measurable_cong_sets)
 | 
|  |    728 |       also have "\<dots> = (\<integral>\<^sup>+M''. emeasure (g x) (f M'' -` A \<inter> space M) \<partial>?R)"
 | 
|  |    729 |         apply (subst emeasure_pair_measure_alt)
 | 
|  |    730 |         apply (rule measurable_sets[OF _ A])
 | 
|  |    731 |         apply (auto simp add: f_M' cong: measurable_cong_sets)
 | 
|  |    732 |         apply (intro nn_integral_cong arg_cong[where f="emeasure (g x)"])
 | 
|  |    733 |         apply (auto simp: space_subprob_algebra space_pair_measure)
 | 
|  |    734 |         done
 | 
|  |    735 |       also have "\<dots> = emeasure (g x) (f x -` A \<inter> space M)"
 | 
|  |    736 |         by (subst nn_integral_return)
 | 
|  |    737 |            (auto simp: x intro!: measurable_emeasure)
 | 
|  |    738 |       also have "\<dots> = emeasure ?l A"
 | 
|  |    739 |         by (simp add: emeasure_distr f_M' cong: measurable_cong_sets)
 | 
|  |    740 |       finally show "emeasure ?l A = emeasure ?r A" ..
 | 
|  |    741 |     qed
 | 
|  |    742 |   qed
 | 
|  |    743 |   also have "\<dots>"
 | 
|  |    744 |     apply (intro measurable_compose[OF measurable_pair_measure measurable_distr])
 | 
|  |    745 |     apply (rule return_measurable)
 | 
|  |    746 |     apply measurable
 | 
|  |    747 |     done
 | 
|  |    748 |   finally show ?thesis .
 | 
|  |    749 | qed
 | 
|  |    750 | 
 | 
|  |    751 | (* END TODO *)
 | 
|  |    752 | 
 | 
|  |    753 | lemma measurable_bind':
 | 
|  |    754 |   assumes M1: "f \<in> measurable M (subprob_algebra N)" and
 | 
|  |    755 |           M2: "split g \<in> measurable (M \<Otimes>\<^sub>M N) (subprob_algebra R)"
 | 
|  |    756 |   shows "(\<lambda>x. bind (f x) (g x)) \<in> measurable M (subprob_algebra R)"
 | 
|  |    757 | proof (subst measurable_cong)
 | 
|  |    758 |   fix x assume x_in_M: "x \<in> space M"
 | 
|  |    759 |   with assms have "space (f x) \<noteq> {}" 
 | 
|  |    760 |       by (blast dest: subprob_space_kernel subprob_space.subprob_not_empty)
 | 
|  |    761 |   moreover from M2 x_in_M have "g x \<in> measurable (f x) (subprob_algebra R)"
 | 
|  |    762 |       by (subst measurable_cong_sets[OF sets_kernel[OF M1 x_in_M] refl])
 | 
|  |    763 |          (auto dest: measurable_Pair2)
 | 
|  |    764 |   ultimately show "bind (f x) (g x) = join (distr (f x) (subprob_algebra R) (g x))" 
 | 
|  |    765 |       by (simp_all add: bind_nonempty'')
 | 
|  |    766 | next
 | 
|  |    767 |   show "(\<lambda>w. join (distr (f w) (subprob_algebra R) (g w))) \<in> measurable M (subprob_algebra R)"
 | 
|  |    768 |     apply (rule measurable_compose[OF _ measurable_join])
 | 
|  |    769 |     apply (rule measurable_distr2[OF M2 M1])
 | 
|  |    770 |     done
 | 
|  |    771 | qed
 | 
|  |    772 | 
 | 
|  |    773 | lemma measurable_bind:
 | 
|  |    774 |   assumes M1: "f \<in> measurable M (subprob_algebra N)" and
 | 
|  |    775 |           M2: "(\<lambda>x. g (fst x) (snd x)) \<in> measurable (M \<Otimes>\<^sub>M N) (subprob_algebra R)"
 | 
|  |    776 |   shows "(\<lambda>x. bind (f x) (g x)) \<in> measurable M (subprob_algebra R)"
 | 
|  |    777 |   using assms by (auto intro: measurable_bind' simp: measurable_split_conv)
 | 
|  |    778 | 
 | 
|  |    779 | lemma measurable_bind2:
 | 
|  |    780 |   assumes "f \<in> measurable M (subprob_algebra N)" and "g \<in> measurable N (subprob_algebra R)"
 | 
|  |    781 |   shows "(\<lambda>x. bind (f x) g) \<in> measurable M (subprob_algebra R)"
 | 
|  |    782 |     using assms by (intro measurable_bind' measurable_const) auto
 | 
|  |    783 | 
 | 
|  |    784 | lemma subprob_space_bind:
 | 
|  |    785 |   assumes "subprob_space M" "f \<in> measurable M (subprob_algebra N)"
 | 
|  |    786 |   shows "subprob_space (M \<guillemotright>= f)"
 | 
|  |    787 | proof (rule subprob_space_kernel[of "\<lambda>x. x \<guillemotright>= f"])
 | 
|  |    788 |   show "(\<lambda>x. x \<guillemotright>= f) \<in> measurable (subprob_algebra M) (subprob_algebra N)"
 | 
|  |    789 |     by (rule measurable_bind, rule measurable_ident_sets, rule refl, 
 | 
|  |    790 |         rule measurable_compose[OF measurable_snd assms(2)])
 | 
|  |    791 |   from assms(1) show "M \<in> space (subprob_algebra M)"
 | 
|  |    792 |     by (simp add: space_subprob_algebra)
 | 
|  |    793 | qed
 | 
|  |    794 | 
 | 
|  |    795 | lemma double_bind_assoc:
 | 
|  |    796 |   assumes Mg: "g \<in> measurable N (subprob_algebra N')"
 | 
|  |    797 |   assumes Mf: "f \<in> measurable M (subprob_algebra M')"
 | 
|  |    798 |   assumes Mh: "split h \<in> measurable (M \<Otimes>\<^sub>M M') N"
 | 
|  |    799 |   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)} \<guillemotright>= g"
 | 
|  |    800 | proof-
 | 
|  |    801 |   have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y)} \<guillemotright>= g = 
 | 
|  |    802 |             do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<guillemotright>= g}"
 | 
|  |    803 |     using Mh by (auto intro!: bind_assoc measurable_bind'[OF Mf] Mf Mg
 | 
|  |    804 |                       measurable_compose[OF _ return_measurable] simp: measurable_split_conv)
 | 
|  |    805 |   also from Mh have "\<And>x. x \<in> space M \<Longrightarrow> h x \<in> measurable M' N" by measurable
 | 
|  |    806 |   hence "do {x \<leftarrow> M; do {y \<leftarrow> f x; return N (h x y)} \<guillemotright>= g} = 
 | 
|  |    807 |             do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<guillemotright>= g}"
 | 
|  |    808 |     apply (intro ballI bind_cong bind_assoc)
 | 
|  |    809 |     apply (subst measurable_cong_sets[OF sets_kernel[OF Mf] refl], simp)
 | 
|  |    810 |     apply (rule measurable_compose[OF _ return_measurable], auto intro: Mg)
 | 
|  |    811 |     done
 | 
|  |    812 |   also have "\<And>x. x \<in> space M \<Longrightarrow> space (f x) = space M'"
 | 
|  |    813 |     by (intro sets_eq_imp_space_eq sets_kernel[OF Mf])
 | 
|  |    814 |   with measurable_space[OF Mh] 
 | 
|  |    815 |     have "do {x \<leftarrow> M; y \<leftarrow> f x; return N (h x y) \<guillemotright>= g} = do {x \<leftarrow> M; y \<leftarrow> f x; g (h x y)}"
 | 
|  |    816 |     by (intro ballI bind_cong bind_return[OF Mg]) (auto simp: space_pair_measure)
 | 
|  |    817 |   finally show ?thesis ..
 | 
|  |    818 | qed
 | 
|  |    819 | 
 | 
| 58608 |    820 | section {* Measures form a $\omega$-chain complete partial order *}
 | 
| 58606 |    821 | 
 | 
|  |    822 | definition SUP_measure :: "(nat \<Rightarrow> 'a measure) \<Rightarrow> 'a measure" where
 | 
|  |    823 |   "SUP_measure M = measure_of (\<Union>i. space (M i)) (\<Union>i. sets (M i)) (\<lambda>A. SUP i. emeasure (M i) A)"
 | 
|  |    824 | 
 | 
|  |    825 | lemma
 | 
|  |    826 |   assumes const: "\<And>i j. sets (M i) = sets (M j)"
 | 
|  |    827 |   shows space_SUP_measure: "space (SUP_measure M) = space (M i)" (is ?sp)
 | 
|  |    828 |     and sets_SUP_measure: "sets (SUP_measure M) = sets (M i)" (is ?st)
 | 
|  |    829 | proof -
 | 
|  |    830 |   have "(\<Union>i. sets (M i)) = sets (M i)"
 | 
|  |    831 |     using const by auto
 | 
|  |    832 |   moreover have "(\<Union>i. space (M i)) = space (M i)"
 | 
|  |    833 |     using const[THEN sets_eq_imp_space_eq] by auto
 | 
|  |    834 |   moreover have "\<And>i. sets (M i) \<subseteq> Pow (space (M i))"
 | 
|  |    835 |     by (auto dest: sets.sets_into_space)
 | 
|  |    836 |   ultimately show ?sp ?st
 | 
|  |    837 |     by (simp_all add: SUP_measure_def)
 | 
|  |    838 | qed
 | 
|  |    839 | 
 | 
|  |    840 | lemma emeasure_SUP_measure:
 | 
|  |    841 |   assumes const: "\<And>i j. sets (M i) = sets (M j)"
 | 
|  |    842 |     and mono: "mono (\<lambda>i. emeasure (M i))"
 | 
|  |    843 |   shows "emeasure (SUP_measure M) A = (SUP i. emeasure (M i) A)"
 | 
|  |    844 | proof cases
 | 
|  |    845 |   assume "A \<in> sets (SUP_measure M)"
 | 
|  |    846 |   show ?thesis
 | 
|  |    847 |   proof (rule emeasure_measure_of[OF SUP_measure_def])
 | 
|  |    848 |     show "countably_additive (sets (SUP_measure M)) (\<lambda>A. SUP i. emeasure (M i) A)"
 | 
|  |    849 |     proof (rule countably_additiveI)
 | 
|  |    850 |       fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (SUP_measure M)"
 | 
|  |    851 |       then have "\<And>i j. A i \<in> sets (M j)"
 | 
|  |    852 |         using sets_SUP_measure[of M, OF const] by simp
 | 
|  |    853 |       moreover assume "disjoint_family A"
 | 
|  |    854 |       ultimately show "(\<Sum>i. SUP ia. emeasure (M ia) (A i)) = (SUP i. emeasure (M i) (\<Union>i. A i))"
 | 
|  |    855 |         using mono by (subst suminf_SUP_eq) (auto simp: mono_def le_fun_def intro!: SUP_cong suminf_emeasure)
 | 
|  |    856 |     qed
 | 
|  |    857 |     show "positive (sets (SUP_measure M)) (\<lambda>A. SUP i. emeasure (M i) A)"
 | 
|  |    858 |       by (auto simp: positive_def intro: SUP_upper2)
 | 
|  |    859 |     show "(\<Union>i. sets (M i)) \<subseteq> Pow (\<Union>i. space (M i))"
 | 
|  |    860 |       using sets.sets_into_space by auto
 | 
|  |    861 |   qed fact
 | 
|  |    862 | next
 | 
|  |    863 |   assume "A \<notin> sets (SUP_measure M)"
 | 
|  |    864 |   with sets_SUP_measure[of M, OF const] show ?thesis
 | 
|  |    865 |     by (simp add: emeasure_notin_sets)
 | 
|  |    866 | qed
 | 
|  |    867 | 
 | 
|  |    868 | end
 |