diff --git a/thys/S_Finite_Measure_Monad/Kernels.thy b/thys/S_Finite_Measure_Monad/Kernels.thy --- a/thys/S_Finite_Measure_Monad/Kernels.thy +++ b/thys/S_Finite_Measure_Monad/Kernels.thy @@ -1,2961 +1,2989 @@ (* Title: Kernels.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) section \ Kernels \ theory Kernels imports Lemmas_S_Finite_Measure_Monad begin subsection \S-Finite Measures\ locale s_finite_measure = fixes M :: "'a measure" assumes s_finite_sum: "\Mi :: nat \ 'a measure. (\i. sets (Mi i) = sets M) \ (\i. finite_measure (Mi i)) \ (\A\sets M. M A = (\i. Mi i A))" lemma(in sigma_finite_measure) s_finite_measure: "s_finite_measure M" proof obtain A :: "nat \ _" where A: "range A \ sets M" "\ (range A) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" by(metis sigma_finite_disjoint) define Mi where "Mi \ (\i. measure_of (space M) (sets M) (\a. M (a \ A i)))" have emeasure_Mi:"Mi i a = M (a \ A i)" if "a \ sets M" for i a proof - have "positive (sets (Mi i)) (\a. M (a \ A i))" "countably_additive (sets (Mi i)) (\a. M (a \ A i))" unfolding positive_def countably_additive_def proof safe fix B :: "nat \ _" assume "range B \ sets (Mi i)" "disjoint_family B" with A(1) have "range (\j. B j \ A i) \ sets M" "disjoint_family (\j. B j \ A i)" by(fastforce simp: Mi_def disjoint_family_on_def)+ thus "(\j. M (B j \ A i)) = M (\ (range B) \ A i)" by (metis UN_extend_simps(4) suminf_emeasure) qed simp from emeasure_measure_of[OF _ _ this] that show ?thesis by(auto simp add: Mi_def sets.space_closed) qed have sets_Mi:"sets (Mi i) = sets M" for i by(simp add: Mi_def) show "\Mi. (\i. sets (Mi i) = sets M) \ (\i. finite_measure (Mi i)) \ (\A\sets M. emeasure M A = (\i. emeasure (Mi i) A))" proof(safe intro!: exI[where x=Mi]) fix i show "finite_measure (Mi i)" using A by(auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_Mi] emeasure_Mi) next fix B assume B:"B \ sets M" with A(1,4) have "range (\i. B \ A i) \ sets M" "disjoint_family (\i. B \ A i)" by(auto simp: disjoint_family_on_def) then show "M B = (\i. (Mi i) B)" by(simp add: emeasure_Mi[OF B] suminf_emeasure A(2) B) qed(simp_all add: sets_Mi) qed lemmas(in finite_measure) s_finite_measure_finite_measure = s_finite_measure lemmas(in subprob_space) s_finite_measure_subprob = s_finite_measure lemmas(in prob_space) s_finite_measure_prob = s_finite_measure sublocale sigma_finite_measure \ s_finite_measure by(rule s_finite_measure) lemma s_finite_measureI: assumes "\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. A\sets M \ M A = (\i. Mi i A)" shows "s_finite_measure M" by standard (use assms in blast) lemma s_finite_measure_prodI: assumes "\i j. sets (Mij i j) = sets M" "\i j. Mij i j (space M) < \" "\A. A \ sets M \ M A = (\i. (\j. Mij i j A))" shows "s_finite_measure M" proof - define Mi' where "Mi' \ (\n. case_prod Mij (prod_decode n))" have sets_Mi'[measurable_cong]:"\i. sets (Mi' i) = sets M" using assms(1) by(simp add: Mi'_def split_beta') have Mi'_finite:"\i. finite_measure (Mi' i)" using assms(2) sets_eq_imp_space_eq[OF sets_Mi'[symmetric]] top.not_eq_extremum by(fastforce intro!: finite_measureI simp: Mi'_def split_beta') show ?thesis proof(safe intro!: s_finite_measureI[where Mi=Mi'] sets_Mi' Mi'_finite) fix A show "A \ sets M \ M A = (\i. Mi' i A)" by(simp add: assms(3) suminf_ennreal_2dimen[where f="\(x,y). Mij x y A", simplified,OF refl,symmetric] Mi'_def split_beta') qed qed corollary s_finite_measure_s_finite_sumI: assumes "\i. sets (Mi i) = sets M" "\i. s_finite_measure (Mi i)" "\A. A \ sets M \ M A = (\i. Mi i A)" shows "s_finite_measure M" proof - from s_finite_measure.s_finite_sum[OF assms(2)] obtain Mij where Mij[measurable]: "\i j. sets (Mij i j) = sets M" "\i j. finite_measure (Mij i j)" "\i j A. A \ sets M \ Mi i A = (\j. Mij i j A)" by (metis assms(1)) show ?thesis using finite_measure.emeasure_finite[OF Mij(2)] by(auto intro!: s_finite_measure_prodI[where Mij = Mij] simp: assms(3) Mij top.not_eq_extremum) qed +lemma s_finite_measure_finite_sumI: + assumes "finite I" "\i. i \ I \ s_finite_measure (Mi i)" "\i. i \ I \ sets (Mi i) = sets M" + and "\A. A \ sets M \ M A = (\i\I. Mi i A)" + shows "s_finite_measure M" +proof - + let ?Mi = "\n. if n < card I then Mi (from_nat_into I n) else null_measure M" + show ?thesis + proof(rule s_finite_measure_s_finite_sumI[of ?Mi]) + show "\i. s_finite_measure (?Mi i)" + by (metis (full_types) assms(2) bot_nat_0.extremum_strict card.empty emeasure_null_measure ennreal_top_neq_zero finite_measure.s_finite_measure_finite_measure finite_measureI from_nat_into infinity_ennreal_def) + next + show "\i. sets (?Mi i) = sets M" + by (metis (full_types) assms(3) card_gt_0_iff dual_order.strict_trans2 from_nat_into less_zeroE linorder_le_less_linear sets_null_measure) + next + fix A + assume "A \ sets M" + then have "M A = (\i\I. Mi i A)" + by fact + also have "... = (\nnn. ?Mi n A)" + by(rule suminf_finite[symmetric]) auto + finally show "M A = (\n. ?Mi n A)" . + qed +qed + lemma countable_space_s_finite_measure: assumes "countable (space M)" "sets M = Pow (space M)" shows "s_finite_measure M" proof - define Mi where "Mi \ (\i. measure_of (space M) (sets M) (\A. emeasure M (A \ {from_nat_into (space M) i})))" have sets_Mi[measurable_cong,simp]: "sets (Mi i) = sets M" for i by(auto simp: Mi_def) have emeasure_Mi: "emeasure (Mi i) A = emeasure M (A \ {from_nat_into (space M) i})" if [measurable]: "A \ sets M" and i:"i \ to_nat_on (space M) ` (space M)" for i A proof - have "from_nat_into (space M) i \ space M" by (simp add: from_nat_into_def i inv_into_into) hence [measurable]: "{from_nat_into (space M) i} \ sets M" using assms(2) by auto have 1:"countably_additive (sets M) (\A. emeasure M (A \ {from_nat_into (space M) i}))" unfolding countably_additive_def proof safe fix B :: "nat \ _" assume "range B \ sets M" "disjoint_family B" then have [measurable]:"\i. B i \ sets M" and "disjoint_family (\j. B j \ {from_nat_into (space M) i})" by(auto simp: disjoint_family_on_def) then have "(\j. emeasure M (B j \ {from_nat_into (space M) i})) = emeasure M (\ (range (\j. B j \ {from_nat_into (space M) i})))" by(intro suminf_emeasure) auto thus "(\j. emeasure M (B j \ {from_nat_into (space M) i})) = emeasure M (\ (range B) \ {from_nat_into (space M) i})" by simp qed have 2:"positive (sets M) (\A. emeasure M (A \ {from_nat_into (space M) i}))" by(auto simp: positive_def) show ?thesis by(simp add: Mi_def emeasure_measure_of_sigma[OF sets.sigma_algebra_axioms 2 1]) qed define Mi' where "Mi' \ (\i. if i \ to_nat_on (space M) ` (space M) then Mi i else null_measure M)" have [measurable_cong, simp]: "sets (Mi' i) = sets M" for i by(auto simp: Mi'_def) show ?thesis proof(rule s_finite_measure_s_finite_sumI[where Mi=Mi']) fix A assume A[measurable]: "A \ sets M" show "emeasure M A = (\i. emeasure (Mi' i) A)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. emeasure M {x} \count_space A)" using sets.sets_into_space[OF A] by(auto intro!: emeasure_countable_singleton simp: assms(2) countable_subset[OF _ assms(1)]) also have "... = (\\<^sup>+ x. emeasure (Mi (to_nat_on (space M) x)) A \count_space A)" proof(safe intro!: nn_integral_cong) fix x assume "x \ space (count_space A)" then have 1:"x \ A" by simp hence 2:"to_nat_on (space M) x \ to_nat_on (space M) ` (space M)" using A assms(2) by auto have [simp]: "from_nat_into (space M) (to_nat_on (space M) x) = x" by (metis 1 2 A assms(1) eq_from_nat_into_iff in_mono sets.sets_into_space) show "emeasure M {x} = emeasure (Mi (to_nat_on (space M) x)) A" using 1 by(simp add: emeasure_Mi[OF A 2]) qed also have "... = (\\<^sup>+ x\A. emeasure (Mi (to_nat_on (space M) x)) A \count_space UNIV)" by (simp add: nn_integral_count_space_indicator) also have "... = (\\<^sup>+ i\to_nat_on (space M) ` A. emeasure (Mi i) A \count_space UNIV)" by(rule nn_integral_count_compose_inj[OF inj_on_subset[OF inj_on_to_nat_on[OF assms(1)] sets.sets_into_space[OF A]]]) also have "... = (\\<^sup>+ i\to_nat_on (space M) ` A. emeasure (Mi' i) A \count_space UNIV)" proof - { fix x assume "x \ A" then have "to_nat_on (space M) x \ to_nat_on (space M) ` (space M)" using sets.sets_into_space[OF A] by auto hence "emeasure (Mi (to_nat_on (space M) x)) A = emeasure (Mi' (to_nat_on (space M) x)) A" by(auto simp: Mi'_def) } thus ?thesis by(auto intro!: nn_integral_cong simp: indicator_def) qed also have "... = (\\<^sup>+ i. emeasure (Mi' i) A \count_space UNIV)" proof - { fix i assume i:"i \ to_nat_on (space M) ` A" have "from_nat_into (space M) i \ A" if "i \ to_nat_on (space M) ` (space M)" by (metis i image_eqI that to_nat_on_from_nat_into) with emeasure_Mi have "emeasure (Mi' i) A = 0" by(auto simp: Mi'_def) } thus ?thesis by(auto intro!: nn_integral_cong simp: indicator_def) qed also have "... = ?rhs" by(rule nn_integral_count_space_nat) finally show ?thesis . qed next fix i show "s_finite_measure (Mi' i)" proof - { fix x assume h:"x \ space M" "i = to_nat_on (space M) x" then have i:"i \ to_nat_on (space M) ` space M" by blast have x: "from_nat_into (space M) i = x" using h by (simp add: assms(1)) consider "M {x} = 0" | "M {x} \ 0" "M {x} < \" | "M {x} = \" using top.not_eq_extremum by fastforce hence "s_finite_measure (Mi (to_nat_on (space M) x))" proof cases case 1 then have [simp]:"Mi i = null_measure M" by(auto intro!: measure_eqI simp: emeasure_Mi[OF _ i] x Int_insert_right) show ?thesis by(auto simp: h(2)[symmetric] intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) next case 2 then show ?thesis unfolding h(2)[symmetric] by(auto intro!: finite_measure.s_finite_measure_finite_measure finite_measureI simp: sets_eq_imp_space_eq[OF sets_Mi] emeasure_Mi[OF _ i] x h(1)) next case 3 show ?thesis unfolding h(2)[symmetric] s_finite_measure_def proof(safe intro!: exI[where x="\j. return M x"] prob_space.finite_measure prob_space_return h(1)) fix A assume "A \ sets (Mi i)" then have [measurable]: "A \ sets M" by(simp add: Mi_def) thus "emeasure (Mi i) A = (\i. emeasure (return M x) A)" by(simp add: emeasure_Mi[OF _ i] x) (cases "x \ A",auto simp: 3 nn_integral_count_space_nat[symmetric]) qed(auto simp: Mi_def) qed } thus ?thesis by(auto simp: Mi'_def) (auto intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) qed qed simp qed lemma s_finite_measure_subprob_space: "s_finite_measure M \ (\Mi :: nat \ 'a measure. (\i. sets (Mi i) = sets M) \ (\i. (Mi i) (space M) \ 1) \ (\A\sets M. M A = (\i. Mi i A)))" proof assume "\Mi. (\i. sets (Mi i) = sets M) \ (\i. emeasure (Mi i) (space M) \ 1) \ (\A\sets M. M A = (\i. (Mi i) A))" then obtain Mi where 1:"\i. sets (Mi i) = sets M" "\i. emeasure (Mi i) (space M) \ 1" "(\A\sets M. M A = (\i. (Mi i) A))" by auto thus "s_finite_measure M" by(auto simp: s_finite_measure_def sets_eq_imp_space_eq[OF 1(1)] intro!: finite_measureI exI[where x=Mi]) (metis ennreal_one_less_top linorder_not_le) next assume "s_finite_measure M" then obtain Mi' where Mi': "\i. sets (Mi' i) = sets M" "\i. finite_measure (Mi' i)" "\A. A\sets M \ M A = (\i. Mi' i A)" by (metis s_finite_measure.s_finite_sum) obtain u where u:"\i. u i < \" "\i. Mi' i (space M) = u i" using Mi'(2) finite_measure.emeasure_finite top.not_eq_extremum by fastforce define Mmn where "Mmn \ (\(m,n). if n < nat \enn2real (u m)\ then scale_measure (1 / ennreal (real_of_int \enn2real (u m)\)) (Mi' m) else (sigma (space M) (sets M)))" have sets_Mmn : "sets (Mmn k) = sets M" for k by(simp add: Mmn_def split_beta Mi') have emeasure_Mmn: "(Mmn (m, n)) A = (Mi' m A) / ennreal (real_of_int \enn2real (u m)\)" if "n < nat \enn2real (u m)\" "A \ sets M" for n m A by(auto simp: Mmn_def that ennreal_divide_times) have emeasure_Mmn_less1: "(Mmn (m, n)) A \ 1" for m n A proof (cases "n < nat \enn2real (u m)\ \ A \ sets M") case h:True have "(Mi' m) A \ ennreal (real_of_int \enn2real (u m)\)" by(rule order.trans[OF emeasure_mono[OF sets.sets_into_space sets.top]]) (insert u(1) h, auto simp: u(2)[symmetric] enn2real_le top.not_eq_extremum sets_eq_imp_space_eq[OF Mi'(1)] Mi'(1)) with h show ?thesis by(simp add: emeasure_Mmn) (metis divide_le_posI_ennreal dual_order.eq_iff ennreal_zero_divide mult.right_neutral not_gr_zero zero_le) qed(auto simp: Mmn_def emeasure_sigma emeasure_notin_sets Mi') have Mi'_sum:"Mi' m A = (\n. Mmn (m, n) A)" if "A \ sets M" for m A proof - have "(\n. Mmn (m, n) A) = (\n. Mmn (m, n + nat \enn2real (u m)\) A) + (\n< nat \enn2real (u m)\. Mmn (m, n) A)" by(simp add: suminf_offset[where f="\n. Mmn (m, n) A"]) also have "... = (\n< nat \enn2real (u m)\. Mmn (m, n) A)" by(simp add: emeasure_sigma Mmn_def) also have "... = (\n< nat \enn2real (u m)\. (Mi' m A) / ennreal (real_of_int \enn2real (u m)\))" by(rule Finite_Cartesian_Product.sum_cong_aux) (auto simp: emeasure_Mmn that) also have "... = Mi' m A" proof (cases "nat \enn2real (u m)\") case 0 with u[of m] show ?thesis by simp (metis Mi'(1) emeasure_mono enn2real_positive_iff less_le_not_le linorder_less_linear not_less_zero sets.sets_into_space sets.top that) next case (Suc n') then have "ennreal (real_of_int \enn2real (u m)\) > 0" using ennreal_less_zero_iff by fastforce with u(1)[of m] have "of_nat (nat \enn2real (u m)\) / ennreal (real_of_int \enn2real (u m)\) = 1" by (simp add: ennreal_eq_0_iff ennreal_of_nat_eq_real_of_nat) thus ?thesis by (simp add: ennreal_divide_times[symmetric]) qed finally show ?thesis .. qed define Mi where "Mi \ (\i. Mmn (prod_decode i))" show "\Mi. (\i. sets (Mi i) = sets M) \ (\i. emeasure (Mi i) (space M) \ 1) \ (\A\sets M. M A = (\i. (Mi i) A))" by(auto intro!: exI[where x=Mi] simp: Mi_def sets_Mmn suminf_ennreal_2dimen[OF Mi'_sum] Mi'(3)) (metis emeasure_Mmn_less1 prod_decode_aux.cases) qed lemma(in s_finite_measure) finite_measures: obtains Mi where "\i. sets (Mi i) = sets M" "\i. (Mi i) (space M) \ 1" "\A. M A = (\i. Mi i A)" proof - obtain Mi where Mi:"\i. sets (Mi i) = sets M" "\i. (Mi i) (space M) \ 1" "\A. A \ sets M \ M A = (\i. Mi i A)" using s_finite_measure_axioms by(metis s_finite_measure_subprob_space) hence "M A = (\i. Mi i A)" for A by(cases "A \ sets M") (auto simp: emeasure_notin_sets) with Mi(1,2) show ?thesis using that by blast qed lemma(in s_finite_measure) finite_measures_ne: assumes "space M \ {}" obtains Mi where "\i. sets (Mi i) = sets M" "\i. subprob_space (Mi i)" "\A. M A = (\i. Mi i A)" by (metis assms finite_measures sets_eq_imp_space_eq subprob_spaceI) lemma(in s_finite_measure) finite_measures': obtains Mi where "\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by (metis ennreal_top_neq_one finite_measureI finite_measures infinity_ennreal_def sets_eq_imp_space_eq top.extremum_uniqueI) lemma(in s_finite_measure) s_finite_measure_distr: assumes f[measurable]:"f \ M \\<^sub>M N" shows "s_finite_measure (distr M N f)" proof - obtain Mi where Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by(metis finite_measures') show ?thesis by(auto intro!: s_finite_measureI[where Mi="(\i. distr (Mi i) N f)"] finite_measure.finite_measure_distr[OF Mi(2)] simp: emeasure_distr Mi(3) sets_eq_imp_space_eq[OF Mi(1)]) qed lemma nn_integral_measure_suminf: assumes [measurable_cong]:"\i. sets (Mi i) = sets M" and "\A. A\sets M \ M A = (\i. Mi i A)" "f \ borel_measurable M" shows "(\i. \\<^sup>+x. f x \(Mi i)) = (\\<^sup>+x. f x \M)" using assms(3) proof induction case (cong f g) then show ?case by (metis (no_types, lifting) assms(1) nn_integral_cong sets_eq_imp_space_eq suminf_cong) next case (set A) then show ?case using assms(1,2) by simp next case (mult u c) then show ?case by(simp add: nn_integral_cmult) next case (add u v) then show ?case by(simp add: nn_integral_add suminf_add[symmetric]) next case ih:(seq U) have "(\i. integral\<^sup>N (Mi i) (\ range U)) = (\i. \\<^sup>+ x. (\j. U j x) \(Mi i))" by(auto intro!: suminf_cong) (metis SUP_apply) also have "... = (\i. \j. \\<^sup>+ x. U j x \(Mi i))" using ih by(auto intro!: suminf_cong nn_integral_monotone_convergence_SUP) also have "... = (\j. (\i. \\<^sup>+ x. U j x \(Mi i)))" using ih(3) by(auto intro!: ennreal_suminf_SUP_eq incseq_nn_integral) also have "... = (\j. integral\<^sup>N M (U j))" by(simp add: ih) also have "... = (\\<^sup>+ x. (\j. U j x) \M)" using ih by(auto intro!: nn_integral_monotone_convergence_SUP[symmetric]) also have "... = integral\<^sup>N M (\ range U)" by(metis SUP_apply) finally show ?case . qed text \ A @{term \density M f\} of $s$-finite measure @{term M} and @{term \f \ borel_measurable M\} is again s-finite. We do not require additional assumption, unlike $\sigma$-finite measures.\ lemma(in s_finite_measure) s_finite_measure_density: assumes f[measurable]:"f \ borel_measurable M" shows "s_finite_measure (density M f)" proof - obtain Mi where Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by(metis finite_measures') show ?thesis proof(rule s_finite_measure_s_finite_sumI[where Mi="\i. density (Mi i) f"]) show "s_finite_measure (density (Mi i) f)" for i proof - define Mij where "Mij = (\j::nat. if j = 0 then density (Mi i) (\x. \ * indicator {x\space M. f x = \} x) else if j = 1 then density (Mi i) (\x. f x * indicator {x\space M. f x < \} x) else null_measure M)" have sets_Mij[measurable_cong]: "sets (Mij j) = sets M" for j by(auto simp: Mij_def Mi) have emeasure_Mi:"density (Mi i) f A = (\j. Mij j A)" (is "?lhs = ?rhs") if A[measurable]: "A \ sets M" for A proof - have "?lhs = (\\<^sup>+x \ A. f x \Mi i)" by(simp add: emeasure_density) also have "... = (\\<^sup>+x. \ * indicator {x\space M. f x = \} x * indicator A x + f x * indicator {x\space M. f x < \} x * indicator A x \Mi i)" by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF Mi(1)] indicator_def) (simp add: top.not_eq_extremum) also have "... = density (Mi i) (\x. \ * indicator {x\space M. f x = \} x) A + density (Mi i) (\x. f x * indicator {x\space M. f x < \} x) A" by(simp add: nn_integral_add emeasure_density) also have "... = ?rhs" using suminf_finite[of "{..j. Mij j A"] by(auto simp: Mij_def) finally show ?thesis . qed show ?thesis proof(rule s_finite_measure_s_finite_sumI[OF _ _ emeasure_Mi]) fix j :: nat consider "j = 0" | "j = 1" | "j \ 0" "j \ 1" by auto then show "s_finite_measure (Mij j)" proof cases case 1 have 2:"Mij j A = (\k. density (Mi i) (indicator {x\space M. f x = \}) A)" if A[measurable]:"A \ sets M" for A by(auto simp: Mij_def 1 emeasure_density nn_integral_suminf[symmetric] sets_eq_imp_space_eq[OF Mi(1)] indicator_def intro!: nn_integral_cong) (simp add: nn_integral_count_space_nat[symmetric]) show ?thesis by(auto simp: s_finite_measure_def 2 Mi(1)[of i] sets_Mij[of j] intro!: exI[where x="\k. density (Mi i) (indicator {x\space M. f x = \})"] finite_measure.finite_measure_restricted Mi(2)) next case 2 show ?thesis by(auto intro!: sigma_finite_measure.s_finite_measure AE_mono_measure[OF Mi(1)] sum_le_suminf[where I="{i}" and f="\i. Mi i _",simplified] simp: sigma_finite_measure.sigma_finite_iff_density_finite[OF finite_measure.sigma_finite_measure[OF Mi(2)[of i]]] le_measure[OF Mi(1)] Mi indicator_def 2 Mij_def) auto next case 3 then show ?thesis by(auto simp: Mij_def intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) qed qed(auto simp: sets_Mij Mi) qed qed(auto simp: emeasure_density nn_integral_measure_suminf[OF Mi(1,3)] Mi(1)) qed lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable_cong]:"\i. sets (Mi i) = sets M" and "\A. A\sets M \ M A = (\i. Mi i A)" "integrable M f" shows lebesgue_integral_measure_suminf:"(\i. \x. f x \(Mi i)) = (\x. f x \M)" (is "?suminf") and lebesgue_integral_measure_suminf_summable_norm: "summable (\i. norm (\x. f x \(Mi i)))" (is "?summable2") and lebesgue_integral_measure_suminf_summable_norm_in: "summable (\i. \x. norm (f x) \(Mi i))" (is "?summable") proof - have Mi:"Mi i \ M" for i using assms(2) ennreal_suminf_lessD linorder_not_le by(fastforce simp: assms(1) le_measure[OF assms(1)]) have sum2: "summable (\i. norm (\x. g x \(Mi i)))" if "summable (\i. \x. norm (g x) \(Mi i))" for g :: "'a \ 'b" proof(rule summable_suminf_not_top) have "(\i. ennreal (norm (\x. g x \(Mi i)))) \ (\i. ennreal (\x. norm (g x) \(Mi i)))" by(auto intro!: suminf_le) thus "(\i. ennreal (norm (\x. g x \(Mi i)))) \ \" by (metis ennreal_suminf_neq_top[OF that] Bochner_Integration.integral_nonneg neq_top_trans norm_ge_zero) qed simp have "?suminf \ ?summable" using assms(3) proof induction case h[measurable]:(base A c) have Mi_fin:"Mi i A < \" for i by(rule order.strict_trans1[OF _ h(2)], auto simp: le_measureD3[OF Mi assms(1)]) have 1: "(\x. (indicat_real A x *\<^sub>R c) \Mi i) = measure (Mi i) A *\<^sub>R c" for i using Mi_fin by simp have 2:"summable (\i. \x. norm (indicat_real A x *\<^sub>R c) \Mi i)" proof(rule summable_suminf_not_top) show "(\i. ennreal (\x. norm (indicat_real A x *\<^sub>R c) \Mi i)) \ \" (is "?l \ _") proof - have "?l = (\i. Mi i A ) * norm c" using Mi_fin by(auto intro!: suminf_cong simp: measure_def enn2real_mult ennreal_mult) also have "... = M A * norm c" by(simp add: assms(2)) also have "... \ \" using h(2) by (simp add: ennreal_mult_less_top top.not_eq_extremum) finally show ?thesis . qed qed simp have 3: "(\i. \x. indicat_real A x *\<^sub>R c \Mi i) = (\x. indicat_real A x *\<^sub>R c \M)" (is "?l = ?r") proof - have [simp]: "summable (\i. enn2real (Mi i A))" using Mi_fin h by(auto intro!: summable_suminf_not_top simp: assms(2)[symmetric]) have "?l = (\i. measure (Mi i) A) *\<^sub>R c" by(auto intro!: suminf_cong simp: 1 measure_def suminf_scaleR_left) also have "... = ?r" using h(2) Mi_fin by(simp add: ennreal_inj[where a="(\i. measure (Mi i) A)" and b="measure M A",OF suminf_nonneg measure_nonneg,symmetric,simplified measure_def] measure_def suminf_ennreal2[symmetric] assms(2)[symmetric]) finally show ?thesis . qed from 2 3 show ?case by simp next case ih[measurable]:(add f g) have 1:"summable (\i. \x. norm (f x + g x) \Mi i)" proof(rule summable_suminf_not_top) show "(\i. ennreal (\x. norm (f x + g x) \Mi i)) \ \" (is "?l \ _") proof - have "?l = (\i. (\\<^sup>+x. ennreal (norm (f x + g x)) \Mi i))" using ih by(auto intro!: suminf_cong nn_integral_eq_integral[symmetric] integrable_mono_measure[OF assms(1) Mi]) also have "... \ (\i. (\\<^sup>+x. ennreal (norm (f x) + norm (g x)) \Mi i))" by(auto intro!: suminf_le nn_integral_mono norm_triangle_ineq simp del: ennreal_plus) also have "... = (\i. (\\<^sup>+x. ennreal (norm (f x)) \Mi i)) + (\i. (\\<^sup>+x. ennreal (norm (g x)) \Mi i))" by(auto intro!: suminf_cong simp: nn_integral_add suminf_add) also have "... = (\i. ennreal (\x. norm (f x) \Mi i)) + (\i. ennreal (\x. norm (g x) \Mi i))" using ih by(simp add: nn_integral_eq_integral integrable_mono_measure[OF assms(1) Mi]) also have "... < \" using ennreal_suminf_neq_top[OF conjunct2[OF ih(3)]] ennreal_suminf_neq_top[OF conjunct2[OF ih(4)]] by (meson Bochner_Integration.integral_nonneg ennreal_add_eq_top norm_ge_zero top.not_eq_extremum) finally show ?thesis using order.strict_iff_order by blast qed qed simp with ih show ?case by(auto simp: Bochner_Integration.integral_add[OF integrable_mono_measure[OF assms(1) Mi ih(1)] integrable_mono_measure[OF assms(1) Mi ih(2)]] suminf_add[symmetric,OF summable_norm_cancel[OF sum2[OF conjunct2[OF ih(3)]]] summable_norm_cancel[OF sum2[OF conjunct2[OF ih(4)]]]]) next case ih[measurable]:(lim f fn) have 1:"summable (\i. \x. norm (f x) \(Mi i))" proof(rule summable_suminf_not_top) show "(\i. ennreal (\x. norm (f x) \(Mi i))) \ \" (is "?lhs \ _") proof - have "?lhs = (\i. \\<^sup>+ x. ennreal (norm (f x)) \Mi i)" by(auto intro!: suminf_cong nn_integral_eq_integral[symmetric] integrable_mono_measure[OF assms(1) Mi] simp: ih) also have "... = (\\<^sup>+ x. ennreal (norm (f x)) \M)" by(simp add: nn_integral_measure_suminf[OF assms(1,2)]) also have "... = ennreal (\ x. norm (f x) \M)" by(auto intro!: nn_integral_eq_integral ih(4)) also have "... < \" by simp finally show "?lhs \ \" using linorder_neq_iff by blast qed qed simp have "(\i. \x. f x \(Mi i)) = (\i. \x. f x \(Mi i) \(count_space UNIV))" by(rule integral_count_space_nat[symmetric]) (simp add: integrable_count_space_nat_iff sum2[OF 1]) also have "... = lim (\m. \i. \x. fn m x \(Mi i) \(count_space UNIV))" proof(rule limI[OF integral_dominated_convergence[where w="\i. 2 * (\x. norm (f x) \(Mi i))"],symmetric],auto simp: AE_count_space integrable_count_space_nat_iff 1) show "(\m. \x. fn m x \(Mi i)) \ \x. f x \(Mi i)" for i by(rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"],insert ih) (auto intro!: integrable_mono_measure[OF assms(1) Mi] simp: sets_eq_imp_space_eq[OF assms(1)]) next fix i j show "norm (\x. fn j x \(Mi i)) \ 2 * (\x. norm (f x) \(Mi i))" (is "?l \ ?r") proof - have "?l \ (\x. norm (fn j x) \(Mi i))" by simp also have "... \ (\x. 2 * norm (f x) \(Mi i))" by(rule integral_mono,insert ih) (auto intro!: integrable_mono_measure[OF assms(1) Mi] simp: sets_eq_imp_space_eq[OF assms(1)]) finally show "?l \ ?r" by simp qed qed also have "... = lim (\m. (\i. \x. fn m x \(Mi i)))" proof - have "(\i. \x. fn m x \(Mi i) \(count_space UNIV)) = (\i. \x. fn m x \(Mi i))" for m by(auto intro!: integral_count_space_nat sum2 simp: integrable_count_space_nat_iff) (use ih(5) in auto) thus ?thesis by simp qed also have "... = lim (\m. \x. fn m x \M)" by(simp add: ih(5)) also have "... = (\x. f x \M)" using ih by(auto intro!: limI[OF integral_dominated_convergence[where w="\x. 2 * norm (f x)"]]) finally show ?case using 1 by auto qed thus ?suminf ?summable ?summable2 by(simp_all add: sum2) qed (* Ported from sigma-finite measure. The following proof is easier than the sigma-finite measure version. *) lemma (in s_finite_measure) measurable_emeasure_Pair': assumes "Q \ sets (N \\<^sub>M M)" shows "(\x. emeasure M (Pair x -` Q)) \ borel_measurable N" (is "?s Q \ _") proof - obtain Mi where Mi:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by(metis finite_measures') show ?thesis using Mi(1,2) assms finite_measure.finite_measure_cut_measurable[of "Mi _" Q N] by(simp add: Mi(3)) qed lemma (in s_finite_measure) measurable_emeasure'[measurable (raw)]: assumes space: "\x. x \ space N \ A x \ space M" assumes A: "{x\space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M)" shows "(\x. emeasure M (A x)) \ borel_measurable N" proof - from space have "\x. x \ space N \ Pair x -` {x \ space (N \\<^sub>M M). snd x \ A (fst x)} = A x" by (auto simp: space_pair_measure) with measurable_emeasure_Pair'[OF A] show ?thesis by (auto cong: measurable_cong) qed lemma(in s_finite_measure) emeasure_pair_measure': assumes "X \ sets (N \\<^sub>M M)" shows "emeasure (N \\<^sub>M M) X = (\\<^sup>+ x. \\<^sup>+ y. indicator X (x, y) \M \N)" (is "_ = ?\ X") proof (rule emeasure_measure_of[OF pair_measure_def]) show "positive (sets (N \\<^sub>M M)) ?\" by (auto simp: positive_def) have eq[simp]: "\A x y. indicator A (x, y) = indicator (Pair x -` A) y" by (auto simp: indicator_def) show "countably_additive (sets (N \\<^sub>M M)) ?\" proof (rule countably_additiveI) fix F :: "nat \ ('b \ 'a) set" assume F: "range F \ sets (N \\<^sub>M M)" "disjoint_family F" from F have *: "\i. F i \ sets (N \\<^sub>M M)" by auto moreover have "\x. disjoint_family (\i. Pair x -` F i)" by (intro disjoint_family_on_bisimulation[OF F(2)]) auto moreover have "\x. range (\i. Pair x -` F i) \ sets M" using F by (auto simp: sets_Pair1) ultimately show "(\n. ?\ (F n)) = ?\ (\i. F i)" by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure intro!: nn_integral_cong nn_integral_indicator[symmetric]) qed show "{a \ b |a b. a \ sets N \ b \ sets M} \ Pow (space N \ space M)" using sets.space_closed[of N] sets.space_closed[of M] by auto qed fact lemma (in s_finite_measure) emeasure_pair_measure_alt': assumes X: "X \ sets (N \\<^sub>M M)" shows "emeasure (N \\<^sub>M M) X = (\\<^sup>+x. emeasure M (Pair x -` X) \N)" proof - have [simp]: "\x y. indicator X (x, y) = indicator (Pair x -` X) y" by (auto simp: indicator_def) show ?thesis using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure' sets_Pair1) qed proposition (in s_finite_measure) emeasure_pair_measure_Times': assumes A: "A \ sets N" and B: "B \ sets M" shows "emeasure (N \\<^sub>M M) (A \ B) = emeasure N A * emeasure M B" proof - have "emeasure (N \\<^sub>M M) (A \ B) = (\\<^sup>+x. emeasure M B * indicator A x \N)" using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt') also have "\ = emeasure M B * emeasure N A" using A by (simp add: nn_integral_cmult_indicator) finally show ?thesis by (simp add: ac_simps) qed lemma(in s_finite_measure) measure_times: assumes[measurable]: "A \ sets N" "B \ sets M" shows "measure (N \\<^sub>M M) (A \ B) = measure N A * measure M B" by(auto simp: measure_def emeasure_pair_measure_Times' enn2real_mult) lemma pair_measure_s_finite_measure_suminf: assumes Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" and Ni[measurable_cong]:"\i. sets (Ni i) = sets N" "\i. finite_measure (Ni i)" "\A. N A = (\i. Ni i A)" shows "(M \\<^sub>M N) A = (\i j. (Mi i \\<^sub>M Ni j) A)" (is "?lhs = ?rhs") proof - interpret N: s_finite_measure N by(auto intro!: s_finite_measureI[where Mi=Mi] s_finite_measureI[where Mi=Ni] assms) show ?thesis proof(cases "A \ sets (M \\<^sub>M N)") case [measurable]:True show ?thesis proof - have "?lhs = (\\<^sup>+x. N (Pair x -` A) \M)" by(simp add: N.emeasure_pair_measure_alt') also have "... = (\i. \\<^sup>+x. N (Pair x -` A) \Mi i)" using N.measurable_emeasure_Pair'[of A] by(auto intro!: nn_integral_measure_suminf[OF Mi(1,3),symmetric]) also have "... = (\i. \\<^sup>+x. (\j. Ni j (Pair x -` A)) \Mi i)" by(simp add: Ni(3)) also have "... = (\i j. \\<^sup>+x. Ni j (Pair x -` A) \Mi i)" using s_finite_measure.measurable_emeasure_Pair'[OF finite_measure.s_finite_measure_finite_measure[OF Ni(2)],of A] by(auto simp: nn_integral_suminf intro!: suminf_cong) also have "... = ?rhs" by(auto intro!: suminf_cong simp: s_finite_measure.emeasure_pair_measure_alt'[OF finite_measure.s_finite_measure_finite_measure[OF Ni(2)]]) finally show ?thesis . qed next case False with Mi(1) Ni(1) show ?thesis by(simp add: emeasure_notin_sets) qed qed lemma pair_measure_s_finite_measure_suminf': assumes Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" and Ni[measurable_cong]:"\i. sets (Ni i) = sets N" "\i. finite_measure (Ni i)" "\A. N A = (\i. Ni i A)" shows "(M \\<^sub>M N) A = (\i j. (Mi j \\<^sub>M Ni i) A)" (is "?lhs = ?rhs") proof - interpret N: s_finite_measure N by(auto intro!: s_finite_measureI[where Mi=Mi] s_finite_measureI[where Mi=Ni] assms) show ?thesis proof(cases "A \ sets (M \\<^sub>M N)") case [measurable]:True show ?thesis proof - have "?lhs = (\\<^sup>+x. N (Pair x -` A) \M)" by(simp add: N.emeasure_pair_measure_alt') also have "... = (\\<^sup>+x. (\i. Ni i (Pair x -` A)) \M)" by(auto intro!: nn_integral_cong simp: Ni) also have "... = (\i. (\\<^sup>+x. Ni i (Pair x -` A) \M))" by(auto intro!: nn_integral_suminf simp: finite_measure.finite_measure_cut_measurable[OF Ni(2)]) also have "... = (\i j. \\<^sup>+x. Ni i (Pair x -` A) \Mi j)" by(auto intro!: suminf_cong nn_integral_measure_suminf[symmetric] simp: finite_measure.finite_measure_cut_measurable[OF Ni(2)] Mi) also have "... = ?rhs" by(auto intro!: suminf_cong simp: s_finite_measure.emeasure_pair_measure_alt'[OF finite_measure.s_finite_measure_finite_measure[OF Ni(2)]]) finally show ?thesis . qed next case False with Mi(1) Ni(1) show ?thesis by(simp add: emeasure_notin_sets) qed qed lemma pair_measure_s_finite_measure: assumes "s_finite_measure M" and "s_finite_measure N" shows "s_finite_measure (M \\<^sub>M N)" proof - obtain Mi where Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by(metis s_finite_measure.finite_measures'[OF assms(1)]) obtain Ni where Ni[measurable_cong]:"\i. sets (Ni i) = sets N" "\i. finite_measure (Ni i)" "\A. N A = (\i. Ni i A)" by(metis s_finite_measure.finite_measures'[OF assms(2)]) show ?thesis proof(rule s_finite_measure_prodI[where Mij="\i j. Mi i \\<^sub>M Ni j"]) show "emeasure (Mi i \\<^sub>M Ni j) (space (M \\<^sub>M N)) < \" for i j using finite_measure.emeasure_finite[OF Mi(2)[of i]] finite_measure.emeasure_finite[OF Ni(2)[of j]] by(auto simp: sets_eq_imp_space_eq[OF Mi(1)[of i],symmetric] sets_eq_imp_space_eq[OF Ni(1)[of j],symmetric] space_pair_measure s_finite_measure.emeasure_pair_measure_Times'[OF finite_measure.s_finite_measure_finite_measure[OF Ni(2)[of j]]] ennreal_mult_less_top top.not_eq_extremum) qed(auto simp: pair_measure_s_finite_measure_suminf Mi Ni) qed lemma(in s_finite_measure) borel_measurable_nn_integral_fst': assumes [measurable]: "f \ borel_measurable (N \\<^sub>M M)" shows "(\x. \\<^sup>+ y. f (x, y) \M) \ borel_measurable N" proof - obtain Mi where Mi[measurable_cong]:"\i. sets (Mi i) = sets M" "\i. finite_measure (Mi i)" "\A. M A = (\i. Mi i A)" by(metis finite_measures') show ?thesis by(rule measurable_cong[where g="\x. \i. \\<^sup>+ y. f (x, y) \Mi i",THEN iffD2]) (auto simp: nn_integral_measure_suminf[OF Mi(1,3)] intro!: borel_measurable_suminf_order sigma_finite_measure.borel_measurable_nn_integral_fst[OF finite_measure.sigma_finite_measure[OF Mi(2)]]) qed lemma (in s_finite_measure) nn_integral_fst': assumes f: "f \ borel_measurable (M1 \\<^sub>M M)" shows "(\\<^sup>+ x. \\<^sup>+ y. f (x, y) \M \M1) = integral\<^sup>N (M1 \\<^sub>M M) f" (is "?I f = _") using f proof induct case (cong u v) then have "?I u = ?I v" by (intro nn_integral_cong) (auto simp: space_pair_measure) with cong show ?case by (simp cong: nn_integral_cong) qed (simp_all add: emeasure_pair_measure' nn_integral_cmult nn_integral_add nn_integral_monotone_convergence_SUP measurable_compose_Pair1 borel_measurable_nn_integral_fst' nn_integral_mono incseq_def le_fun_def image_comp cong: nn_integral_cong) lemma (in s_finite_measure) borel_measurable_nn_integral'[measurable (raw)]: "case_prod f \ borel_measurable (N \\<^sub>M M) \ (\x. \\<^sup>+ y. f x y \M) \ borel_measurable N" using borel_measurable_nn_integral_fst'[of "case_prod f" N] by simp lemma distr_pair_swap_s_finite: assumes "s_finite_measure M1" and "s_finite_measure M2" shows "M1 \\<^sub>M M2 = distr (M2 \\<^sub>M M1) (M1 \\<^sub>M M2) (\(x, y). (y, x))" (is "?P = ?D") proof - { from s_finite_measure.finite_measures'[OF assms(1)] s_finite_measure.finite_measures'[OF assms(2)] obtain Mi1 Mi2 where Mi1:"\i. sets (Mi1 i) = sets M1" "\i. finite_measure (Mi1 i)" "\A. M1 A = (\i. Mi1 i A)" and Mi2:"\i. sets (Mi2 i) = sets M2" "\i. finite_measure (Mi2 i)" "\A. M2 A = (\i. Mi2 i A)" by metis fix A assume A[measurable]:"A \ sets (M1 \\<^sub>M M2)" have "emeasure (M1 \\<^sub>M M2) A = emeasure (M2 \\<^sub>M M1) ((\(x, y). (y, x)) -` A \ space (M2 \\<^sub>M M1))" proof - { fix i j interpret pair_sigma_finite "Mi1 i" "Mi2 j" by(auto simp: pair_sigma_finite_def Mi1(2) Mi2(2) finite_measure.sigma_finite_measure) have "emeasure (Mi1 i \\<^sub>M Mi2 j) A = emeasure (Mi2 j \\<^sub>M Mi1 i) ((\(x, y). (y, x)) -` A \ space (M2 \\<^sub>M M1))" using Mi1(1) Mi2(1) by(simp add: arg_cong[OF distr_pair_swap,of emeasure] emeasure_distr sets_eq_imp_space_eq[OF sets_pair_measure_cong[OF Mi2(1) Mi1(1)]]) } thus ?thesis by(auto simp: pair_measure_s_finite_measure_suminf'[OF Mi2 Mi1] pair_measure_s_finite_measure_suminf[OF Mi1 Mi2] intro!: suminf_cong) qed } thus ?thesis by(auto intro!: measure_eqI simp: emeasure_distr) qed proposition nn_integral_snd': assumes "s_finite_measure M1" "s_finite_measure M2" and f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = integral\<^sup>N (M1 \\<^sub>M M2) f" proof - interpret M1: s_finite_measure M1 by fact interpret M2: s_finite_measure M2 by fact note measurable_pair_swap[OF f] from M1.nn_integral_fst'[OF this] have "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1))" by simp also have "(\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1)) = integral\<^sup>N (M1 \\<^sub>M M2) f" by (subst distr_pair_swap_s_finite[OF assms(1,2)]) (auto simp add: nn_integral_distr intro!: nn_integral_cong) finally show ?thesis . qed lemma (in s_finite_measure) borel_measurable_lebesgue_integrable'[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "Measurable.pred N (\x. integrable M (f x))" proof - have [simp]: "\x. x \ space N \ integrable M (f x) \ (\\<^sup>+y. norm (f x y) \M) < \" unfolding integrable_iff_bounded by simp show ?thesis by (simp cong: measurable_cong) qed lemma (in s_finite_measure) measurable_measure'[measurable (raw)]: "(\x. x \ space N \ A x \ space M) \ {x \ space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M) \ (\x. measure M (A x)) \ borel_measurable N" unfolding measure_def by (intro measurable_emeasure' borel_measurable_enn2real) auto proposition (in s_finite_measure) borel_measurable_lebesgue_integral'[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "(\x. \y. f x y \M) \ borel_measurable N" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where s: "\i. simple_function (N \\<^sub>M M) (s i)" and "\x\space (N \\<^sub>M M). (\i. s i x) \ (case x of (x, y) \ f x y)" and "\i. \x\space (N \\<^sub>M M). dist (s i x) 0 \ 2 * dist (case x of (x, xa) \ f x xa) 0" by auto then have *: "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" by (rule borel_measurable_simple_function) fact have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" by (rule measurable_simple_function) fact define f' where [abs_def]: "f' i x = (if integrable M (f x) then Bochner_Integration.simple_bochner_integral M (\y. s i (x, y)) else 0)" for i x { fix i x assume "x \ space N" then have "Bochner_Integration.simple_bochner_integral M (\y. s i (x, y)) = (\z\s i ` (space N \ space M). measure M {y \ space M. s i (x, y) = z} *\<^sub>R z)" using s[THEN simple_functionD(1)] unfolding simple_bochner_integral_def by (intro sum.mono_neutral_cong_left) (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } note eq = this show ?thesis proof (rule borel_measurable_LIMSEQ_metric) fix i show "f' i \ borel_measurable N" unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) next fix x assume x: "x \ space N" { assume int_f: "integrable M (f x)" have int_2f: "integrable M (\y. 2 * norm (f x y))" by (intro integrable_norm integrable_mult_right int_f) have "(\i. integral\<^sup>L M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" proof (rule integral_dominated_convergence) from int_f show "f x \ borel_measurable M" by auto show "\i. (\y. s i (x, y)) \ borel_measurable M" using x by simp show "AE xa in M. (\i. s i (x, xa)) \ f x xa" using x * by auto show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" using x * by auto qed fact moreover { fix i have "Bochner_Integration.simple_bochner_integrable M (\y. s i (x, y))" proof (rule simple_bochner_integrableI_bounded) have "(\y. s i (x, y)) ` space M \ s i ` (space N \ space M)" using x by auto then show "simple_function M (\y. s i (x, y))" using simple_functionD(1)[OF s(1), of i] x by (intro simple_function_borel_measurable) (auto simp: space_pair_measure dest: finite_subset) have "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \M) \ (\\<^sup>+ y. 2 * norm (f x y) \M)" using x * by (intro nn_integral_mono) auto also have "(\\<^sup>+ y. 2 * norm (f x y) \M) < \" using int_2f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ xa. ennreal (norm (s i (x, xa))) \M) < \" . qed then have "integral\<^sup>L M (\y. s i (x, y)) = Bochner_Integration.simple_bochner_integral M (\y. s i (x, y))" by (rule simple_bochner_integrable_eq_integral[symmetric]) } ultimately have "(\i. Bochner_Integration.simple_bochner_integral M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" by simp } then show "(\i. f' i x) \ integral\<^sup>L M (f x)" unfolding f'_def by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) qed qed lemma integrable_product_swap_s_finite: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes M1:"s_finite_measure M1" and M2:"s_finite_measure M2" and "integrable (M1 \\<^sub>M M2) f" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" proof - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (rule integrable_distr[OF measurable_pair_swap']) (simp add: distr_pair_swap_s_finite[OF M1 M2,symmetric] assms) qed lemma integrable_product_swap_iff_s_finite: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes M1:"s_finite_measure M1" and M2:"s_finite_measure M2" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" proof - from integrable_product_swap_s_finite[OF M2 M1,of "\(x,y). f (y,x)"] integrable_product_swap_s_finite[OF M1 M2,of f] show ?thesis by auto qed lemma integral_product_swap_s_finite: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes M1:"s_finite_measure M1" and M2:"s_finite_measure M2" and f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" proof - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap_s_finite[OF M1 M2,symmetric]) qed theorem(in s_finite_measure) Fubini_integrable': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M)" and integ1: "integrable M1 (\x. \ y. norm (f (x, y)) \M)" and integ2: "AE x in M1. integrable M (\y. f (x, y))" shows "integrable (M1 \\<^sub>M M) f" proof (rule integrableI_bounded) have "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M)) = (\\<^sup>+ x. (\\<^sup>+ y. norm (f (x, y)) \M) \M1)" by (simp add: nn_integral_fst'[symmetric]) also have "\ = (\\<^sup>+ x. \\y. norm (f (x, y)) \M\ \M1)" proof(rule nn_integral_cong_AE) show "AE x in M1. (\\<^sup>+ y. ennreal (norm (f (x, y))) \M) = ennreal \LINT y|M. norm (f (x, y))\" using integ2 proof eventually_elim fix x assume "integrable M (\y. f (x, y))" then have f: "integrable M (\y. norm (f (x, y)))" by simp then have "(\\<^sup>+y. ennreal (norm (f (x, y))) \M) = ennreal (LINT y|M. norm (f (x, y)))" by (rule nn_integral_eq_integral) simp also have "\ = ennreal \LINT y|M. norm (f (x, y))\" using f by simp finally show "(\\<^sup>+y. ennreal (norm (f (x, y))) \M) = ennreal \LINT y|M. norm (f (x, y))\" . qed qed also have "\ < \" using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) finally show "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M)) < \" . qed fact lemma(in s_finite_measure) emeasure_pair_measure_finite': assumes A: "A \ sets (M1 \\<^sub>M M)" and finite: "emeasure (M1 \\<^sub>M M) A < \" shows "AE x in M1. emeasure M {y\space M. (x, y) \ A} < \" proof - from emeasure_pair_measure_alt'[OF A] finite have "(\\<^sup>+ x. emeasure M (Pair x -` A) \M1) \ \" by simp then have "AE x in M1. emeasure M (Pair x -` A) \ \" by (rule nn_integral_PInf_AE[rotated]) (intro measurable_emeasure_Pair' A) moreover have "\x. x \ space M1 \ Pair x -` A = {y\space M. (x, y) \ A}" using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) ultimately show ?thesis by (auto simp: less_top) qed lemma(in s_finite_measure) AE_integrable_fst''': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M) f" shows "AE x in M1. integrable M (\y. f (x, y))" proof - have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M))" by (rule nn_integral_fst') simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M)) \ \" using f unfolding integrable_iff_bounded by simp finally have "AE x in M1. (\\<^sup>+y. norm (f (x, y)) \M) \ \" by (intro nn_integral_PInf_AE borel_measurable_nn_integral') (auto simp: measurable_split_conv) with AE_space show ?thesis by eventually_elim (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) qed lemma(in s_finite_measure) integrable_fst_norm': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M) f" shows "integrable M1 (\x. \y. norm (f (x, y)) \M)" unfolding integrable_iff_bounded proof show "(\x. \ y. norm (f (x, y)) \M) \ borel_measurable M1" by (rule borel_measurable_lebesgue_integral') simp have "(\\<^sup>+ x. ennreal (norm (\y. norm (f (x, y)) \M)) \M1) = (\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M) \M1)" using AE_integrable_fst'''[OF f] by (auto intro!: nn_integral_cong_AE simp: nn_integral_eq_integral) also have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M))" by (rule nn_integral_fst') simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M)) < \" using f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ x. ennreal (norm (\y. norm (f (x, y)) \M)) \M1) < \" . qed lemma(in s_finite_measure) integrable_fst''': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M) f" shows "integrable M1 (\x. \y. f (x, y) \M)" by(auto intro!: Bochner_Integration.integrable_bound[OF integrable_fst_norm'[OF f]]) proposition(in s_finite_measure) integral_fst''': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M) f" shows "(\x. (\y. f (x, y) \M) \M1) = integral\<^sup>L (M1 \\<^sub>M M) f" using f proof induct case (base A c) have A[measurable]: "A \ sets (M1 \\<^sub>M M)" by fact have eq: "\x y. x \ space M1 \ indicator A (x, y) = indicator {y\space M. (x, y) \ A} y" using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) have int_A: "integrable (M1 \\<^sub>M M) (indicator A :: _ \ real)" using base by (rule integrable_real_indicator) have "(\ x. \ y. indicator A (x, y) *\<^sub>R c \M \M1) = (\x. measure M {y\space M. (x, y) \ A} *\<^sub>R c \M1)" proof (intro integral_cong_AE) from AE_integrable_fst'''[OF int_A] AE_space show "AE x in M1. (\y. indicator A (x, y) *\<^sub>R c \M) = measure M {y\space M. (x, y) \ A} *\<^sub>R c" by eventually_elim (simp add: eq integrable_indicator_iff) qed simp_all also have "\ = measure (M1 \\<^sub>M M) A *\<^sub>R c" proof (subst integral_scaleR_left) have "(\\<^sup>+x. ennreal (measure M {y \ space M. (x, y) \ A}) \M1) = (\\<^sup>+x. emeasure M {y \ space M. (x, y) \ A} \M1)" using emeasure_pair_measure_finite'[OF base] by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) also have "\ = emeasure (M1 \\<^sub>M M) A" using sets.sets_into_space[OF A] by (subst emeasure_pair_measure_alt') (auto intro!: nn_integral_cong arg_cong[where f="emeasure M"] simp: space_pair_measure) finally have *: "(\\<^sup>+x. ennreal (measure M {y \ space M. (x, y) \ A}) \M1) = emeasure (M1 \\<^sub>M M) A" . from base * show "integrable M1 (\x. measure M {y \ space M. (x, y) \ A})" by (simp add: integrable_iff_bounded) then have "(\x. measure M {y \ space M. (x, y) \ A} \M1) = (\\<^sup>+x. ennreal (measure M {y \ space M. (x, y) \ A}) \M1)" by (rule nn_integral_eq_integral[symmetric]) simp also note * finally show "(\x. measure M {y \ space M. (x, y) \ A} \M1) *\<^sub>R c = measure (M1 \\<^sub>M M) A *\<^sub>R c" using base by (simp add: emeasure_eq_ennreal_measure) qed also have "\ = (\ a. indicator A a *\<^sub>R c \(M1 \\<^sub>M M))" using base by simp finally show ?case . next case (add f g) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M)" "g \ borel_measurable (M1 \\<^sub>M M)" by auto have "(\ x. \ y. f (x, y) + g (x, y) \M \M1) = (\ x. (\ y. f (x, y) \M) + (\ y. g (x, y) \M) \M1)" apply (rule integral_cong_AE) apply simp_all using AE_integrable_fst'''[OF add(1)] AE_integrable_fst'''[OF add(3)] apply eventually_elim apply simp done also have "\ = (\ x. f x \(M1 \\<^sub>M M)) + (\ x. g x \(M1 \\<^sub>M M))" using integrable_fst'''[OF add(1)] integrable_fst'''[OF add(3)] add(2,4) by simp finally show ?case using add by simp next case (lim f s) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M)" "\i. s i \ borel_measurable (M1 \\<^sub>M M)" by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (M1 \\<^sub>M M) (s i)) \ integral\<^sup>L (M1 \\<^sub>M M) f" proof (rule integral_dominated_convergence) show "integrable (M1 \\<^sub>M M) (\x. 2 * norm (f x))" using lim(5) by auto qed (insert lim, auto) have "(\i. \ x. \ y. s i (x, y) \M \M1) \ \ x. \ y. f (x, y) \M \M1" proof (rule integral_dominated_convergence) have "AE x in M1. \i. integrable M (\y. s i (x, y))" unfolding AE_all_countable using AE_integrable_fst'''[OF lim(1)] .. with AE_space AE_integrable_fst'''[OF lim(5)] show "AE x in M1. (\i. \ y. s i (x, y) \M) \ \ y. f (x, y) \M" proof eventually_elim fix x assume x: "x \ space M1" and s: "\i. integrable M (\y. s i (x, y))" and f: "integrable M (\y. f (x, y))" show "(\i. \ y. s i (x, y) \M) \ \ y. f (x, y) \M" proof (rule integral_dominated_convergence) show "integrable M (\y. 2 * norm (f (x, y)))" using f by auto show "AE xa in M. (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) by (auto simp: space_pair_measure) show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) by (auto simp: space_pair_measure) qed (insert x, measurable) qed show "integrable M1 (\x. (\ y. 2 * norm (f (x, y)) \M))" by (intro integrable_mult_right integrable_norm integrable_fst''' lim) fix i show "AE x in M1. norm (\ y. s i (x, y) \M) \ (\ y. 2 * norm (f (x, y)) \M)" using AE_space AE_integrable_fst'''[OF lim(1), of i] AE_integrable_fst'''[OF lim(5)] proof eventually_elim fix x assume x: "x \ space M1" and s: "integrable M (\y. s i (x, y))" and f: "integrable M (\y. f (x, y))" from s have "norm (\ y. s i (x, y) \M) \ (\\<^sup>+y. norm (s i (x, y)) \M)" by (rule integral_norm_bound_ennreal) also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \M)" using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure) also have "\ = (\y. 2 * norm (f (x, y)) \M)" using f by (intro nn_integral_eq_integral) auto finally show "norm (\ y. s i (x, y) \M) \ (\ y. 2 * norm (f (x, y)) \M)" by simp qed qed simp_all then show "(\i. integral\<^sup>L (M1 \\<^sub>M M) (s i)) \ \ x. \ y. f (x, y) \M \M1" using lim by simp qed qed lemma (in s_finite_measure) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M) (case_prod f)" shows AE_integrable_fst'': "AE x in M1. integrable M (\y. f x y)" and integrable_fst'': "integrable M1 (\x. \y. f x y \M)" and integrable_fst_norm: "integrable M1 (\x. \y. norm (f x y) \M)" and integral_fst'': "(\x. (\y. f x y \M) \M1) = integral\<^sup>L (M1 \\<^sub>M M) (\(x, y). f x y)" using AE_integrable_fst'''[OF f] integrable_fst'''[OF f] integral_fst'''[OF f] integrable_fst_norm'[OF f] by auto lemma fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes M1:"s_finite_measure M1" and M2:"s_finite_measure M2" and f[measurable]: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_snd_s_finite: "AE y in M2. integrable M1 (\x. f x y)" (is "?AE") and integrable_snd_s_finite: "integrable M2 (\y. \x. f x y \M1)" (is "?INT") and integrable_snd_norm_s_finite: "integrable M2 (\y. \x. norm (f x y) \M1)" (is "?INT2") and integral_snd_s_finite: "(\y. (\x. f x y \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) (case_prod f)" (is "?EQ") proof - interpret Q: s_finite_measure M1 by fact have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f y x)" using f unfolding integrable_product_swap_iff_s_finite[OF M1 M2,symmetric] by simp show ?AE using Q.AE_integrable_fst'''[OF Q_int] by simp show ?INT using Q.integrable_fst'''[OF Q_int] by simp show ?INT2 using Q.integrable_fst_norm[OF Q_int] by simp show ?EQ using Q.integral_fst'''[OF Q_int] using integral_product_swap_s_finite[OF M1 M2,of "case_prod f"] by simp qed proposition Fubini_integral': fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes M1:"s_finite_measure M1" and M2:"s_finite_measure M2" and f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows "(\y. (\x. f x y \M1) \M2) = (\x. (\y. f x y \M2) \M1)" unfolding integral_snd_s_finite[OF assms] s_finite_measure.integral_fst''[OF assms(2,3)] .. locale product_s_finite = fixes M :: "'i \ 'a measure" assumes s_finite_measures: "\i. s_finite_measure (M i)" sublocale product_s_finite \ M?: s_finite_measure "M i" for i by (rule s_finite_measures) locale finite_product_s_finite = product_s_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes finite_index: "finite I" lemma (in product_s_finite) emeasure_PiM: "finite I \ (\i. i\I \ A i \ sets (M i)) \ emeasure (PiM I M) (Pi\<^sub>E I A) = (\i\I. emeasure (M i) (A i))" proof (induct I arbitrary: A rule: finite_induct) case (insert i I) interpret finite_product_s_finite M I by standard fact have "finite (insert i I)" using \finite I\ by auto interpret I': finite_product_s_finite M "insert i I" by standard fact let ?h = "(\(f, y). f(i := y))" let ?P = "distr (Pi\<^sub>M I M \\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h" let ?\ = "emeasure ?P" let ?I = "{j \ insert i I. emeasure (M j) (space (M j)) \ 1}" let ?f = "\J E j. if j \ J then emeasure (M j) (E j) else emeasure (M j) (space (M j))" have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) = (\i\insert i I. emeasure (M i) (A i))" proof (subst emeasure_extend_measure_Pair[OF PiM_def]) fix J E assume "(J \ {} \ insert i I = {}) \ finite J \ J \ insert i I \ E \ (\ j\J. sets (M j))" then have J: "J \ {}" "finite J" "J \ insert i I" and E: "\j\J. E j \ sets (M j)" by auto let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)" let ?p' = "prod_emb I M (J - {i}) (\\<^sub>E j\J-{i}. E j)" have "?\ ?p = emeasure (Pi\<^sub>M I M \\<^sub>M (M i)) (?h -` ?p \ space (Pi\<^sub>M I M \\<^sub>M M i))" by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+ also have "?h -` ?p \ space (Pi\<^sub>M I M \\<^sub>M M i) = ?p' \ (if i \ J then E i else space (M i))" using J E[rule_format, THEN sets.sets_into_space] by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm) also have "emeasure (Pi\<^sub>M I M \\<^sub>M (M i)) (?p' \ (if i \ J then E i else space (M i))) = emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \ J then (E i) else space (M i))" using J E by (intro M.emeasure_pair_measure_Times' sets_PiM_I) auto also have "?p' = (\\<^sub>E j\I. if j \ J-{i} then E j else space (M j))" using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+ also have "emeasure (Pi\<^sub>M I M) (\\<^sub>E j\I. if j \ J-{i} then E j else space (M j)) = (\ j\I. if j \ J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))" using E by (subst insert) (auto intro!: prod.cong) also have "(\j\I. if j \ J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * emeasure (M i) (if i \ J then E i else space (M i)) = (\j\insert i I. ?f J E j)" using insert by (auto simp: mult.commute intro!: arg_cong2[where f="(*)"] prod.cong) also have "\ = (\j\J \ ?I. ?f J E j)" using insert(1,2) J E by (intro prod.mono_neutral_right) auto finally show "?\ ?p = \" . show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \ Pow (\\<^sub>E i\insert i I. space (M i))" using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def) next show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\" using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all next show "(insert i I \ {} \ insert i I = {}) \ finite (insert i I) \ insert i I \ insert i I \ A \ (\ j\insert i I. sets (M j))" using insert by auto qed (auto intro!: prod.cong) with insert show ?case by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space) qed simp lemma (in finite_product_s_finite) measure_times: "(\i. i \ I \ A i \ sets (M i)) \ emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\i\I. emeasure (M i) (A i))" using emeasure_PiM[OF finite_index] by auto lemma (in product_s_finite) nn_integral_empty: "0 \ f (\k. undefined) \ integral\<^sup>N (Pi\<^sub>M {} M) f = f (\k. undefined)" by (simp add: PiM_empty nn_integral_count_space_finite) text \ Every s-finite measure is represented as the push-forward measure of a $\sigma$-finite measure.\ definition Mi_to_NM :: "(nat \ 'a measure) \ 'a measure \ (nat \ 'a) measure" where "Mi_to_NM Mi M \ measure_of (space (count_space UNIV \\<^sub>M M)) (sets (count_space UNIV \\<^sub>M M)) (\A. \i. distr (Mi i) (count_space UNIV \\<^sub>M M) (\x. (i,x)) A)" lemma shows sets_Mi_to_NM[measurable_cong,simp]: "sets (Mi_to_NM Mi M) = sets (count_space UNIV \\<^sub>M M)" and space_Mi_to_NM[simp]: "space (Mi_to_NM Mi M) = space (count_space UNIV \\<^sub>M M)" by(simp_all add: Mi_to_NM_def) context fixes M :: "'a measure" and Mi :: "nat \ 'a measure" assumes sets_Mi[measurable_cong,simp]: "\i. sets (Mi i) = sets M" and emeasure_Mi: "\A. A \ sets M \ M A = (\i. Mi i A)" begin lemma emeasure_Mi_to_NM: assumes [measurable]: "A \ sets (count_space UNIV \\<^sub>M M)" shows "emeasure (Mi_to_NM Mi M) A = (\i. distr (Mi i) (count_space UNIV \\<^sub>M M) (\x. (i,x)) A)" proof(rule emeasure_measure_of[where \="space (count_space UNIV \\<^sub>M M)" and A="sets (count_space UNIV \\<^sub>M M)"]) show "countably_additive (sets (Mi_to_NM Mi M)) (\A. \i. emeasure (distr (Mi i) (count_space UNIV \\<^sub>M M) (Pair i)) A)" unfolding countably_additive_def proof safe fix A :: "nat \ (nat \ _) set" assume "range A \ sets (Mi_to_NM Mi M)" and dA:"disjoint_family A" hence [measurable]: "\i. A i \ sets (count_space UNIV \\<^sub>M M)" by auto show "(\j i. emeasure (distr (Mi i) (count_space UNIV \\<^sub>M M) (Pair i)) (A j)) = (\i. emeasure (distr (Mi i) (count_space UNIV \\<^sub>M M) (Pair i)) (\ (range A)))" (is "?lhs = ?rhs") proof - have "?lhs = (\i j. emeasure (distr (Mi i) (count_space UNIV \\<^sub>M M) (Pair i)) (A j))" by(auto simp: nn_integral_count_space_nat[symmetric] pair_sigma_finite_def sigma_finite_measure_count_space intro!: pair_sigma_finite.Fubini') also have "... = ?rhs" proof(rule suminf_cong) fix n have [simp]:"Pair n -` \ (range A) = (\ (range (\j. Pair n -` A j)))" by auto show " (\j. emeasure (distr (Mi n) (count_space UNIV \\<^sub>M M) (Pair n)) (A j)) = emeasure (distr (Mi n) (count_space UNIV \\<^sub>M M) (Pair n)) (\ (range A))" using dA by(fastforce intro!: suminf_emeasure simp: disjoint_family_on_def emeasure_distr) qed finally show ?thesis . qed qed qed(auto simp: positive_def sets.space_closed Mi_to_NM_def) lemma sigma_finite_Mi_to_NM_measure: assumes "\i. finite_measure (Mi i)" shows "sigma_finite_measure (Mi_to_NM Mi M)" proof - { fix n assume "emeasure (Mi_to_NM Mi M) ({n} \ space M) = \" moreover have "emeasure (Mi_to_NM Mi M) ({n} \ space M) = emeasure (Mi n) (space M)" by(simp add: emeasure_Mi_to_NM emeasure_distr suminf_offset[of _ "Suc n"]) ultimately have False using finite_measure.finite_emeasure_space[OF assms[of n]] by(auto simp: sets_eq_imp_space_eq[OF sets_Mi]) } thus ?thesis by(auto intro!: exI[where x="\i. {{i} \ space M}"] simp: space_pair_measure sigma_finite_measure_def) qed lemma distr_Mi_to_NM_M: "distr (Mi_to_NM Mi M) M snd = M" proof - have [simp]:"Pair i -` snd -` A \ Pair i -` space (count_space UNIV \\<^sub>M M) = A" if "A \ sets M" for A and i :: nat using sets.sets_into_space[OF that] by(auto simp: space_pair_measure) show ?thesis by(auto intro!: measure_eqI simp: emeasure_distr emeasure_Mi_to_NM emeasure_Mi) qed end context fixes \ :: "'a measure" assumes standard_borel_ne: "standard_borel_ne \" and s_finite: "s_finite_measure \" begin interpretation \ : s_finite_measure \ by fact interpretation n_\: standard_borel_ne "count_space (UNIV :: nat set) \\<^sub>M \" by (simp add: pair_standard_borel_ne standard_borel_ne) lemma exists_push_forward: "\(\' :: real measure) f. f \ borel \\<^sub>M \ \ sets \' = sets borel \ sigma_finite_measure \' \ distr \' \ f = \" proof - obtain \i where \i: "\i. sets (\i i) = sets \" "\i. finite_measure (\i i)" "\A. \ A = (\i. \i i A)" by(metis \.finite_measures') show ?thesis proof(safe intro!: exI[where x="distr (Mi_to_NM \i \) borel n_\.to_real"] exI[where x="snd \ n_\.from_real"]) have [simp]:"distr (distr (Mi_to_NM \i \) borel n_\.to_real) (count_space UNIV \\<^sub>M \) n_\.from_real = Mi_to_NM \i \" by(auto simp: distr_distr comp_def intro!:distr_id') show "sigma_finite_measure (distr (Mi_to_NM \i \) borel n_\.to_real)" by(rule sigma_finite_measure_distr[where N="count_space UNIV \\<^sub>M \" and f=n_\.from_real]) (auto intro!: sigma_finite_Mi_to_NM_measure \i) next have [simp]: "distr (Mi_to_NM \i \) \ (snd \ n_\.from_real \ n_\.to_real) = distr (Mi_to_NM \i \) \ snd" by(auto intro!: distr_cong[OF refl]) show "distr (distr (Mi_to_NM \i \) borel n_\.to_real) \ (snd \ n_\.from_real) = \" by(auto simp: distr_distr distr_Mi_to_NM_M[OF \i(1,3)]) qed auto qed abbreviation "\'_and_f \ (SOME (\'::real measure,f). f \ borel \\<^sub>M \ \ sets \' = sets borel \ sigma_finite_measure \' \ distr \' \ f = \)" definition "sigma_pair_\ \ fst \'_and_f" definition "sigma_pair_f \ snd \'_and_f" lemma shows sigma_pair_f_measurable : "sigma_pair_f \ borel \\<^sub>M \" (is ?g1) and sets_sigma_pair_\: "sets sigma_pair_\ = sets borel" (is ?g2) and sigma_finite_sigma_pair_\: "sigma_finite_measure sigma_pair_\" (is ?g3) and distr_sigma_pair: "distr sigma_pair_\ \ sigma_pair_f = \" (is ?g4) proof - have "case \'_and_f of (\',f) \ f \ borel \\<^sub>M \ \ sets \' = sets borel \ sigma_finite_measure \' \ distr \' \ f = \" by(rule someI_ex) (use exists_push_forward in auto) then show ?g1 ?g2 ?g3 ?g4 by(auto simp: sigma_pair_\_def sigma_pair_f_def split_beta) qed end definition s_finite_measure_algebra :: "'a measure \ 'a measure measure" where "s_finite_measure_algebra K = (SUP A \ sets K. vimage_algebra {M. s_finite_measure M \ sets M = sets K} (\M. emeasure M A) borel)" lemma space_s_finite_measure_algebra: "space (s_finite_measure_algebra K) = {M. s_finite_measure M \ sets M = sets K}" by (auto simp add: s_finite_measure_algebra_def space_Sup_eq_UN) lemma s_finite_measure_algebra_cong: "sets M = sets N \ s_finite_measure_algebra M = s_finite_measure_algebra N" by (simp add: s_finite_measure_algebra_def) lemma measurable_emeasure_s_finite_measure_algebra[measurable]: "a \ sets A \ (\M. emeasure M a) \ borel_measurable (s_finite_measure_algebra A)" by (auto intro!: measurable_Sup1 measurable_vimage_algebra1 simp: s_finite_measure_algebra_def) lemma measurable_measure_s_finite_measure_algebra[measurable]: "a \ sets A \ (\M. measure M a) \ borel_measurable (s_finite_measure_algebra A)" unfolding measure_def by measurable lemma s_finite_measure_algebra_measurableD: assumes N: "N \ measurable M (s_finite_measure_algebra S)" and x: "x \ space M" shows "space (N x) = space S" and "sets (N x) = sets S" and "measurable (N x) K = measurable S K" and "measurable K (N x) = measurable K S" using measurable_space[OF N x] by (auto simp: space_s_finite_measure_algebra intro!: measurable_cong_sets dest: sets_eq_imp_space_eq) context fixes K M N assumes K: "K \ measurable M (s_finite_measure_algebra N)" begin lemma s_finite_measure_algebra_kernel: "a \ space M \ s_finite_measure (K a)" using measurable_space[OF K] by (simp add: space_s_finite_measure_algebra) lemma s_finite_measure_algebra_sets_kernel: "a \ space M \ sets (K a) = sets N" using measurable_space[OF K] by (simp add: space_s_finite_measure_algebra) lemma measurable_emeasure_kernel_s_finite_measure_algebra[measurable]: "A \ sets N \ (\a. emeasure (K a) A) \ borel_measurable M" using measurable_compose[OF K measurable_emeasure_s_finite_measure_algebra] . end lemma measurable_s_finite_measure_algebra: "(\a. a \ space M \ s_finite_measure (K a)) \ (\a. a \ space M \ sets (K a) = sets N) \ (\A. A \ sets N \ (\a. emeasure (K a) A) \ borel_measurable M) \ K \ measurable M (s_finite_measure_algebra N)" by (auto intro!: measurable_Sup2 measurable_vimage_algebra2 simp: s_finite_measure_algebra_def) definition bind_kernel :: "'a measure \ ('a \ 'b measure) \ 'b measure" (infixl "\\<^sub>k" 54) where "bind_kernel M k = (if space M = {} then count_space {} else let Y = k (SOME x. x \ space M) in measure_of (space Y) (sets Y) (\B. \\<^sup>+x. (k x B) \M))" lemma bind_kernel_cong_All: assumes "\x. x \ space M \ f x = g x" shows "M \\<^sub>k f = M \\<^sub>k g" proof(cases "space M = {}") case 1:False have "(SOME x. x \ space M) \ space M" by (rule someI_ex) (use 1 in blast) with assms have [simp]:"f (SOME x. x \ space M) = g (SOME x. x \ space M)" by simp have "(\B. \\<^sup>+ x. emeasure (f x) B \M) = (\B. \\<^sup>+ x. emeasure (g x) B \M)" by standard (auto intro!: nn_integral_cong simp: assms) thus ?thesis by(auto simp: bind_kernel_def 1) qed(simp add: bind_kernel_def) lemma sets_bind_kernel: assumes "\x. x \ space M \ sets (k x) = sets N" "space M \ {}" shows "sets (M \\<^sub>k k) = sets N" proof - have "sets (k (SOME x. x \ space M)) = sets N" by(rule someI2_ex) (use assms in auto) with sets_eq_imp_space_eq[OF this] show ?thesis by(simp add: bind_kernel_def assms(2)) qed subsection \ Measure Kernel \ locale measure_kernel = fixes X :: "'a measure" and Y :: "'b measure" and \ :: "'a \ 'b measure" assumes kernel_sets[measurable_cong]: "\x. x \ space X \ sets (\ x) = sets Y" and emeasure_measurable[measurable]: "\B. B \ sets Y \ (\x. emeasure (\ x) B) \ borel_measurable X" and Y_not_empty: "space X \ {} \ space Y \ {}" begin lemma kernel_space :"\x. x \ space X \ space (\ x) = space Y" by (meson kernel_sets sets_eq_imp_space_eq) lemma measure_measurable: assumes "B \ sets Y" shows "(\x. measure (\ x) B) \ borel_measurable X" using emeasure_measurable[OF assms] by (simp add: Sigma_Algebra.measure_def) lemma set_nn_integral_measure: assumes [measurable_cong]: "sets \ = sets X" and [measurable]: "A \ sets X" "B \ sets Y" defines "\ \ measure_of (space Y) (sets Y) (\B. \\<^sup>+x\A. (\ x B) \\)" shows "\ B = (\\<^sup>+x\A. (\ x B) \\)" proof - have nu_sets[measurable_cong]: "sets \ = sets Y" by(simp add: \_def) have "positive (sets Y) (\B. \\<^sup>+x\A. (\ x B) \\)" by(simp add: positive_def) moreover have "countably_additive (sets Y) (\B. \\<^sup>+x\A. (\ x B) \\)" unfolding countably_additive_def proof safe fix C :: "nat \ _" assume h:"range C \ sets Y" "disjoint_family C" thus "(\i. \\<^sup>+x\A. (\ x) (C i)\\) = (\\<^sup>+x\A. (\ x) (\ (range C))\\)" by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF assms(1)] kernel_sets suminf_emeasure nn_integral_suminf[symmetric]) qed ultimately show ?thesis using \_def assms(3) emeasure_measure_of_sigma sets.sigma_algebra_axioms by blast qed corollary nn_integral_measure: assumes "sets \ = sets X" "B \ sets Y" defines "\ \ measure_of (space Y) (sets Y) (\B. \\<^sup>+x. (\ x B) \\)" shows "\ B = (\\<^sup>+x. (\ x B) \\)" using set_nn_integral_measure[OF assms(1) sets.top assms(2)] by(simp add: \_def sets_eq_imp_space_eq[OF assms(1),symmetric]) lemma distr_measure_kernel: assumes [measurable]:"f \ Y \\<^sub>M Z" shows "measure_kernel X Z (\x. distr (\ x) Z f)" unfolding measure_kernel_def proof safe fix B assume B[measurable]: "B \ sets Z" show "(\x. emeasure (distr (\ x) Z f) B) \ borel_measurable X" by(rule measurable_cong[where g= "(\x. \ x (f -` B \ space Y))",THEN iffD2]) (auto simp: emeasure_distr sets_eq_imp_space_eq[OF kernel_sets]) next show "\x. space Z = {} \ x \ space X \ x \ {}" by (metis Y_not_empty assms measurable_empty_iff) qed auto lemma measure_kernel_comp: assumes [measurable]: "f \ W \\<^sub>M X" shows "measure_kernel W Y (\x. \ (f x))" using measurable_space[OF assms] kernel_sets Y_not_empty by(auto simp: measure_kernel_def) lemma emeasure_bind_kernel: assumes "sets \ = sets X" "B \ sets Y" "space X \ {}" shows "(\ \\<^sub>k \) B = (\\<^sup>+x. (\ x B) \\)" proof - have "sets (\ (SOME x. x \ space \)) = sets Y" by(rule someI2_ex) (use assms(3) kernel_sets sets_eq_imp_space_eq[OF assms(1)] in auto) with sets_eq_imp_space_eq[OF this] show ?thesis by(simp add: bind_kernel_def sets_eq_imp_space_eq[OF assms(1) ]assms(3) nn_integral_measure[OF assms(1,2)]) qed lemma measure_bind_kernel: assumes [measurable_cong]:"sets \ = sets X" and [measurable]:"B \ sets Y" "space X \ {}" "AE x in \. \ x B < \" shows "measure (\ \\<^sub>k \) B = (\x. measure (\ x) B \\)" using assms(4) by(auto simp: emeasure_bind_kernel[OF assms(1-3)] measure_def integral_eq_nn_integral intro!: arg_cong[of _ _ enn2real] nn_integral_cong_AE) lemma sets_bind_kernel: assumes "space X \ {}" "sets \ = sets X" shows "sets (\ \\<^sub>k \) = sets Y" using sets_bind_kernel[of \ \, OF kernel_sets,simplified sets_eq_imp_space_eq[OF assms(2)]] by(auto simp: assms(1)) lemma distr_bind_kernel: assumes "space X \ {}" and [measurable_cong]:"sets \ = sets X" and [measurable]: "f \ Y \\<^sub>M Z" shows "distr (\ \\<^sub>k \) Z f = \ \\<^sub>k (\x. distr (\ x) Z f)" proof - { fix A assume A[measurable]:"A \ sets Z" have sets[measurable_cong]:"sets (\ \\<^sub>k \) = sets Y" by(rule sets_bind_kernel[OF assms(1,2)]) have "emeasure (distr (\ \\<^sub>k \) Z f) A = emeasure (\ \\<^sub>k (\x. distr (\ x) Z f)) A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. emeasure (\ x) (f -` A \ space Y) \\)" by(simp add: emeasure_distr sets_eq_imp_space_eq[OF sets] emeasure_bind_kernel[OF assms(2) _ assms(1)]) also have "... = (\\<^sup>+ x. emeasure (distr (\ x) Z f) A \\)" by(auto simp: emeasure_distr sets_eq_imp_space_eq[OF assms(2)] sets_eq_imp_space_eq[OF kernel_sets] intro!: nn_integral_cong) also have "... = ?rhs" by(simp add: measure_kernel.emeasure_bind_kernel[OF distr_measure_kernel[OF assms(3)] assms(2) _ assms(1)]) finally show ?thesis . qed } thus ?thesis by(auto intro!: measure_eqI simp: measure_kernel.sets_bind_kernel[OF distr_measure_kernel[OF assms(3)] assms(1,2)]) qed lemma bind_kernel_distr: assumes [measurable]: "f \ W \\<^sub>M X" and "space W \ {}" shows "distr W X f \\<^sub>k \ = W \\<^sub>k (\x. \ (f x))" proof - have X: "space X \ {}" using measurable_space[OF assms(1)] assms(2) by auto show ?thesis by(rule measure_eqI, insert X) (auto simp: sets_bind_kernel[OF X] measure_kernel.sets_bind_kernel[OF measure_kernel_comp[OF assms(1)] assms(2) refl] emeasure_bind_kernel nn_integral_distr measure_kernel.emeasure_bind_kernel[OF measure_kernel_comp[OF assms(1)] refl _ assms(2)]) qed lemma bind_kernel_return: assumes "x \ space X" shows "return X x \\<^sub>k \ = \ x" proof - have X: "space X \ {}" using assms by auto show ?thesis by(rule measure_eqI) (auto simp: sets_bind_kernel[OF X sets_return] kernel_sets[OF assms] emeasure_bind_kernel[OF sets_return,simplified,OF _ X] nn_integral_return[OF assms]) qed lemma kernel_nn_integral_measurable: assumes "f \ borel_measurable Y" shows "(\x. \\<^sup>+ y. f y \(\ x)) \ borel_measurable X" using assms proof induction case (cong f g) then show ?case by(auto intro!: measurable_cong[where f="\x. \\<^sup>+ y. f y \(\ x)" and g= "\x. \\<^sup>+ y. g y \(\ x)",THEN iffD2] nn_integral_cong simp: sets_eq_imp_space_eq[OF kernel_sets]) next case (set A) then show ?case by(auto intro!: measurable_cong[where f="\x. integral\<^sup>N (\ x) (indicator A)" and g="\x. \ x A",THEN iffD2]) next case (mult u c) then show ?case by(auto intro!: measurable_cong[where f="\x. \\<^sup>+ y. c * u y \\ x" and g="\x. c * \\<^sup>+ y. u y \\ x",THEN iffD2] simp: nn_integral_cmult) next case (add u v) then show ?case by(auto intro!: measurable_cong[where f="\x. \\<^sup>+ y. v y + u y \\ x" and g="\x. (\\<^sup>+ y. v y \\ x) + (\\<^sup>+ y. u y \\ x)",THEN iffD2] simp: nn_integral_add) next case (seq U) then show ?case by(intro measurable_cong[where f="\x. integral\<^sup>N (\ x) (\ range U)" and g="\x. \i. integral\<^sup>N (\ x) (U i)",THEN iffD2]) (auto simp: nn_integral_monotone_convergence_SUP[of U,simplified SUP_apply[symmetric]]) qed lemma bind_kernel_measure_kernel: assumes "measure_kernel Y Z k'" shows "measure_kernel X Z (\x. \ x \\<^sub>k k')" proof(cases "space X = {}") case True then show ?thesis by(auto simp: measure_kernel_def measurable_def) next case X:False then have Y: "space Y \ {}" by(simp add: Y_not_empty) interpret k': measure_kernel Y Z k' by fact show ?thesis proof fix B assume "B \ sets Z" with k'.emeasure_bind_kernel[OF kernel_sets,of _ B] show "(\x. emeasure (\ x \\<^sub>k k') B) \ borel_measurable X" by(auto intro!: measurable_cong[where f="\x. emeasure (\ x \\<^sub>k k') B" and g="\x. \\<^sup>+ y. emeasure (k' y) B \\ x",THEN iffD2] kernel_nn_integral_measurable simp: sets_eq_imp_space_eq[OF kernel_sets] Y) qed(use k'.Y_not_empty Y k'.sets_bind_kernel[OF Y kernel_sets] in auto) qed lemma restrict_measure_kernel: "measure_kernel (restrict_space X A) Y \" proof fix B assume "B \ sets Y" from emeasure_measurable[OF this] show "(\x. emeasure (\ x) B) \ borel_measurable (restrict_space X A)" using measurable_restrict_space1 by blast qed(insert Y_not_empty,auto simp add: space_restrict_space kernel_sets) end lemma measure_kernel_cong_sets: assumes "sets X = sets X'" "sets Y = sets Y'" shows "measure_kernel X Y = measure_kernel X' Y'" by standard (simp add: measure_kernel_def measurable_cong_sets[OF assms(1) refl] sets_eq_imp_space_eq[OF assms(1)] assms(2) sets_eq_imp_space_eq[OF assms(2)]) lemma measure_kernel_pair_countble1: assumes "countable A" "\i. i \ A \ measure_kernel X Y (\x. k (i,x))" shows "measure_kernel (count_space A \\<^sub>M X) Y k" using assms by(auto simp: measure_kernel_def space_pair_measure intro!: measurable_pair_measure_countable1) lemma measure_kernel_empty_trivial: assumes "space X = {}" shows "measure_kernel X Y k" using assms by(auto simp: measure_kernel_def measurable_def) subsection \ Finite Kernel \ locale finite_kernel = measure_kernel + assumes finite_measure_spaces: "\r<\. \x\ space X. \ x (space Y) < r" begin lemma finite_measures: assumes "x \ space X" shows "finite_measure (\ x)" proof- obtain r where "\ x (space Y) < r" using finite_measure_spaces assms by metis then show ?thesis by(auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF kernel_sets[OF assms]]) qed end lemma finite_kernel_empty_trivial: "space X = {} \ finite_kernel X Y f" by(auto simp: finite_kernel_def finite_kernel_axioms_def measure_kernel_empty_trivial intro!: exI[where x=1]) lemma finite_kernel_cong_sets: assumes "sets X = sets X'" "sets Y = sets Y'" shows "finite_kernel X Y = finite_kernel X' Y'" by standard (auto simp: measure_kernel_cong_sets[OF assms] finite_kernel_def finite_kernel_axioms_def sets_eq_imp_space_eq[OF assms(1)] sets_eq_imp_space_eq[OF assms(2)]) subsection \ Sub-Probability Kernel\ locale subprob_kernel = measure_kernel + assumes subprob_spaces: "\x. x \ space X \ subprob_space (\ x)" begin lemma subprob_space: "\x. x \ space X \ \ x (space Y) \ 1" by (simp add: subprob_space.subprob_emeasure_le_1 subprob_spaces) lemma subprob_measurable[measurable]: "\ \ X \\<^sub>M subprob_algebra Y" by(auto intro!: measurable_subprob_algebra_generated[OF sets.sigma_sets_eq[symmetric] sets.Int_stable sets.space_closed] simp: subprob_spaces kernel_sets emeasure_measurable) lemma finite_kernel: "finite_kernel X Y \" by(auto simp: finite_kernel_def finite_kernel_axioms_def intro!: measure_kernel_axioms exI[where x=2] order.strict_trans1[OF subprob_space.subprob_emeasure_le_1[OF subprob_spaces]]) sublocale finite_kernel by (rule finite_kernel) end lemma subprob_kernel_def': "subprob_kernel X Y \ \ \ \ X \\<^sub>M subprob_algebra Y" by(auto simp: subprob_kernel.subprob_measurable subprob_kernel_def subprob_kernel_axioms_def measure_kernel_def measurable_subprob_algebra measurable_empty_iff space_subprob_algebra_empty_iff) (auto simp: subprob_measurableD(2) subprob_space_kernel) lemmas subprob_kernelI = measurable_subprob_algebra[simplified subprob_kernel_def'[symmetric]] lemma subprob_kernel_cong_sets: assumes "sets X = sets X'" "sets Y = sets Y'" shows "subprob_kernel X Y = subprob_kernel X' Y'" by standard (auto simp: subprob_kernel_def' subprob_algebra_cong[OF assms(2)] measurable_cong_sets[OF assms(1) refl]) lemma subprob_kernel_empty_trivial: assumes "space X = {}" shows "subprob_kernel X Y k" using assms by(auto simp: subprob_kernel_def subprob_kernel_axioms_def intro!: measure_kernel_empty_trivial) lemma bind_kernel_bind: assumes "f \ M \\<^sub>M subprob_algebra N" shows "M \\<^sub>k f = M \ f" proof(cases "space M = {}") case True then show ?thesis by(simp add: bind_kernel_def bind_def) next case h:False interpret subprob_kernel M N f using assms(1) by(simp add: subprob_kernel_def') show ?thesis by(rule measure_eqI,insert sets_kernel[OF assms]) (auto simp: h sets_bind_kernel emeasure_bind_kernel[OF refl _ h] emeasure_bind[OF h assms]) qed lemma(in measure_kernel) subprob_kernel_sum: assumes "\x. x \ space X \ finite_measure (\ x)" obtains ki where "\i. subprob_kernel X Y (ki i)" "\A x. x \ space X \ \ x A = (\i. ki i x A)" proof - obtain u where u: "\x. x \ space X \ u x < \" "\x. x \ space X \ u x = \ x (space Y)" using finite_measure.emeasure_finite[OF assms] by (simp add: top.not_eq_extremum) have [measurable]: "u \ borel_measurable X" by(simp cong: measurable_cong add: u(2)) define ki where "ki \ (\i x. if i < nat \enn2real (u x)\ then scale_measure (1 / ennreal (real_of_int \enn2real (u x)\)) (\ x) else (sigma (space Y) (sets Y)))" have 1:"\i x. x \ space X \ sets (ki i x) = sets Y" by(auto simp: ki_def kernel_sets) have "subprob_kernel X Y (ki i)" for i proof - { fix i B assume [measurable]: "B \ sets Y" have "(\x. emeasure (ki i x) B) = (\x. if i < nat \enn2real (u x)\ then (1 / ennreal (real_of_int \enn2real (u x)\)) * emeasure (\ x) B else 0)" by(auto simp: ki_def emeasure_sigma) also have "... \ borel_measurable X" by simp finally have "(\x. emeasure (ki i x) B) \ borel_measurable X" . } moreover { fix i x assume x:"x \ space X" have "emeasure (ki i x) (space Y) \ 1" by(cases "u x = 0",auto simp: ki_def emeasure_sigma u(2)[OF x,symmetric]) (metis u(1)[OF x,simplified] divide_ennreal_def divide_le_posI_ennreal enn2real_le le_of_int_ceiling mult.commute mult.right_neutral not_gr_zero order.strict_iff_not) hence "subprob_space (ki i x)" using x Y_not_empty by(fastforce intro!: subprob_spaceI simp: sets_eq_imp_space_eq[OF 1[OF x]]) } ultimately show ?thesis by(auto simp: subprob_kernel_def measure_kernel_def 1 Y_not_empty subprob_kernel_axioms_def) qed moreover have "\ x A = (\i. ki i x A)" if x:"x \ space X" for x A proof (cases "A \ sets Y") case A[measurable]:True have "emeasure (\ x) A = (\ienn2real (u x)\. emeasure (ki i x) A)" proof(cases "u x = 0") case True then show ?thesis using u(2)[OF that] by simp (metis A emeasure_eq_0 kernel_sets sets.sets_into_space sets.top x) next case u0:False hence "real_of_int \enn2real (u x)\ > 0" by (metis enn2real_nonneg ennreal_0 ennreal_enn2real_if infinity_ennreal_def linorder_not_le nat_0_iff nle_le of_int_le_0_iff of_nat_eq_0_iff real_nat_ceiling_ge u(1) x) with u(1)[OF x] have "of_nat (nat \enn2real (u x)\) / ennreal (real_of_int \enn2real (u x)\) = 1" by(simp add: ennreal_eq_0_iff ennreal_of_nat_eq_real_of_nat) thus ?thesis by(simp add: ki_def ennreal_divide_times[symmetric] mult.assoc[symmetric]) qed then show ?thesis by(auto simp: suminf_offset[of "\i. emeasure (ki i x) A" "nat \enn2real (u x)\"]) (simp add: ki_def emeasure_sigma) next case False then show ?thesis using kernel_sets[OF x] 1[OF x] by(simp add: emeasure_notin_sets) qed ultimately show ?thesis using that by blast qed subsection \ Probability Kernel \ locale prob_kernel = measure_kernel + assumes prob_spaces: "\x. x \ space X \ prob_space (\ x)" begin lemma prob_space: "\x. x \ space X \ \ x (space Y) = 1" using kernel_space prob_space.emeasure_space_1 prob_spaces by fastforce lemma prob_measurable[measurable]: "\ \ X \\<^sub>M prob_algebra Y" by(auto intro!: measurable_prob_algebra_generated[OF sets.sigma_sets_eq[symmetric] sets.Int_stable sets.space_closed] simp: prob_spaces kernel_sets emeasure_measurable) lemma subprob_kernel: "subprob_kernel X Y \" by (simp add: measurable_prob_algebraD subprob_kernel_def') sublocale subprob_kernel by (simp add: subprob_kernel) lemma restrict_probability_kernel: "prob_kernel (restrict_space X A) Y \" by(auto simp: prob_kernel_def restrict_measure_kernel prob_kernel_axioms_def space_restrict_space prob_spaces) end lemma prob_kernel_def': "prob_kernel X Y \ \ \ \ X \\<^sub>M prob_algebra Y" proof assume h:"\ \ X \\<^sub>M prob_algebra Y" show "prob_kernel X Y \" using subprob_measurableD(2)[OF measurable_prob_algebraD[OF h]] measurable_space[OF h] measurable_emeasure_kernel[OF measurable_prob_algebraD[OF h]] by(auto simp: prob_kernel_def measure_kernel_def prob_kernel_axioms_def space_prob_algebra ) (metis prob_space.not_empty sets_eq_imp_space_eq) qed(auto simp: prob_kernel.prob_measurable prob_kernel_def prob_kernel_axioms_def measure_kernel_def) lemma bind_kernel_return'': assumes "sets M = sets N" shows "M \\<^sub>k return N = M" proof(cases "space M = {}") case True then show ?thesis by(simp add: bind_kernel_def space_empty[symmetric]) next case False then have 1: "space N \ {}" by(simp add: sets_eq_imp_space_eq[OF assms]) interpret prob_kernel N N "return N" by(simp add: prob_kernel_def') show ?thesis by(rule measure_eqI) (auto simp: emeasure_bind_kernel[OF assms _ 1] sets_bind_kernel[OF 1 assms] assms) qed subsection\ S-Finite Kernel \ locale s_finite_kernel = measure_kernel + assumes s_finite_kernel_sum: "\ki. (\i. finite_kernel X Y (ki i) \ (\x\space X. \A\sets Y. \ x A = (\i. ki i x A)))" lemma s_finite_kernel_subI: assumes "\x. x \ space X \ sets (\ x) = sets Y" "\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ A \ sets Y \ emeasure (\ x) A = (\i. ki i x A)" shows "s_finite_kernel X Y \" proof - interpret measure_kernel X Y \ proof show "B \ sets Y \ (\x. emeasure (\ x) B) \ borel_measurable X" for B using assms(2) by(simp add: assms(3) subprob_kernel_def' cong: measurable_cong) next show "space X \ {} \ space Y \ {}" using assms(2)[of 0] by(auto simp: subprob_kernel_def measure_kernel_def) qed fact show ?thesis by (auto simp: s_finite_kernel_def measure_kernel_axioms s_finite_kernel_axioms_def assms(2,3) intro!: exI[where x=ki] subprob_kernel.finite_kernel) qed context s_finite_kernel begin lemma s_finite_kernels_fin: obtains ki where "\i. finite_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" proof - obtain ki where ki:"\i. finite_kernel X Y (ki i)" "\x A. x \ space X \ A \ sets Y \ \ x A = (\i. ki i x A)" by(metis s_finite_kernel_sum) hence "\ x A = (\i. ki i x A)" if "x \ space X " for x A by(cases "A \ sets Y", insert that kernel_sets[OF that]) (auto simp: finite_kernel_def measure_kernel_def emeasure_notin_sets) with ki show ?thesis using that by auto qed lemma s_finite_kernels: obtains ki where "\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" proof - obtain ki where ki:"\i. finite_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by(metis s_finite_kernels_fin) have "\kij. (\j. subprob_kernel X Y (kij j)) \ (\x A. x \ space X \ ki i x A = (\j. kij j x A))" for i using measure_kernel.subprob_kernel_sum[of X Y "ki i", OF _ finite_kernel.finite_measures[OF ki(1)[of i]]] ki(1)[of i] by(metis finite_kernel_def) then obtain kij where kij: "\i j. subprob_kernel X Y (kij i j)" "\x A i. x \ space X \ ki i x A = (\j. kij i j x A)" by metis have "\i. subprob_kernel X Y (case_prod kij (prod_decode i))" using kij(1) by(auto simp: split_beta) moreover have "x \ space X \ \ x A = (\i. case_prod kij (prod_decode i) x A)" for x A using suminf_ennreal_2dimen[of "\i. ki i x A" "\(i,j). kij i j x A"] by(auto simp: ki(2) kij(2) split_beta') ultimately show ?thesis using that by fastforce qed lemma image_s_finite_measure: assumes "x \ space X" shows "s_finite_measure (\ x)" proof - obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by(metis s_finite_kernels) show ?thesis using ki(1)[simplified subprob_kernel_def'] measurable_space[OF ki(1)[simplified subprob_kernel_def'] assms] by(auto intro!: s_finite_measureI[where Mi="\i. ki i x"] subprob_space.axioms(1) simp: kernel_sets[OF assms] space_subprob_algebra ki(2)[OF assms]) qed corollary kernel_measurable_s_finite[measurable]:"\ \ X \\<^sub>M s_finite_measure_algebra Y" by(auto intro!: measurable_s_finite_measure_algebra simp: kernel_sets image_s_finite_measure) lemma comp_measurable: assumes f[measurable]:"f \ M \\<^sub>M X" shows "s_finite_kernel M Y (\x. \ (f x))" proof - obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by(metis s_finite_kernels) show ?thesis using ki(1) measurable_space[OF f] by(auto intro!: s_finite_kernel_subI[where ki="\i x. ki i (f x)"] simp: subprob_kernel_def' ki(2) kernel_sets) qed lemma distr_s_finite_kernel: assumes f[measurable]: "f \ Y \\<^sub>M Z" shows "s_finite_kernel X Z (\x. distr (\ x) Z f)" proof - obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by(metis s_finite_kernels) hence 1:"x \ space X \ space (ki i x) = space Y" for x i by(auto simp: subprob_kernel_def' intro!: subprob_measurableD(1)[of _ X Y]) have [measurable]:"B \ sets Z \ (\x. emeasure (distr (\ x) Z f) B) \ borel_measurable X" for B by(rule measurable_cong[where g="\x. \ x (f -` B \ space Y)", THEN iffD2]) (auto simp: emeasure_distr sets_eq_imp_space_eq[OF kernel_sets]) show ?thesis using ki(1) measurable_distr[OF f] by(auto intro!: s_finite_kernel_subI[where ki="\i x. distr (ki i x) Z f"] simp: subprob_kernel_def' emeasure_distr ki(2) sets_eq_imp_space_eq[OF kernel_sets] 1) qed lemma comp_s_finite_measure: assumes "s_finite_measure \" and [measurable_cong]: "sets \ = sets X" shows "s_finite_measure (\ \\<^sub>k \)" proof(cases "space X = {}") case 1:True show ?thesis by(auto simp: sets_eq_imp_space_eq[OF assms(2)] 1 bind_kernel_def intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) next case 0:False then have 1: "space \ \ {}" by(simp add: sets_eq_imp_space_eq[OF assms(2)]) have 2: "sets (\ (SOME x. x \ space \)) = sets Y" by(rule someI2_ex, insert 1 kernel_sets) (auto simp: sets_eq_imp_space_eq[OF assms(2)]) have sets_bind[measurable_cong]: "sets (\ \\<^sub>k \) = sets Y" by(simp add: bind_kernel_def 1 sets_eq_imp_space_eq[OF 2] 2) obtain \i where mui[measurable_cong]: "\i. sets (\i i) = sets X" "\i. (\i i) (space X) \ 1" "\A. \ A = (\i. \i i A)" using s_finite_measure.finite_measures[OF assms(1)] assms(2) sets_eq_imp_space_eq[OF assms(2)] by metis obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by(metis s_finite_kernels) define Mi where "Mi \ (\n. (\(i,j). measure_of (space Y) (sets Y) (\A. \\<^sup>+x. (ki i x A) \(\i j))) (prod_decode n))" have emeasure:"(\ \\<^sub>k \) A = (\i. (Mi i) A)" (is "?lhs = ?rhs") if "A \ sets Y" for A proof - have "?lhs = (\\<^sup>+x. (\ x A) \\)" by(simp add: emeasure_bind_kernel[OF assms(2) that 0]) also have "... = (\\<^sup>+x. (\i. (ki i x A)) \\)" by(auto intro!: nn_integral_cong simp: ki sets_eq_imp_space_eq[OF assms(2)]) also have "... = (\i. \\<^sup>+x. (ki i x A) \\)" by(auto intro!: nn_integral_suminf) (metis ki(1) assms(2) measurable_cong_sets measure_kernel.emeasure_measurable subprob_kernel_def that) also have "... = ?rhs" unfolding Mi_def proof(rule suminf_ennreal_2dimen[symmetric]) fix m interpret kim: subprob_kernel X Y "ki m" by(simp add: ki) show "(\\<^sup>+ x. (ki m x) A \\) = (\n. emeasure (case (m, n) of (i, j) \ measure_of (space Y) (sets Y) (\A. \\<^sup>+ x. emeasure (ki i x) A \\i j)) A)" using kim.emeasure_measurable[OF that] by(simp add: kim.nn_integral_measure[OF mui(1) that] nn_integral_measure_suminf[OF mui(1)[simplified assms(2)[symmetric]] mui(3)]) qed finally show ?thesis . qed have fin:"finite_measure (Mi i)" for i proof(rule prod.exhaust[where y="prod_decode i"]) fix j1 j2 interpret kij: subprob_kernel X Y "ki j1" by(simp add: ki) assume pd:"prod_decode i = (j1, j2)" have "Mi i (space (Mi i)) = (\\<^sup>+x. (ki j1 x (space Y)) \\i j2)" by(auto simp: Mi_def pd kij.nn_integral_measure[OF mui(1) sets.top]) also have "... \ (\\<^sup>+x. 1 \\i j2)" by(intro nn_integral_mono) (metis kij.subprob_space mui(1) sets_eq_imp_space_eq) also have "... \ 1" using mui by (simp add: sets_eq_imp_space_eq[OF mui(1)]) finally show "finite_measure (Mi i)" by (metis ennreal_one_less_top finite_measureI infinity_ennreal_def less_le_not_le) qed have 3: "sets (Mi i) = sets (\ \\<^sub>k \)" for i by(simp add: Mi_def split_beta sets_bind) show "s_finite_measure (\ \\<^sub>k \)" using emeasure fin 3 by (auto intro!: exI[where x=Mi] simp: s_finite_measure_def sets_bind) qed end lemma s_finite_kernel_empty_trivial: assumes "space X = {}" shows "s_finite_kernel X Y k" using assms by(auto simp: s_finite_kernel_def s_finite_kernel_axioms_def intro!: measure_kernel_empty_trivial finite_kernel_empty_trivial) lemma s_finite_kernel_def': "s_finite_kernel X Y \ \ ((\x. x \ space X \ sets (\ x) = sets Y) \ (\ki. (\i. subprob_kernel X Y (ki i)) \ (\x A. x \ space X \ A \ sets Y \ emeasure (\ x) A = (\i. ki i x A))))" (is "?l \ ?r") proof assume ?l then interpret s_finite_kernel X Y \ . from s_finite_kernels obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ emeasure (\ x) A = (\i. emeasure (ki i x) A)" by metis thus ?r by(auto simp: kernel_sets) qed(auto intro!: s_finite_kernel_subI) lemma(in finite_kernel) s_finite_kernel_finite_kernel: "s_finite_kernel X Y \" proof consider "space X = {}" | "space X \ {}" by auto then show "\ki. \i. finite_kernel X Y (ki i) \ (\x\space X. \A\sets Y. (\ x) A = (\i. (ki i x) A))" proof cases case 1 then show ?thesis by(auto simp: finite_kernel_def measure_kernel_def finite_kernel_axioms_def measurable_def intro!: exI[where x=0]) next case 2 then have y:"space Y \ {}" by(simp add: Y_not_empty) define ki where "ki i \ case i of 0 \ \ | Suc _ \ (\_. sigma (space Y) (sets Y))" for i have "finite_kernel X Y (ki i)" for i by (cases i, auto simp: ki_def finite_kernel_axioms) (auto simp: emeasure_sigma finite_kernel_def measure_kernel_def finite_kernel_axioms_def y intro!: finite_measureI exI[where x=1]) moreover have "(\ x) A = (\i. (ki i x) A)" for x A by(simp add: suminf_offset[where i="Suc 0" and f="\i. ki i x A",simplified],simp add: ki_def emeasure_sigma) ultimately show ?thesis by auto qed qed lemmas(in subprob_kernel) s_finite_kernel_subprob_kernel = s_finite_kernel_finite_kernel lemmas(in prob_kernel) s_finite_kernel_prob_kernel = s_finite_kernel_subprob_kernel sublocale finite_kernel \ s_finite_kernel by(rule s_finite_kernel_finite_kernel) lemma s_finite_kernel_cong_sets: assumes "sets X = sets X'" "sets Y = sets Y'" shows "s_finite_kernel X Y = s_finite_kernel X' Y'" by standard (simp add: s_finite_kernel_def measurable_cong_sets[OF assms(1) refl] sets_eq_imp_space_eq[OF assms(1)] assms(2) measure_kernel_cong_sets[OF assms] s_finite_kernel_axioms_def finite_kernel_cong_sets[OF assms]) lemma(in s_finite_kernel) s_finite_kernel_cong: assumes "\x. x \ space X \ \ x = g x" shows "s_finite_kernel X Y g" using assms s_finite_kernel_axioms by(auto simp: s_finite_kernel_def s_finite_kernel_axioms_def measure_kernel_def cong: measurable_cong) lemma(in s_finite_measure) s_finite_kernel_const: assumes "space M \ {}" shows "s_finite_kernel X M (\x. M)" proof obtain Mi where Mi:"\i. sets (Mi i) = sets M" "\i. (Mi i) (space M) \ 1" "\A. M A = (\i. Mi i A)" by(metis finite_measures) hence "\i. subprob_kernel X M (\x. Mi i)" by(auto simp: subprob_kernel_def' space_subprob_algebra sets_eq_imp_space_eq[OF Mi(1)] assms intro!:measurable_const subprob_spaceI) thus "\ki. \i. finite_kernel X M (ki i) \ (\x\space X. \A\sets M. M A = (\i. (ki i x) A))" by(auto intro!: exI[where x="\i x. Mi i"] Mi(3) subprob_kernel.finite_kernel) qed (auto simp: assms) lemma s_finite_kernel_pair_countble1: assumes "countable A" "\i. i \ A \ s_finite_kernel X Y (\x. k (i,x))" shows "s_finite_kernel (count_space A \\<^sub>M X) Y k" proof - have "\ki. (\j. subprob_kernel X Y (ki j)) \ (\x B. x \ space X \ B \ sets Y \ k (i,x) B = (\j. ki j x B))" if "i \ A" for i using s_finite_kernel.s_finite_kernels[OF assms(2)[OF that]] by metis then obtain ki where ki:"\i j. i \ A \ subprob_kernel X Y (ki i j)" "\i x B. i \ A \ x \ space X \ B \ sets Y \ k (i,x) B = (\j. ki i j x B)" by metis then show ?thesis using assms(2) by(auto simp: s_finite_kernel_def' measure_kernel_pair_countble1[OF assms(1)] subprob_kernel_def' space_pair_measure intro!: exI[where x="\j (i,x). ki i j x"] measurable_pair_measure_countable1 assms(1)) qed lemma s_finite_kernel_s_finite_kernel: assumes "\i. s_finite_kernel X Y (ki i)" "\x. x \ space X \ sets (k x) = sets Y" "\x A. x \ space X \ A \ sets Y \ emeasure (k x) A = (\i. (ki i) x A)" shows "s_finite_kernel X Y k" proof - have "\kij. (\j. subprob_kernel X Y (kij j)) \ (\x A. x \ space X \ ki i x A = (\j. kij j x A))" for i using s_finite_kernel.s_finite_kernels[OF assms(1)[of i]] by metis then obtain kij where kij:"\i j. subprob_kernel X Y (kij i j)" "\i x A. x \ space X \ ki i x A = (\j. kij i j x A)" by metis define ki' where "ki' \ (\n. case_prod kij (prod_decode n))" have emeasure_sumk':"emeasure (k x) A = (\i. emeasure (ki' i x) A)" if x:"x \ space X" and A: "A \ sets Y" for x A by(auto simp: assms(3)[OF that] kij(2)[OF x] ki'_def intro!: suminf_ennreal_2dimen[symmetric]) have "subprob_kernel X Y (ki' i)" for i using kij(1) by(auto simp: ki'_def split_beta') thus ?thesis by(auto simp: s_finite_kernel_def' measure_kernel_def assms(2) s_finite_kernel_axioms_def emeasure_sumk' intro!: exI[where x=ki']) qed lemma s_finite_kernel_finite_sumI: assumes [measurable_cong]: "\x. x \ space X \ sets (\ x) = sets Y" and "\i. i \ I \ subprob_kernel X Y (ki i)" "\x A. x \ space X \ A \ sets Y \ emeasure (\ x) A = (\i\I. ki i x A)" "finite I" "I \ {}" shows "s_finite_kernel X Y \" proof - consider "space X = {}" | "space X \ {}" by auto then show ?thesis proof cases case 1 then show ?thesis by(rule s_finite_kernel_empty_trivial) next case 2 then have Y:"space Y \ {}" using assms measure_kernel.Y_not_empty by (fastforce simp: subprob_kernel_def) define ki' where "ki' \ (\i x. if i < card I then ki (from_nat_into I i) x else null_measure Y)" have [simp]:"subprob_kernel X Y (ki' i)" for i by(cases "i < card I") (simp add: ki'_def from_nat_into assms, auto simp: ki'_def subprob_kernel_def measure_kernel_def subprob_kernel_axioms_def Y intro!: subprob_spaceI) have [simp]: "(\i. emeasure (ki' i x) A) = (\i\I. ki i x A)" for x A using suminf_finite[of "{..i. (if i < card I then ki (from_nat_into I i) x else null_measure Y) A"] by(auto simp: sum.reindex_bij_betw[OF bij_betw_from_nat_into_finite[OF assms(4)],symmetric] ki'_def) have [measurable]:"B \ sets Y \ (\x. emeasure (\ x) B) \ borel_measurable X" for B using assms(2) by(auto simp: assms(3) subprob_kernel_def' cong: measurable_cong) show ?thesis by (auto simp: s_finite_kernel_def' intro!: exI[where x=ki'] assms) qed qed text \ Each kernel does not need to be bounded by a uniform upper-bound in the definition of @{term s_finite_kernel} \ lemma s_finite_kernel_finite_bounded_sum: assumes [measurable_cong]: "\x. x \ space X \ sets (\ x) = sets Y" and "\i. measure_kernel X Y (ki i)" "\x A. x \ space X \ A \ sets Y \ \ x A = (\i. ki i x A)" "\i x. x \ space X \ ki i x (space Y) < \" shows "s_finite_kernel X Y \" proof(cases "space X = {}") case True then show ?thesis by(simp add: s_finite_kernel_empty_trivial) next case X:False then have Y: "space Y \ {}" using assms(2)[of 0] by(simp add: measure_kernel_def) show ?thesis proof(rule s_finite_kernel_s_finite_kernel[where ki=ki,OF _ assms(1) assms(3)]) fix i interpret m: measure_kernel X Y "ki i" by fact define kij where "kij \ (\(j :: nat) x. if j < nat \enn2real (ki i x (space Y))\ then scale_measure (1 / ennreal \enn2real (ki i x (space Y))\) (ki i x) else sigma (space Y) (sets Y))" have sets_kij: "sets (kij j x) = sets Y" if "x \ space X" for j x by(auto simp: m.kernel_sets[OF that] kij_def) have emeasure_kij: "ki i x A = (\j. kij j x A)" if "x \ space X" "A \ sets Y" for x A proof - have "(\j. kij j x A) = (\j< nat \enn2real (ki i x (space Y))\. scale_measure (1 / ennreal \enn2real (ki i x (space Y))\) (ki i x) A)" by(simp add: suminf_offset[where i="nat \enn2real (ki i x (space Y))\" and f="\j. kij j x A"], simp add: kij_def emeasure_sigma) also have "... = ki i x A" proof(cases "nat \enn2real (ki i x (space Y))\") case 0 then show ?thesis by simp (metis assms(4) emeasure_eq_0 enn2real_le ennreal_0 infinity_ennreal_def le_zero_eq linorder_not_le m.kernel_space nle_le sets.sets_into_space sets.top that) next case (Suc n') then have "ennreal (real_of_int \enn2real (emeasure (ki i x) (space Y))\) > 0" using ennreal_less_zero_iff by fastforce with assms(4)[OF that(1),of i] have [simp]: "of_nat (nat \enn2real (emeasure (ki i x) (space Y))\) / ennreal (real_of_int \enn2real (emeasure (ki i x) (space Y))\) = 1" by (simp add: ennreal_eq_0_iff ennreal_of_nat_eq_real_of_nat) show ?thesis by(simp add: mult.assoc[symmetric] ennreal_times_divide) qed finally show ?thesis by simp qed have sk: "subprob_kernel X Y (kij j)" for j proof - { fix B assume [measurable]:"B \ sets Y" have "emeasure (kij j x) B = (if j < nat \enn2real (ki i x (space Y))\ then (ki i x) B / ennreal (real_of_int \enn2real (ki i x (space Y))\) else 0)" if "x \ space X" for x by(auto simp: kij_def emeasure_sigma divide_ennreal_def mult.commute) hence " (\x. emeasure (kij j x) B) \ borel_measurable X" by(auto simp: kij_def cong: measurable_cong) } moreover { fix x assume x:"x \ space X" have "subprob_space (kij j x)" proof - have "emeasure (kij j x) (space Y) \ 1" proof - { assume 1:"j < nat \enn2real (emeasure (ki i x) (space Y))\" then have "emeasure (ki i x) (space Y) > 0" by (metis ceiling_zero enn2real_0 nat_zero_as_int not_gr_zero not_less_zero) with assms(4)[OF x] have [simp]:"emeasure (ki i x) (space Y) / emeasure (ki i x) (space Y) = 1" by simp have [simp]:"emeasure (ki i x) (space Y) / ennreal (real_of_int \enn2real (ki i x (space Y))\) \ 1" proof(rule order.trans[where b="emeasure (ki i x) (space Y) / ki i x (space Y)",OF divide_le_posI_ennreal]) show "0 < ennreal (real_of_int \enn2real (ki i x (space Y))\)" using 1 assms(4)[OF x] enn2real_positive_iff top.not_eq_extremum by fastforce next have 1:"ennreal (real_of_int \enn2real (ki i x (space Y))\) \ ki i x (space Y)" using assms(4)[OF x] enn2real_le by (simp add: linorder_neq_iff) have "ennreal (real_of_int \enn2real (ki i x (space Y))\) / ki i x (space Y) \ 1" by(rule order.trans[OF _ divide_right_mono_ennreal[OF 1,of "ki i x (space Y)"]]) simp thus "emeasure (ki i x) (space Y) \ ennreal (real_of_int \enn2real (ki i x (space Y))\) * (emeasure (ki i x) (space Y) / ki i x (space Y))" by (simp add: "1") qed simp have "1 / ennreal (real_of_int \enn2real (emeasure (ki i x) (space Y))\) * emeasure (ki i x) (space Y) \ 1" by (simp add: ennreal_divide_times) } thus ?thesis by(auto simp: kij_def emeasure_sigma) qed thus ?thesis by(auto intro!: subprob_spaceI simp: sets_eq_imp_space_eq[OF sets_kij[OF x,of j]] Y) qed } ultimately show ?thesis by(auto simp: subprob_kernel_def measure_kernel_def sets_kij m.Y_not_empty subprob_kernel_axioms_def) qed show "s_finite_kernel X Y (ki i)" by(auto intro!: s_finite_kernel_subI simp: emeasure_kij sk m.kernel_sets) qed simp_all qed lemma(in measure_kernel) s_finite_kernel_finite_bounded: assumes "\x. x \ space X \ \ x (space Y) < \" shows "s_finite_kernel X Y \" proof(cases "space X = {}") case True then show ?thesis by(simp add: s_finite_kernel_empty_trivial) next case False then have Y:"space Y \ {}" by(simp add: Y_not_empty) have "measure_kernel X Y (case i of 0 \ \ | Suc x \ \x. null_measure Y)" for i by(cases i,auto simp: measure_kernel_axioms) (auto simp: measure_kernel_def Y) moreover have "\ x A = (\i. emeasure ((case i of 0 \ \ | Suc x \ \x. null_measure Y) x) A)" for x A by(simp add: suminf_offset[where i="Suc 0"]) moreover have "x \ space X \ emeasure ((case i of 0 \ \ | Suc x \ \x. null_measure Y) x) (space Y) < \" for x i by(cases i) (use assms in auto) ultimately show ?thesis by(auto intro!: s_finite_kernel_finite_bounded_sum[where ki="\i. case i of 0 \ \ | Suc _ \ (\x. null_measure Y)" and X=X and Y=Y] simp: kernel_sets) qed lemma(in s_finite_kernel) density_s_finite_kernel: assumes f[measurable]: "case_prod f \ X \\<^sub>M Y \\<^sub>M borel" shows "s_finite_kernel X Y (\x. density (\ x) (f x))" proof(cases "space X = {}") case True then show ?thesis by(simp add: s_finite_kernel_empty_trivial) next case False note Y = Y_not_empty[OF this] obtain ki' where ki': "\i. subprob_kernel X Y (ki' i)" "\x A. x \ space X \ \ x A = (\i. ki' i x A)" by(metis s_finite_kernels) hence sets_ki'[measurable_cong]:"\x i. x \ space X \ sets (ki' i x) = sets Y" by(auto simp: subprob_kernel_def measure_kernel_def) define ki where "ki \ (\i x. density (ki' i x) (f x))" have sets_ki: "x \ space X \ sets (ki i x) = sets Y" for i x using ki'(1) by(auto simp: subprob_kernel_def measure_kernel_def ki_def) have emeasure_k:"density (\ x) (f x) A = (\i. ki i x A)" if x:"x \ space X" and A[measurable]:"A \ sets Y" for x A using kernel_sets[OF x] x ki'(1) sets_ki'[OF x] by(auto simp: emeasure_density nn_integral_measure_suminf[OF _ ki'(2),of x] ki_def) show ?thesis proof(rule s_finite_kernel_s_finite_kernel[where ki="ki",OF _ _ emeasure_k]) fix i note nn_integral_measurable_subprob_algebra2[OF _ ki'(1)[of i,simplified subprob_kernel_def'],measurable] define kij where "kij \ (\j x. if j = 0 then density (ki' i x) (\y. \ * indicator {y\space Y. f x y = \} y) else if j = (Suc 0) then density (ki' i x) (\y. f x y * indicator {y\space Y. f x y < \} y) else null_measure Y)" have emeasure_kij: "ki i x A = (\j. kij j x A)" (is "?lhs = ?rhs") if x:"x \ space X" and [measurable]: "A \ sets Y" for x A proof - have "?lhs = (\\<^sup>+y\A. f x y \ki' i x)" using sets_ki[OF x,of i] x by(auto simp: ki_def emeasure_density) also have "... = (\\<^sup>+y. (\ * indicator {y \ space Y. f x y = \} y * indicator A y + f x y * indicator {y \ space Y. f x y < \} y * indicator A y) \ki' i x)" by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF sets_ki'[OF x]] indicator_def) (simp add: top.not_eq_extremum) also have "... = density (ki' i x) (\y. \ * indicator {y\space Y. f x y = \} y) A + density (ki' i x) (\y. f x y * indicator {y\space Y. f x y < \} y) A" using sets_ki[OF x,of i] x by(auto simp: ki_def emeasure_density nn_integral_add) also have "... = ?rhs" using suminf_finite[of "{..j. kij j x A"] by(simp add: kij_def) finally show ?thesis . qed have sets_kij[measurable_cong]:"x \ space X \ sets (kij j x) = sets Y" for j x by(auto simp: kij_def sets_ki') show "s_finite_kernel X Y (ki i)" proof(rule s_finite_kernel_s_finite_kernel[where ki=kij,OF _ _ emeasure_kij]) fix j consider "j = 0" | "j = Suc 0" | "j \ 0" "j \ Suc 0" by auto then show "s_finite_kernel X Y (kij j)" proof cases case 1 have emeasure_ki: "emeasure (kij j x) A = (\j. emeasure (density (ki' i x) (indicator {y \ space Y. f x y = \})) A)" if x:"x \ space X" and [measurable]: "A \ sets Y" for x A using sets_ki'[OF x] x by(auto simp: 1 kij_def emeasure_density nn_integral_suminf[symmetric] indicator_def intro!: nn_integral_cong) (simp add: nn_integral_count_space_nat[symmetric]) have [simp]:"subprob_kernel X Y (\x. density (ki' i x) (indicator {y \ space Y. f x y = \}))" proof - have [simp]:"x \ space X \ set_nn_integral (ki' i x) (space Y) (indicator {y \ space Y. f x y = \}) \ 1" for x by(rule order.trans[OF nn_integral_mono[where v="\x. 1"]],insert ki'(1)[of i]) (auto simp: indicator_def subprob_kernel_def subprob_kernel_axioms_def intro!: subprob_space.emeasure_space_le_1) show ?thesis by(auto simp: subprob_kernel_def measure_kernel_def emeasure_density subprob_kernel_axioms_def sets_ki' sets_eq_imp_space_eq[OF sets_ki'] Y cong: measurable_cong intro!: subprob_spaceI) qed show ?thesis by (auto simp: s_finite_kernel_def' sets_kij intro!: exI[where x="\k x. density (ki' i x) (indicator {y \ space Y. f x y = \})"] simp: emeasure_ki ) next case j:2 have emeasure_ki: "emeasure (kij j x) A = (\k. density (ki' i x) (\y. f x y * indicator {y \ space Y. of_nat k \ f x y \ f x y < 1 + of_nat k} y) A)" if x:"x \ space X" and [measurable]:"A \ sets Y" for x A proof - have [simp]: "f x y * indicator {y \ space Y. f x y < \} y * indicator A y = f x y * (\k. indicator {y \ space Y. of_nat k \ f x y \ f x y < 1 + of_nat k} y) * indicator A y" if y:"y \ space Y" for y proof(cases "f x y < \") case f:True define l where "l \ floor (enn2real (f x y))" have "nat l \ enn2real (f x y)" "enn2real (f x y) < 1 + nat l" by (simp_all add: l_def) linarith with y have l:"of_nat (nat l) \ f x y" "f x y < 1 + of_nat (nat l)" using Orderings.order_eq_iff enn2real_positive_iff ennreal_enn2real_if ennreal_of_nat_eq_real_of_nat linorder_not_le of_nat_0_le_iff f by fastforce+ then have "(\j. indicator {y \ space Y. of_nat j \ f x y \ f x y < 1 + of_nat j} y :: ennreal) = (\j. if j = nat l then 1 else 0)" by(auto intro!: suminf_cong simp: indicator_def y) (metis Suc_leI linorder_neqE_nat linorder_not_less of_nat_Suc of_nat_le_iff order_trans) also have "... = 1" using suminf_finite[where N="{nat l}" and f="\j. if j = nat l then 1 else (0 :: ennreal)"] by simp finally show ?thesis by(auto, insert f) (auto simp: indicator_def) qed(use top.not_eq_extremum in fastforce) show ?thesis using sets_ki[OF x] sets_ki'[OF x] x by(auto simp: kij_def j emeasure_density nn_integral_suminf[symmetric] sets_eq_imp_space_eq[OF sets_ki'[OF x]] intro!: nn_integral_cong) qed show ?thesis proof(rule s_finite_kernel_finite_bounded_sum[OF sets_kij _ emeasure_ki]) fix k show "measure_kernel X Y (\x. density (ki' i x) (\y. f x y * indicator {y \ space Y. of_nat k \ f x y \ f x y < 1 + of_nat k} y))" using sets_ki'[of _ i] by(auto simp: measure_kernel_def emeasure_density Y cong: measurable_cong) next fix k x assume x:"x \space X" have "emeasure (density (ki' i x) (\y. f x y * indicator {y \ space Y. of_nat k \ f x y \ f x y < 1 + of_nat k} y)) (space Y) \ 1 + of_nat k" by(auto simp: emeasure_density x,rule order.trans[OF nn_integral_mono[where v="\x. 1 + of_nat k"]]) (insert subprob_kernel.subprob_space[OF ki'(1)[of i] x], auto simp: indicator_def subprob_kernel_def subprob_kernel_axioms_def sets_eq_imp_space_eq[OF sets_ki'[OF x]] intro!: mult_mono[where d="1 :: ennreal",OF order.refl,simplified]) also have "... < \" by (simp add: of_nat_less_top) finally show "emeasure (density (ki' i x) (\y. f x y * indicator {y \ space Y. of_nat k \ f x y \ f x y < 1 + of_nat k} y)) (space Y) < \" . qed auto next case 3 then show ?thesis by(auto simp: kij_def s_finite_kernel_cong_sets[of X X Y,OF _ sets_null_measure[symmetric]] Y intro!: s_finite_measure.s_finite_kernel_const finite_measure.s_finite_measure_finite_measure finite_measureI) qed qed(auto simp: sets_ki) qed(auto simp: kernel_sets) qed lemma(in s_finite_kernel) nn_integral_measurable_f: assumes [measurable]:"(\(x,y). f x y) \ borel_measurable (X \\<^sub>M Y)" shows "(\x. \\<^sup>+y. f x y \(\ x)) \ borel_measurable X" proof - obtain \i where \i:"\i. subprob_kernel X Y (\i i)" "\x A. x \ space X \ \ x A = (\i. \i i x A)" by(metis s_finite_kernels) show ?thesis proof(rule measurable_cong[THEN iffD2]) fix x assume "x \ space X" with \i show "(\\<^sup>+ y. f x y \\ x) = (\i. \\<^sup>+ y. f x y \\i i x)" by(auto intro!: nn_integral_measure_suminf[symmetric] simp: subprob_kernel_def kernel_sets measure_kernel_def) next show "(\x. \i. \\<^sup>+ y. f x y \\i i x) \ borel_measurable X" using \i(1) nn_integral_measurable_subprob_algebra2[OF assms] by(simp add: subprob_kernel_def' ) qed qed lemma(in s_finite_kernel) nn_integral_measurable_f': assumes "f \ borel_measurable (X \\<^sub>M Y)" shows "(\x. \\<^sup>+y. f (x, y) \(\ x)) \ borel_measurable X" using nn_integral_measurable_f[where f="curry f",simplified,OF assms] by simp lemma(in s_finite_kernel) bind_kernel_s_finite_kernel': assumes "s_finite_kernel (X \\<^sub>M Y) Z (case_prod g)" shows "s_finite_kernel X Z (\x. \ x \\<^sub>k g x)" proof(cases "space X = {}") case True then show ?thesis by (simp add: s_finite_kernel_empty_trivial) next case X:False then have Y:"space Y \ {}" by(simp add: Y_not_empty) from s_finite_kernels obtain ki where ki: "\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" by metis interpret g:s_finite_kernel "X \\<^sub>M Y" Z "case_prod g" by fact from g.s_finite_kernels[simplified space_pair_measure] obtain gi where gi: "\i. subprob_kernel (X \\<^sub>M Y) Z (gi i)" "\x y A. x \ space X \ y \ space Y \ g x y A = (\i. gi i (x,y) A)" by auto metis define kgi where "kgi = (\i x. case prod_decode i of (i,j) \ (ki j x \ curry (gi i) x))" have emeasure:"emeasure (\ x \\<^sub>k g x) A = (\i. emeasure (kgi i x) A)" (is "?lhs = ?rhs") if x:"x \ space X" and A:"A \ sets Z" for x A proof - interpret gx: s_finite_kernel Y Z "g x" using g.comp_measurable[OF measurable_Pair1'[OF x]] by auto have "?lhs = (\\<^sup>+ y. g x y A \\ x)" using gx.emeasure_bind_kernel[OF kernel_sets[OF x] A] by(auto simp: sets_eq_imp_space_eq[OF kernel_sets[OF x]] Y) also have "... = (\\<^sup>+ y. (\i. gi i (x, y) A) \\ x)" by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF kernel_sets[OF x]] gi(2)[OF x]) also have "... = (\i. \\<^sup>+ y. gi i (x, y) A \\ x)" using gi(1) x A by(auto intro!: nn_integral_suminf simp: subprob_kernel_def') also have "... = (\i. (\j. \\<^sup>+ y. gi i (x, y) A \ki j x))" by(rule suminf_cong, rule nn_integral_measure_suminf[symmetric], insert kernel_sets[OF x] ki gi(1) x A) (auto simp: subprob_kernel_def measure_kernel_def measurable_cong_sets[OF sets_pair_measure_cong[OF refl kernel_sets[OF x]]] intro!: measurable_Pair2[OF _ x]) also have "... = (\i. (\j. emeasure (ki j x \ (curry (gi i) x)) A))" using sets_eq_imp_space_eq[of "ki _ x" Y] ki(1) x gi(1) measurable_cong_sets[of _ _ "subprob_algebra Z" "subprob_algebra Z", OF sets_pair_measure_cong[of X X Y "ki _ x"]] by(auto intro!: suminf_cong emeasure_bind[OF _ _ A,symmetric] measurable_Pair2[OF _ x] simp: curry_def subprob_kernel_def[of X] subprob_kernel_def'[of "X \\<^sub>M Y"] measure_kernel_def Y) also have "... = ?rhs" unfolding kgi_def by(rule suminf_ennreal_2dimen[symmetric]) (simp add: curry_def) finally show ?thesis . qed have sets: "sets (\ x \\<^sub>k g x) = sets Z" if x:"x \ space X" for x proof - interpret gx: s_finite_kernel Y Z "g x" using g.comp_measurable[OF measurable_Pair1'[OF x]] by auto show ?thesis by(simp add: gx.sets_bind_kernel[OF _ kernel_sets[OF x]] Y) qed have sk:"subprob_kernel X Z (kgi i)" for i using ki(1)[of "snd (prod_decode i)"] gi(1)[of "fst (prod_decode i)"] by(auto simp: subprob_kernel_def' kgi_def split_beta' curry_def) show ?thesis using sk by(auto simp: s_finite_kernel_def' emeasure sets subprob_kernel_def' intro!: exI[where x=kgi] measurable_cong[where g="\x. \i. emeasure (kgi i x) _" and f="\x. emeasure (\ x \\<^sub>k g x) _",THEN iffD2]) qed corollary(in s_finite_kernel) bind_kernel_s_finite_kernel: assumes "s_finite_kernel Y Z k'" shows "s_finite_kernel X Z (\x. \ x \\<^sub>k k')" by(auto intro!: bind_kernel_s_finite_kernel' s_finite_kernel.comp_measurable[OF assms measurable_snd] simp: split_beta') lemma(in s_finite_kernel) nn_integral_bind_kernel: assumes "f \ borel_measurable Y" "sets \ = sets X" shows "(\\<^sup>+ y. f y \(\ \\<^sub>k \)) = (\\<^sup>+x. (\\<^sup>+ y. f y \(\ x)) \\)" proof(cases "space X = {}") case True then show ?thesis by(simp add: sets_eq_imp_space_eq[OF assms(2)] bind_kernel_def nn_integral_empty) next case X:False then have \:"space \ \ {}" by(simp add: sets_eq_imp_space_eq[OF assms(2)]) note 1[measurable_cong] = assms(2) sets_bind_kernel[OF X assms(2)] from assms(1) show ?thesis proof induction case ih:(cong f g) have "(\\<^sup>+ y. f y \(\ \\<^sub>k \)) = (\\<^sup>+ y. g y \(\ \\<^sub>k \))" "(\\<^sup>+ x. integral\<^sup>N (\ x) f \\) = (\\<^sup>+ x. integral\<^sup>N (\ x) g \\)" by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF 1(2)] sets_eq_imp_space_eq[OF assms(2)] sets_eq_imp_space_eq[OF kernel_sets] ih(3)) then show ?case by(simp add: ih) next case (set A) then show ?case by(auto simp: emeasure_bind_kernel[OF 1(1) _ X] sets_eq_imp_space_eq[OF 1(1)] intro!: nn_integral_cong) next case ih:(mult u c) then have "(\\<^sup>+ x. \\<^sup>+ y. c * u y \\ x \\) = (\\<^sup>+ x. c * \\<^sup>+ y. u y \\ x \\)" by(auto intro!: nn_integral_cong nn_integral_cmult simp: sets_eq_imp_space_eq[OF 1(1)]) with ih nn_integral_measurable_f[of "\_ y. u y"] show ?case by(auto simp: nn_integral_cmult intro!: nn_integral_cong) next case ih:(add u v) then have "(\\<^sup>+ x. \\<^sup>+ y. v y + u y \\ x \\) = (\\<^sup>+ x. (\\<^sup>+ y. v y \\ x) + (\\<^sup>+ y. u y \\ x) \\)" by(auto intro!: nn_integral_cong simp: nn_integral_add sets_eq_imp_space_eq[OF 1(1)]) with ih nn_integral_measurable_f[of "\_ y. u y"] nn_integral_measurable_f[of "\_ y. v y"] show ?case by(simp add: nn_integral_add) next case ih[measurable]:(seq U) show ?case (is "?lhs = ?rhs") proof - have "?lhs = ((\i. integral\<^sup>N (\ \\<^sub>k \) (U i)))" by(rule nn_integral_monotone_convergence_SUP[of U,simplified SUP_apply[of U UNIV,symmetric]]) (use ih in auto) also have "... = (\i. \\<^sup>+ x. (\\<^sup>+ y. U i y \\ x) \\)" by(simp add: ih) also have "... = (\\<^sup>+ x. (\i. (\\<^sup>+ y. U i y \\ x)) \\)" proof(rule nn_integral_monotone_convergence_SUP[symmetric]) show "incseq (\i x. \\<^sup>+ y. U i y \\ x)" by standard+ (auto intro!: le_funI nn_integral_mono simp:le_funD[OF incseqD[OF ih(3)]]) qed(use nn_integral_measurable_f[of "\_ y. U _ y"] in simp) also have "... = ?rhs" by(rule nn_integral_cong, rule nn_integral_monotone_convergence_SUP[of U,simplified SUP_apply[of U UNIV,symmetric],OF ih(3),symmetric]) (auto simp: sets_eq_imp_space_eq[OF 1(1)]) finally show ?thesis . qed qed qed lemma(in s_finite_kernel) bind_kernel_assoc: assumes "s_finite_kernel Y Z k'" "sets \ = sets X" shows "\ \\<^sub>k (\x. \ x \\<^sub>k k') = \ \\<^sub>k \ \\<^sub>k k'" proof(cases "space X = {}") case X:False then have \: "space \ \ {}" and Y:"space Y \ {}" by(simp_all add: Y_not_empty sets_eq_imp_space_eq[OF assms(2)]) interpret k':s_finite_kernel Y Z k' by fact interpret k'': s_finite_kernel X Z "\x. \ x \\<^sub>k k'" by(rule bind_kernel_s_finite_kernel[OF assms(1)]) show ?thesis proof(rule measure_eqI) fix A assume "A \ sets (\ \\<^sub>k (\x. \ x \\<^sub>k k'))" then have A[measurable]: "A \ sets Z" by(simp add: k''.sets_bind_kernel[OF X assms(2)]) show "emeasure (\ \\<^sub>k (\x. \ x \\<^sub>k k')) A = emeasure (\ \\<^sub>k \ \\<^sub>k k') A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. emeasure (\ x \\<^sub>k k') A \\)" by(rule k''.emeasure_bind_kernel[OF assms(2) A X]) also have "... = (\\<^sup>+ x. \\<^sup>+ y. k' y A \\ x \\)" using k'.emeasure_bind_kernel[OF kernel_sets A] by(auto intro!: nn_integral_cong simp: sets_eq_imp_space_eq[OF assms(2)] sets_eq_imp_space_eq[OF kernel_sets] Y) also have "... = (\\<^sup>+ y. k' y A \(\ \\<^sub>k \))" by(simp add: nn_integral_bind_kernel[OF k'.emeasure_measurable[OF A] assms(2)]) also have "... = ?rhs" by(simp add: k'.emeasure_bind_kernel[OF sets_bind_kernel[OF X assms(2)] A] sets_eq_imp_space_eq[OF sets_bind_kernel[OF X assms(2)]] Y) finally show ?thesis . qed qed(auto simp: k'.sets_bind_kernel[OF Y sets_bind_kernel[OF X assms(2)]] k''.sets_bind_kernel[OF X assms(2)]) qed(simp add: bind_kernel_def sets_eq_imp_space_eq[OF assms(2)]) lemma s_finite_kernel_pair_measure: assumes "s_finite_kernel X Y k" and "s_finite_kernel X Z k'" shows "s_finite_kernel X (Y \\<^sub>M Z) (\x. k x \\<^sub>M k' x)" proof - interpret k: s_finite_kernel X Y k by fact interpret k': s_finite_kernel X Z k' by fact from k.s_finite_kernels k'.s_finite_kernels obtain ki ki' where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ k x A = (\i. ki i x A)" and ki':"\i. subprob_kernel X Z (ki' i)" "\x A. x \ space X \ k' x A = (\i. ki' i x A)" by metis then have 1[measurable_cong]: "\x i. x \ space X \ sets (ki i x) = sets Y" "\x i. x \ space X \ sets (ki' i x) = sets Z" by(auto simp: subprob_kernel_def measure_kernel_def) define kki where "kki \ (\i x. (\(j,i). ki i x \\<^sub>M ki' j x) (prod_decode i))" have kki1: "\i. subprob_kernel X (Y \\<^sub>M Z) (kki i)" using ki(1) ki'(1) by(auto simp: subprob_kernel_def' kki_def split_beta intro!: measurable_pair_measure) have kki2: "(k x \\<^sub>M k' x) A = (\i. (kki i x) A)" (is "?lhs = ?rhs") if x:"x \ space X" and A[measurable]: "A \ sets (Y \\<^sub>M Z)" for x A proof - have "?lhs = (\\<^sup>+ y. \\<^sup>+ z. indicator A (y, z) \k' x \k x)" using x by(simp add: s_finite_measure.emeasure_pair_measure'[OF k'.image_s_finite_measure]) also have "... = (\\<^sup>+ y. (\j. \\<^sup>+ z. indicator A (y, z) \ki' j x) \k x)" using ki' x by(auto intro!: nn_integral_cong nn_integral_measure_suminf[symmetric] simp: sets_eq_imp_space_eq[OF k.kernel_sets[OF x]] subprob_kernel_def measure_kernel_def k'.kernel_sets) also have "... = (\j. \\<^sup>+ y. \\<^sup>+ z. indicator A (y, z) \ki' j x \k x)" using x by(auto intro!: nn_integral_suminf s_finite_measure.borel_measurable_nn_integral_fst' s_finite_kernel.image_s_finite_measure[OF subprob_kernel.s_finite_kernel_subprob_kernel[OF ki'(1)]]) also have "... = (\j. (\i. (\\<^sup>+ y. \\<^sup>+ z. indicator A (y, z) \ki' j x \ki i x)))" using x ki by(auto intro!: suminf_cong nn_integral_measure_suminf[symmetric] s_finite_measure.borel_measurable_nn_integral_fst' simp: k.kernel_sets[OF x] subprob_kernel_def measure_kernel_def s_finite_kernel.image_s_finite_measure[OF subprob_kernel.s_finite_kernel_subprob_kernel[OF ki'(1)]]) also have "... = (\j. (\i. (ki i x \\<^sub>M ki' j x) A))" using x by(auto simp: s_finite_measure.emeasure_pair_measure'[OF s_finite_kernel.image_s_finite_measure[OF subprob_kernel.s_finite_kernel_subprob_kernel[OF ki'(1)]]]) also have "... = ?rhs" unfolding kki_def by(rule suminf_ennreal_2dimen[symmetric]) auto finally show ?thesis . qed show ?thesis proof fix B assume [measurable]:"B \ sets (Y \\<^sub>M Z)" show "(\x. emeasure (k x \\<^sub>M k' x) B) \ borel_measurable X" by(rule measurable_cong[where g="\x. \i. (kki i x) B",THEN iffD2], insert kki1) (auto simp: subprob_kernel_def' kki2) qed(auto intro!: exI[where x=kki] simp: subprob_kernel.finite_kernel kki1 kki2 k.kernel_sets k'.kernel_sets space_pair_measure k.Y_not_empty k'.Y_not_empty) qed lemma pair_measure_eq_bind_s_finite: assumes "s_finite_measure \" "s_finite_measure \" shows "\ \\<^sub>M \ = \ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x,y)))" proof - consider "space \ = {}" | "space \ = {}" | "space \ \ {}" "space \ \ {}" by auto then show ?thesis proof cases case 1 then show ?thesis by(auto simp: bind_kernel_def space_pair_measure intro!: space_empty) next case 2 then have "\ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x, y))) = count_space {}" by(auto simp: bind_kernel_def space_empty) with 2 show ?thesis by(auto simp: space_pair_measure intro!: space_empty) next case 3 show ?thesis proof(intro measure_eqI sets_bind_kernel[OF _ 3(1),symmetric] sets_bind_kernel[OF _ 3(2)]) fix A assume A[measurable]: "A \ sets (\ \\<^sub>M \)" show "emeasure (\ \\<^sub>M \) A = emeasure (\ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x, y)))) A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. \\<^sup>+ y. return (\ \\<^sub>M \) (x, y) A \\ \\)" by(simp add: s_finite_measure.emeasure_pair_measure'[OF assms(2)]) also have "... = (\\<^sup>+ x. (\ \\<^sub>k (\y. return (\ \\<^sub>M \) (x,y))) A \\)" by(auto intro!: nn_integral_cong measure_kernel.emeasure_bind_kernel[OF _ _ A 3(2),symmetric] prob_kernel.axioms(1) simp: prob_kernel_def' simp del: emeasure_return) also have "... = ?rhs" by(auto intro!: measure_kernel.emeasure_bind_kernel[OF _ _ A 3(1),symmetric] s_finite_kernel.axioms(1) s_finite_kernel.bind_kernel_s_finite_kernel'[where Y=\] s_finite_measure.s_finite_kernel_const[OF assms(2) 3(2)] prob_kernel.s_finite_kernel_prob_kernel[of "\ \\<^sub>M \"] simp: prob_kernel_def') finally show ?thesis . qed qed simp qed qed lemma bind_kernel_rotate_return: assumes "s_finite_measure \" "s_finite_measure \" shows "\ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x,y))) = \ \\<^sub>k (\y. \ \\<^sub>k (\x. return (\ \\<^sub>M \) (x,y)))" proof - consider "space \ = {}" | "space \ = {}" | "space \ \ {}" "space \ \ {}" by auto then show ?thesis proof cases case 1 then have "\ \\<^sub>k (\y. \ \\<^sub>k (\x. return (\ \\<^sub>M \) (x,y))) = count_space {}" by(auto simp: bind_kernel_def space_empty) then show ?thesis by(auto simp: bind_kernel_def space_pair_measure 1 intro!: space_empty) next case 2 then have "\ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x, y))) = count_space {}" by(auto simp: bind_kernel_def space_empty) with 2 show ?thesis by(auto simp: space_pair_measure bind_kernel_def intro!: space_empty) next case 3 show ?thesis unfolding pair_measure_eq_bind_s_finite[OF assms,symmetric] proof(intro measure_eqI) fix A assume A[measurable]:"A \ sets (\ \\<^sub>M \)" show "emeasure (\ \\<^sub>M \) A = emeasure (\ \\<^sub>k (\y. \ \\<^sub>k (\x. return (\ \\<^sub>M \) (x, y)))) A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ \\)" by(rule s_finite_measure.emeasure_pair_measure'[OF assms(2) A]) also have "... = (\\<^sup>+ y. \\<^sup>+ x. return (\ \\<^sub>M \) (x, y) A \\ \\)" by(simp add: nn_integral_snd'[OF assms] s_finite_measure.nn_integral_fst'[OF assms(2)]) also have "... = (\\<^sup>+ y. (\ \\<^sub>k (\x. return (\ \\<^sub>M \) (x, y))) A \\)" by(auto intro!: nn_integral_cong measure_kernel.emeasure_bind_kernel[OF _ _ A 3(1),symmetric] prob_kernel.axioms(1) simp add: prob_kernel_def' simp del: emeasure_return) also have "... = ?rhs" by(auto intro!: measure_kernel.emeasure_bind_kernel[OF _ _ A 3(2),symmetric] s_finite_kernel.axioms(1) s_finite_kernel.bind_kernel_s_finite_kernel'[where Y=\] s_finite_measure.s_finite_kernel_const[OF assms(1) 3(1)] prob_kernel.s_finite_kernel_prob_kernel[of "\ \\<^sub>M \"] simp: prob_kernel_def') finally show ?thesis . qed qed(auto intro!: sets_bind_kernel[OF _ 3(2),symmetric] sets_bind_kernel[OF _ 3(1)]) qed qed lemma bind_kernel_rotate': assumes "s_finite_measure \" "s_finite_measure \" "s_finite_kernel (\ \\<^sub>M \) Z (case_prod f)" shows "\ \\<^sub>k (\x. \ \\<^sub>k (\y. f x y)) = \ \\<^sub>k (\y. \ \\<^sub>k (\x. f x y))" (is "?lhs = ?rhs") proof - interpret sk: s_finite_kernel "\ \\<^sub>M \" Z "case_prod f" by fact consider "space \ = {}" | "space \ = {}" | "space \ \ {}" "space \ \ {}" by auto then show ?thesis proof cases case 1 then have "?rhs = count_space {}" by(auto simp: bind_kernel_def space_empty) then show ?thesis by(auto simp: bind_kernel_def space_pair_measure 1 intro!: space_empty) next case 2 then show ?thesis by(auto simp: space_pair_measure bind_kernel_def intro!: space_empty) next case 3 show ?thesis proof - have "?lhs = \ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x, y)) \\<^sub>k case_prod f)" by(auto intro!: bind_kernel_cong_All simp: s_finite_kernel.bind_kernel_assoc[OF prob_kernel.s_finite_kernel_prob_kernel assms(3) refl,of \ "\y. return (\ \\<^sub>M \) (_, y)",simplified prob_kernel_def',symmetric] sk.bind_kernel_return space_pair_measure) also have "... = \ \\<^sub>k (\x. \ \\<^sub>k (\y. return (\ \\<^sub>M \) (x,y))) \\<^sub>k (case_prod f)" by(auto simp: s_finite_kernel.bind_kernel_assoc[OF s_finite_kernel.bind_kernel_s_finite_kernel'[OF s_finite_measure.s_finite_kernel_const[OF assms(2) 3(2),of \] prob_kernel.s_finite_kernel_prob_kernel,of "\ \\<^sub>M \" "\x y. return (\ \\<^sub>M \) (x,y)",simplified] assms(3) refl, simplified prob_kernel_def',symmetric]) also have "... = \ \\<^sub>k (\y. \ \\<^sub>k (\x. return (\ \\<^sub>M \) (x,y))) \\<^sub>k (case_prod f)" by(simp add: bind_kernel_rotate_return assms) also have "... = \ \\<^sub>k (\y. \ \\<^sub>k (\x. return (\ \\<^sub>M \) (x, y)) \\<^sub>k case_prod f)" by(auto intro!: s_finite_kernel.bind_kernel_assoc[OF _ assms(3),symmetric] s_finite_kernel.bind_kernel_s_finite_kernel'[OF s_finite_measure.s_finite_kernel_const[OF assms(1) 3(1)]] prob_kernel.s_finite_kernel_prob_kernel[of "\ \\<^sub>M \"] simp: prob_kernel_def') also have "... = ?rhs" by(auto intro!: bind_kernel_cong_All simp: s_finite_kernel.bind_kernel_assoc[OF prob_kernel.s_finite_kernel_prob_kernel assms(3) refl,of \ "\x. return (\ \\<^sub>M \) (x, _)",simplified prob_kernel_def',symmetric] sk.bind_kernel_return space_pair_measure) finally show ?thesis . qed qed qed lemma bind_kernel_rotate: assumes "sets \ = sets X" and "sets \ = sets Y" and "s_finite_measure \" "s_finite_measure \" "s_finite_kernel (X \\<^sub>M Y) Z (\(x,y). f x y)" shows "\ \\<^sub>k (\x. \ \\<^sub>k (\y. f x y)) = \ \\<^sub>k (\y. \ \\<^sub>k (\x. f x y))" by(auto intro!: bind_kernel_rotate' assms simp: s_finite_kernel_cong_sets[OF sets_pair_measure_cong[OF assms(1,2)]]) lemma(in s_finite_kernel) emeasure_measurable': assumes A[measurable]: "(SIGMA x:space X. A x) \ sets (X \\<^sub>M Y)" shows "(\x. emeasure (\ x) (A x)) \ borel_measurable X" proof - have **: "A x \ sets Y" if "x \ space X" for x proof - have "Pair x -` Sigma (space X) A = A x" using that by auto with sets_Pair1[OF A, of x] show "A x \ sets Y" by auto qed have *: "\x. fst x \ space X \ snd x \ A (fst x) \ x \ (SIGMA x:space X. A x)" by (auto simp: fun_eq_iff) have "(\(x, y). indicator (A x) y::ennreal) \ borel_measurable (X \\<^sub>M Y)" by (measurable,subst measurable_cong[OF *]) (auto simp: space_pair_measure) then have "(\x. integral\<^sup>N (\ x) (indicator (A x))) \ borel_measurable X" by(rule nn_integral_measurable_f) moreover have "integral\<^sup>N (\ x) (indicator (A x)) = emeasure (\ x) (A x)" if "x \ space X" for x using **[OF that] kernel_sets[OF that] by(auto intro!: nn_integral_indicator) ultimately show "(\x. emeasure (\ x) (A x)) \ borel_measurable X" by(auto cong: measurable_cong) qed lemma(in s_finite_kernel) measure_measurable': assumes "(SIGMA x:space X. A x) \ sets (X \\<^sub>M Y)" shows "(\x. measure (\ x) (A x)) \ borel_measurable X" using emeasure_measurable'[OF assms] by(simp add: measure_def) lemma(in s_finite_kernel) AE_pred: assumes P[measurable]:"Measurable.pred (X \\<^sub>M Y) (case_prod P)" shows "Measurable.pred X (\x. AE y in \ x. P x y)" proof - have [measurable]:"Measurable.pred X (\x. emeasure (\ x) {y \ space Y. \ P x y} = 0)" proof(rule pred_eq_const1[where N=borel],rule emeasure_measurable') have "(SIGMA x:space X. {y \ space Y. \ P x y}) = {xy\space (X \\<^sub>M Y). \ P (fst xy) (snd xy)}" by (auto simp: space_pair_measure) also have "... \ sets (X \\<^sub>M Y)" by simp finally show "(SIGMA x:space X. {y \ space Y. \ P x y}) \ sets (X \\<^sub>M Y)" . qed simp have "{x \ space X. almost_everywhere (\ x) (P x)} = {x \ space X. emeasure (\ x) {y\space Y. \ P x y} = 0}" proof safe fix x assume x:"x \ space X" show "(AE y in \ x. P x y) \ emeasure (\ x) {y \ space Y. \ P x y} = 0" using emeasure_eq_0_AE[of "\y. \ P x y" "\ x"] by(simp add: sets_eq_imp_space_eq[OF kernel_sets[OF x]]) show "emeasure (\ x) {y \ space Y. \ P x y} = 0 \ almost_everywhere (\ x) (P x)" using x by(auto intro!: AE_I[where N="{y \ space Y. \ P x y}"] simp: sets_eq_imp_space_eq[OF kernel_sets[OF x]] kernel_sets[OF x]) qed also have "... \ sets X" by(simp add: pred_def) finally show ?thesis by(simp add: pred_def) qed lemma(in subprob_kernel) integrable_probability_kernel_pred: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]:"(\(x,y). f x y) \ borel_measurable (X \\<^sub>M Y)" shows "Measurable.pred X (\x. integrable (\ x) (f x))" proof(rule measurable_cong[THEN iffD2]) show "x \ space X \ integrable (\ x) (f x) \ (\\<^sup>+y. norm (f x y) \(\ x)) < \" for x by(auto simp: integrable_iff_bounded) next have "(\(x,y). ennreal (norm (f x y))) \ borel_measurable (X \\<^sub>M Y)" by measurable from nn_integral_measurable_f[OF this] show "Measurable.pred X (\x. (\\<^sup>+ y. ennreal (norm (f x y)) \\ x) < \)" by simp qed corollary integrable_measurable_subprob': fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]:"(\(x,y). f x y) \ borel_measurable (X \\<^sub>M Y)" "k \ X \\<^sub>M subprob_algebra Y" shows "Measurable.pred X (\x. integrable (k x) (f x))" by(auto intro!: subprob_kernel.integrable_probability_kernel_pred[where Y=Y] simp: subprob_kernel_def') lemma(in subprob_kernel) integrable_probability_kernel_pred': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "f \ borel_measurable (X \\<^sub>M Y)" shows "Measurable.pred X (\x. integrable (\ x) (curry f x))" using integrable_probability_kernel_pred[of "curry f"] assms by auto lemma(in subprob_kernel) lebesgue_integral_measurable_f_subprob: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes [measurable]:"f \ borel_measurable (X \\<^sub>M Y)" shows "(\x. \y. f (x,y) \(\ x)) \ borel_measurable X" proof - from borel_measurable_implies_sequence_metric[OF assms, of 0] obtain s where s: "\i. simple_function (X \\<^sub>M Y) (s i)" and "\x\space (X \\<^sub>M Y). (\i. s i x) \ f x" and "\i. \x\space (X \\<^sub>M Y). dist (s i x) 0 \ 2 * dist (f x) 0" by auto then have *: "\x y. x \ space X \ y \ space Y \ (\i. s i (x, y)) \ f (x,y)" "\i x y. x \ space X \ y \ space Y \ norm (s i (x, y)) \ 2 * norm (f (x, y))" by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (X \\<^sub>M Y)" by (rule borel_measurable_simple_function) fact have s':"\i. s i \ X \\<^sub>M Y \\<^sub>M count_space UNIV" by (rule measurable_simple_function) fact define f' where [abs_def]: "f' i x = (if integrable (\ x) (curry f x) then Bochner_Integration.simple_bochner_integral (\ x) (\y. s i (x, y)) else 0)" for i x have eq: "Bochner_Integration.simple_bochner_integral (\ x) (\y. s i (x, y)) = (\z\s i ` (space X \ space Y). measure (\ x) {y \ space (\ x). s i (x, y) = z} *\<^sub>R z)" if "x \ space X" for x i proof - have [measurable_cong]: "sets (\ x) = sets Y" and [simp]: "space (\ x) = space Y" using that by (simp_all add: kernel_sets kernel_space) with that show ?thesis using s[THEN simple_functionD(1)] unfolding simple_bochner_integral_def by (intro sum.mono_neutral_cong_left) (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) qed show ?thesis proof (rule borel_measurable_LIMSEQ_metric) fix i note [measurable] = integrable_probability_kernel_pred'[OF assms] have [measurable]:"(SIGMA x:space X. {y \ space Y. s i (x, y) = s i (a, b)}) \ sets (X \\<^sub>M Y)" for a b proof - have "(SIGMA x:space X. {y \ space Y. s i (x, y) = s i (a, b)}) = space (X \\<^sub>M Y) \ s i -` {s i (a,b)}" by(auto simp: space_pair_measure) thus ?thesis using s'[of i] by simp qed show "f' i \ borel_measurable X" by (auto simp : eq kernel_space f'_def cong: measurable_cong if_cong intro!: borel_measurable_sum measurable_If borel_measurable_scaleR measure_measurable') next fix x assume x:"x \ space X" have "(\i. Bochner_Integration.simple_bochner_integral (\ x) (\y. s i (x, y))) \ (\y. f (x,y) \(\ x))" if int_f:"integrable (\ x) (curry f x)" proof - have int_2f: "integrable (\ x) (\y. 2 * norm (f (x,y)))" using int_f by(auto simp: curry_def) have "(\i. integral\<^sup>L (\ x) (\y. s i (x, y))) \ integral\<^sup>L (\ x) (curry f x)" proof (rule integral_dominated_convergence) show "curry f x \ borel_measurable (\ x)" using int_f by auto next show "\i. (\y. s i (x, y)) \ borel_measurable (\ x)" using x kernel_sets by auto next show "AE xa in \ x. (\i. s i (x, xa)) \ curry f x xa" using x *(1) kernel_space by(auto simp: curry_def) next show "\i. AE xa in \ x. norm (s i (x, xa)) \ 2 * norm (f (x,xa))" using x * (2) kernel_space by auto qed fact moreover have "integral\<^sup>L (\ x) (\y. s i (x, y)) = Bochner_Integration.simple_bochner_integral (\ x) (\y. s i (x, y))" for i proof - have "Bochner_Integration.simple_bochner_integrable (\ x) (\y. s i (x, y))" proof (rule simple_bochner_integrableI_bounded) have "(\y. s i (x, y)) ` space Y \ s i ` (space X \ space Y)" using x by auto then show "simple_function (\ x) (\y. s i (x, y))" using simple_functionD(1)[OF s(1), of i] x kernel_space by (intro simple_function_borel_measurable) (auto simp: space_pair_measure dest: finite_subset) next have "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \\ x) \ (\\<^sup>+ y. 2 * norm (f (x,y)) \\ x)" using x *(2) kernel_space by (intro nn_integral_mono) auto also have "... < \" using int_2f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \\ x) < \" . qed then show ?thesis by (rule simple_bochner_integrable_eq_integral[symmetric]) qed ultimately show ?thesis by(simp add: curry_def) qed thus "(\i. f' i x) \ (\y. f (x,y) \(\ x))" by (cases "integrable (\ x) (curry f x)") (simp_all add: f'_def not_integrable_integral_eq curry_def) qed qed lemma(in s_finite_kernel) integrable_measurable_pred[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]:"case_prod f \ borel_measurable (X \\<^sub>M Y)" shows "Measurable.pred X (\x. integrable (\ x) (f x))" proof(cases "space X = {}") case True from space_empty[OF this] show ?thesis by simp next case h:False obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" using s_finite_kernels by metis have [simp]:"integrable (\ x) (f x) = ((\i. \\<^sup>+ y. ennreal (norm (f x y)) \ki i x) < \)" if "x \ space X" for x using ki(1) nn_integral_measure_suminf[of "\i. ki i x" "\ x",OF _ ki(2)] that kernel_sets by(auto simp: integrable_iff_bounded subprob_kernel_def measure_kernel_def) note [measurable] = nn_integral_measurable_subprob_algebra2 show ?thesis by(rule measurable_cong[where g="\x. (\i. \\<^sup>+y. ennreal (norm (f x y)) \(ki i x)) < \",THEN iffD2]) (insert ki(1), auto simp: subprob_kernel_def') qed lemma(in s_finite_kernel) integral_measurable_f: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]:"case_prod f \ borel_measurable (X \\<^sub>M Y)" shows "(\x. \y. f x y \(\ x)) \ borel_measurable X" proof - obtain ki where ki:"\i. subprob_kernel X Y (ki i)" "\x A. x \ space X \ \ x A = (\i. ki i x A)" using s_finite_kernels by metis note [measurable] = integral_measurable_subprob_algebra2 show ?thesis proof(rule measurable_cong[where f="(\x. if integrable (\ x) (f x) then (\i. \y. f x y \(ki i x)) else 0)",THEN iffD1]) fix x assume h:"x \ space X" { assume h':"integrable (\ x) (f x)" have "(\i. \y. f x y \(ki i x)) = (\y. f x y \(\ x))" using lebesgue_integral_measure_suminf[of "\i. ki i x" "\ x",OF _ ki(2) h'] ki(1) kernel_sets[OF h] h by(auto simp: subprob_kernel_def measure_kernel_def) } thus "(if integrable (\ x) (f x) then (\i. \y. f x y \(ki i x)) else 0) = (\y. f x y \(\ x))" using not_integrable_integral_eq by auto qed(insert ki(1), auto simp: subprob_kernel_def') qed lemma(in s_finite_kernel) integral_measurable_f': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes [measurable]:"f \ borel_measurable (X \\<^sub>M Y)" shows "(\x. \y. f (x,y) \(\ x)) \ borel_measurable X" using integral_measurable_f[of "curry f"] by simp lemma(in s_finite_kernel) fixes f :: "_ \ _::{banach, second_countable_topology}" assumes [measurable_cong]: "sets \ = sets X" and "integrable (\ \\<^sub>k \) f" shows integrable_bind_kernelD1: "integrable \ (\x. \y. norm (f y) \\ x)" (is ?g1) and integrable_bind_kernelD1': "integrable \ (\x. \y. f y \\ x)" (is ?g1') and integrable_bind_kernelD2: "AE x in \. integrable (\ x) f" (is ?g2) and integrable_bind_kernelD3: "space X \ {} \ f \ borel_measurable Y" (is "_ \ ?g3") proof - show 1:"space X \ {} \ ?g3" using assms(2) sets_bind_kernel[OF _ assms(1)] by(simp add: integrable_iff_bounded cong: measurable_cong_sets) have "integrable \ (\x. \y. norm (f y) \\ x) \ integrable \ (\x. \y. f y \\ x) \ (AE x in \. integrable (\ x) f)" proof(cases "space X = {}") assume ne: "space X \ {}" then have "space \ \ {}" by(simp add: sets_eq_imp_space_eq[OF assms(1)]) note h = integral_measurable_f[measurable] sets_bind_kernel[OF ne assms(1),measurable_cong] have g2: ?g2 unfolding integrable_iff_bounded AE_conj_iff proof safe show "AE x in \. f \ borel_measurable (\ x)" using assms(2) by(auto simp: sets_eq_imp_space_eq[OF assms(1)] measurable_cong_sets[OF kernel_sets]) next note nn_integral_measurable_f[measurable] have "AE x in \. (\\<^sup>+ x. ennreal (norm (f x)) \\ x) \ \" by(rule nn_integral_PInf_AE,insert assms(2)) (auto simp: integrable_iff_bounded nn_integral_bind_kernel[OF _ assms(1)] intro!: ) thus "AE x in \. (\\<^sup>+ x. ennreal (norm (f x)) \\ x) < \" by (simp add: top.not_eq_extremum) qed have [simp]:"(\\<^sup>+ x. \\<^sup>+ x. ennreal (norm (f x)) \\ x \\) = (\\<^sup>+ x. ennreal (\y. norm (f y) \\ x)\\)" using g2 by(auto intro!: nn_integral_cong_AE simp: nn_integral_eq_integral) have g1: ?g1 using assms(2) by(auto simp: integrable_iff_bounded measurable_cong_sets[OF h(2)] measurable_cong_sets[OF assms(1)] nn_integral_bind_kernel[OF _ assms(1)]) have ?g1' using assms(2) by(auto intro!: Bochner_Integration.integrable_bound[OF g1]) with g2 g1 show ?thesis by auto qed(auto simp: space_empty[of \] sets_eq_imp_space_eq[OF assms(1)] integrable_iff_bounded nn_integral_empty) thus ?g1 ?g1' ?g2 by auto qed lemma(in s_finite_kernel) fixes f :: "_ \ _::{banach, second_countable_topology}" assumes [measurable_cong]: "sets \ = sets X" and [measurable]:"AE x in \. integrable (\ x) f" "integrable \ (\x. \y. norm (f y) \\ x)" "f \ borel_measurable Y" shows integrable_bind_kernel: "integrable (\ \\<^sub>k \) f" and integral_bind_kernel: "(\y. f y \(\ \\<^sub>k \)) = (\x. (\y. f y\\ x)\ \)" (is ?eq) proof - have "integrable (\ \\<^sub>k \) f \ (\y. f y \(\ \\<^sub>k \)) = (\x. (\y. f y\\ x)\ \)" proof(cases "space X = {}") assume ne: "space X \ {}" note sets_bind[measurable_cong] = sets_bind_kernel[OF ne assms(1)] note h = integral_measurable_f[measurable] have 1:"integrable (\ \\<^sub>k \) f" unfolding integrable_iff_bounded proof show "(\\<^sup>+ x. ennreal (norm (f x)) \(\ \\<^sub>k \)) < \" (is "?l < _") proof - have "?l = (\\<^sup>+ x. ennreal (\y. norm (f y) \\ x)\\)" using assms(2) by(auto intro!: nn_integral_cong_AE simp: nn_integral_eq_integral simp: nn_integral_bind_kernel[OF _ assms(1)]) also have "... < \" using assms(3) by(auto simp: integrable_iff_bounded) finally show ?thesis . qed qed simp then have ?eq proof induction case h[measurable]:(base A c) hence 1:"integrable (\ \\<^sub>k \) (indicat_real A)" by simp have 2:"integrable \ (\x. measure (\ x) A)" by(rule Bochner_Integration.integrable_cong[where f="\x. Sigma_Algebra.measure (\ x) (A \ space (\ x))",THEN iffD1,OF refl]) (insert h integrable_bind_kernelD1[OF assms(1) 1] sets_eq_imp_space_eq[OF kernel_sets], auto simp: sets_eq_imp_space_eq[OF assms(1)] sets_eq_imp_space_eq[OF kernel_sets] sets_bind) have "AE x in \. emeasure (\ x) A \ \" by(rule nn_integral_PInf_AE,insert h) (auto simp: emeasure_bind_kernel[OF assms(1) _ ne] sets_bind) hence 0:"AE x in \. emeasure (\ x) A < \" by (simp add: top.not_eq_extremum) have "(\x. (\y. indicat_real A y *\<^sub>R c \\ x)\ \) = (\x. measure (\ x) A *\<^sub>R c\\)" using h integrable_bind_kernelD2[OF assms(1) integrable_real_indicator,of A] by(auto intro!: integral_cong_AE simp: sets_eq_imp_space_eq[OF kernel_sets] sets_bind sets_eq_imp_space_eq[OF assms(1)]) also have "... = (\x. measure (\ x) A \\) *\<^sub>R c" using 2 by(auto intro!: integral_scaleR_left) finally show ?case using h by(auto simp: measure_bind_kernel[OF assms(1) _ ne 0] sets_bind) next case ih:(add f g) show ?case using ih(1,2) integrable_bind_kernelD2[OF assms(1) ih(1)] integrable_bind_kernelD2[OF assms(1) ih(2)] by(auto simp: ih(3,4) Bochner_Integration.integral_add[OF integrable_bind_kernelD1'[OF assms(1) ih(1)] integrable_bind_kernelD1'[OF assms(1) ih(2)],symmetric] intro!: integral_cong_AE) next case ih:(lim f fn) show ?case (is "?lhs = ?rhs") proof - have conv: "AE x in \. (\n. \y. fn n y\\ x) \ (\y. f y \\ x)" proof - have conv:"AE x in \. integrable (\ x) f \ (\n. \y. fn n y\\ x) \ (\y. f y \\ x)" proof fix x assume h:"x \ space \" then show "integrable (\ x) f \ (\n. \y. fn n y\\ x) \ (\y. f y \\ x)" using ih by(auto intro!: integral_dominated_convergence[where w="\x. 2 * norm (f x)"] simp: sets_eq_imp_space_eq[OF sets_bind] sets_eq_imp_space_eq[OF kernel_sets[OF h[simplified sets_eq_imp_space_eq[OF assms(1)]]]] sets_eq_imp_space_eq[OF assms(1)]) qed with conv integrable_bind_kernelD2[OF assms(1) ih(4)] show ?thesis by fastforce qed have "?lhs = lim (\n. \y. fn n y \(\ \\<^sub>k \))" by(rule limI[OF integral_dominated_convergence[where w="\x. 2 * norm (f x)"],symmetric]) (use ih in auto) also have "... = lim (\n. (\x. (\y. fn n y\\ x)\ \))" by(simp add: ih) also have "... = (\x. lim (\n. \y. fn n y\\ x)\ \)" proof(rule limI[OF integral_dominated_convergence[where w="\x. \y. 2 * norm (f y) \\ x"]]) fix n show "AE x in \. norm (\y. fn n y\\ x) \ (\y. 2 * norm (f y) \\ x)" by(rule AE_mp[OF integrable_bind_kernelD2[OF assms(1) ih(1),of n] AE_mp[OF integrable_bind_kernelD2[OF assms(1) ih(4)]]],standard+,rule order.trans[OF integral_norm_bound integral_mono[of "\ _" "\y. norm (fn n y)" _,OF _ _ ih(3)[simplified sets_eq_imp_space_eq[OF sets_bind]]]]) (auto simp: sets_eq_imp_space_eq[OF assms(1)] sets_eq_imp_space_eq[OF kernel_sets]) qed(use ih integrable_bind_kernelD1[OF assms(1) ih(4)] conv limI in auto,fastforce) also have "... = ?rhs" using ih conv limI by(auto intro!: integral_cong_AE, blast) finally show ?thesis . qed qed with 1 show ?thesis by auto qed(auto simp: bind_kernel_def space_empty[of \] sets_eq_imp_space_eq[OF assms(1)] integrable_iff_bounded nn_integral_empty Bochner_Integration.integral_empty) thus "integrable (\ \\<^sub>k \) f" ?eq by auto qed end \ No newline at end of file diff --git a/thys/S_Finite_Measure_Monad/Measure_QuasiBorel_Adjunction.thy b/thys/S_Finite_Measure_Monad/Measure_QuasiBorel_Adjunction.thy --- a/thys/S_Finite_Measure_Monad/Measure_QuasiBorel_Adjunction.thy +++ b/thys/S_Finite_Measure_Monad/Measure_QuasiBorel_Adjunction.thy @@ -1,1532 +1,1577 @@ (* Title: Measure_QuasiBorel_Adjunction.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) subsection \Relation to Measurable Spaces\ theory Measure_QuasiBorel_Adjunction imports "QuasiBorel" "QBS_Morphism" Lemmas_S_Finite_Measure_Monad begin text \ We construct the adjunction between \textbf{Meas} and \textbf{QBS}, where \textbf{Meas} is the category of measurable spaces and measurable functions, and \textbf{QBS} is the category of quasi-Borel spaces and morphisms.\ subsubsection \ The Functor $R$ \ definition measure_to_qbs :: "'a measure \ 'a quasi_borel" where "measure_to_qbs X \ Abs_quasi_borel (space X, borel \\<^sub>M X)" +declare [[coercion measure_to_qbs]] + lemma shows qbs_space_R: "qbs_space (measure_to_qbs X) = space X" (is ?goal1) and qbs_Mx_R: "qbs_Mx (measure_to_qbs X) = borel \\<^sub>M X" (is ?goal2) proof - have "Rep_quasi_borel (measure_to_qbs X) = (space X, borel \\<^sub>M X)" by(auto intro!: Abs_quasi_borel_inverse is_quasi_borel_intro qbs_closed1I qbs_closed2I simp: measure_to_qbs_def dest:measurable_space) (rule qbs_closed3I, auto) thus ?goal1 ?goal2 by (simp_all add: qbs_space_def qbs_Mx_def) qed text \ The following lemma says that @{term measure_to_qbs} is a functor from \textbf{Meas} to \textbf{QBS}. \ lemma r_preserves_morphisms: "X \\<^sub>M Y \ (measure_to_qbs X) \\<^sub>Q (measure_to_qbs Y)" by(auto intro!: qbs_morphismI simp: qbs_Mx_R) +lemma measurable_imp_qbs_morphism: "f \ M \\<^sub>M N \ f \ M \\<^sub>Q N" + using r_preserves_morphisms by blast + subsubsection \ The Functor $L$ \ definition sigma_Mx :: "'a quasi_borel \ 'a set set" where "sigma_Mx X \ {U \ qbs_space X |U. \\\qbs_Mx X. \ -` U \ sets borel}" definition qbs_to_measure :: "'a quasi_borel \ 'a measure" where "qbs_to_measure X \ Abs_measure (qbs_space X, sigma_Mx X, \A. (if A = {} then 0 else if A \ - sigma_Mx X then 0 else \))" lemma measure_space_L: "measure_space (qbs_space X) (sigma_Mx X) (\A. (if A = {} then 0 else if A \ - sigma_Mx X then 0 else \))" unfolding measure_space_def proof safe show "sigma_algebra (qbs_space X) (sigma_Mx X)" proof(rule sigma_algebra.intro) show "algebra (qbs_space X) (sigma_Mx X)" proof have "\ U \ sigma_Mx X. U \ qbs_space X" using sigma_Mx_def subset_iff by fastforce thus "sigma_Mx X \ Pow (qbs_space X)" by auto next show "{} \ sigma_Mx X" unfolding sigma_Mx_def by auto next fix A fix B assume "A \ sigma_Mx X" "B \ sigma_Mx X" then have "\ Ua. A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by (simp add: sigma_Mx_def) then obtain Ua where pa:"A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by auto have "\ Ub. B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" using \B \ sigma_Mx X\ sigma_Mx_def by auto then obtain Ub where pb:"B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" by auto from pa pb have [simp]:"\\\qbs_Mx X. \ -` (Ua \ Ub) \ sets borel" by auto from this pa pb sigma_Mx_def have [simp]:"(Ua \ Ub) \ qbs_space X \ sigma_Mx X" by blast from pa pb have [simp]:"A \ B = (Ua \ Ub) \ qbs_space X" by auto thus "A \ B \ sigma_Mx X" by simp next fix A fix B assume "A \ sigma_Mx X" "B \ sigma_Mx X" then have "\ Ua. A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by (simp add: sigma_Mx_def) then obtain Ua where pa:"A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by auto have "\ Ub. B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" using \B \ sigma_Mx X\ sigma_Mx_def by auto then obtain Ub where pb:"B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" by auto from pa pb have [simp]:"A - B = (Ua \ -Ub) \ qbs_space X" by auto from pa pb have "\\\qbs_Mx X. \ -`(Ua \ -Ub) \ sets borel" by (metis Diff_Compl double_compl sets.Diff vimage_Compl vimage_Int) hence 1:"A - B \ sigma_Mx X" using sigma_Mx_def \A - B = Ua \ - Ub \ qbs_space X\ by blast show "\C\sigma_Mx X. finite C \ disjoint C \ A - B = \ C" proof show "{A - B} \sigma_Mx X \ finite {A-B} \ disjoint {A-B} \ A - B = \ {A-B}" using 1 by auto qed next fix A fix B assume "A \ sigma_Mx X" "B \ sigma_Mx X" then have "\ Ua. A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by (simp add: sigma_Mx_def) then obtain Ua where pa:"A = Ua \ qbs_space X \ (\\\qbs_Mx X. \ -` Ua \ sets borel)" by auto have "\ Ub. B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" using \B \ sigma_Mx X\ sigma_Mx_def by auto then obtain Ub where pb:"B = Ub \ qbs_space X \ (\\\qbs_Mx X. \ -` Ub \ sets borel)" by auto from pa pb have "A \ B = (Ua \ Ub) \ qbs_space X" by auto from pa pb have "\\\qbs_Mx X. \ -`(Ua \ Ub) \ sets borel" by auto then show "A \ B \ sigma_Mx X" unfolding sigma_Mx_def using \A \ B = (Ua \ Ub) \ qbs_space X\ by blast next have "\\\qbs_Mx X. \ -` (UNIV) \ sets borel" by simp thus "qbs_space X \ sigma_Mx X" unfolding sigma_Mx_def by blast qed next show "sigma_algebra_axioms (sigma_Mx X)" unfolding sigma_algebra_axioms_def proof safe fix A :: "nat \ _" assume 1:"range A \ sigma_Mx X" then have 2:"\i. \Ui. A i = Ui \ qbs_space X \ (\\\qbs_Mx X. \ -` Ui \ sets borel)" unfolding sigma_Mx_def by auto then have "\ U :: nat \ _. \i. A i = U i \ qbs_space X \ (\\\qbs_Mx X. \ -` (U i) \ sets borel)" by (rule choice) from this obtain U where pu:"\i. A i = U i \ qbs_space X \ (\\\qbs_Mx X. \ -` (U i) \ sets borel)" by auto hence "\\\qbs_Mx X. \ -` (\ (range U)) \ sets borel" by (simp add: countable_Un_Int(1) vimage_UN) from pu have "\ (range A) = (\i::nat. (U i \ qbs_space X))" by blast hence "\ (range A) = \ (range U) \ qbs_space X" by auto thus "\ (range A) \ sigma_Mx X" using sigma_Mx_def \\\\qbs_Mx X. \ -` \ (range U) \ sets borel\ by blast qed qed next show "countably_additive (sigma_Mx X) (\A. if A = {} then 0 else if A \ - sigma_Mx X then 0 else \)" proof(rule countably_additiveI) fix A :: "nat \ _" assume h:"range A \ sigma_Mx X" "\ (range A) \ sigma_Mx X" consider "\ (range A) = {}" | "\ (range A) \ {}" by auto then show "(\i. if A i = {} then 0 else if A i \ - sigma_Mx X then 0 else \) = (if \ (range A) = {} then 0 else if \ (range A) \ - sigma_Mx X then 0 else (\ :: ennreal))" proof cases case 1 then have "\i. A i = {}" by simp thus ?thesis by(simp add: 1) next case 2 then obtain j where hj:"A j \ {}" by auto have "(\i. if A i = {} then 0 else if A i \ - sigma_Mx X then 0 else \) = (\ :: ennreal)" proof - have hsum:"\N f. sum f {.. (\n. (f n :: ennreal))" by (simp add: sum_le_suminf) have hsum':"\P f. (\j. j \ P \ f j = (\ :: ennreal)) \ finite P \ sum f P = \" by auto have h1:"(\i - sigma_Mx X then 0 else \) = (\ :: ennreal)" proof(rule hsum') show "\ja. ja \ {.. (if A ja = {} then 0 else if A ja \ - sigma_Mx X then 0 else \) = (\ :: ennreal)" proof(rule exI[where x=j],rule conjI) have "A j \ sigma_Mx X" using h(1) by auto then show "(if A j = {} then 0 else if A j \ - sigma_Mx X then 0 else \) = (\ :: ennreal)" using hj by simp qed simp qed simp have "(\i - sigma_Mx X then 0 else \) \ (\i. if A i = {} then 0 else if A i \ - sigma_Mx X then 0 else (\ :: ennreal))" by(rule hsum) thus ?thesis by(simp only: h1) (simp add: top.extremum_unique) qed moreover have "(if \ (range A) = {} then 0 else if \ (range A) \ - sigma_Mx X then 0 else \) = (\ :: ennreal)" using 2 h(2) by simp ultimately show ?thesis by simp qed qed qed(simp add: positive_def) lemma shows space_L: "space (qbs_to_measure X) = qbs_space X" (is ?goal1) and sets_L: "sets (qbs_to_measure X) = sigma_Mx X" (is ?goal2) and emeasure_L: "emeasure (qbs_to_measure X) = (\A. if A = {} \ A \ sigma_Mx X then 0 else \)" (is ?goal3) proof - have "Rep_measure (qbs_to_measure X) = (qbs_space X, sigma_Mx X, \A. (if A = {} then 0 else if A \ - sigma_Mx X then 0 else \))" unfolding qbs_to_measure_def by(auto intro!: Abs_measure_inverse simp: measure_space_L) thus ?goal1 ?goal2 ?goal3 by(auto simp: sets_def space_def emeasure_def) qed lemma qbs_Mx_sigma_Mx_contra: assumes "qbs_space X = qbs_space Y" and "qbs_Mx X \ qbs_Mx Y" shows "sigma_Mx Y \ sigma_Mx X" using assms by(auto simp: sigma_Mx_def) text \ The following lemma says that @{term qbs_to_measure} is a functor from \textbf{QBS} to \textbf{Meas}. \ lemma l_preserves_morphisms: "X \\<^sub>Q Y \ (qbs_to_measure X) \\<^sub>M (qbs_to_measure Y)" proof safe fix f assume h:"f \ X \\<^sub>Q Y" show "f \ (qbs_to_measure X) \\<^sub>M (qbs_to_measure Y)" proof(rule measurableI) fix A assume "A \ sets (qbs_to_measure Y)" then obtain Ua where pa:"A = Ua \ qbs_space Y \ (\\\qbs_Mx Y. \ -` Ua \ sets borel)" by (auto simp: sigma_Mx_def sets_L) have "\\\qbs_Mx X. f \ \ \ qbs_Mx Y" "\\\ qbs_Mx X. \ -` (f -` (qbs_space Y)) = UNIV" using qbs_morphism_space[OF h] qbs_morphism_Mx[OF h] by (auto simp: qbs_Mx_to_X) hence "\\\qbs_Mx X. \ -` (f -` A) = \ -` (f -` Ua)" by (simp add: pa) from pa this qbs_morphism_def have "\\\qbs_Mx X. \ -` (f -` A) \ sets borel" by (simp add: vimage_comp \\\\qbs_Mx X. f \ \ \ qbs_Mx Y\) thus "f -` A \ space (qbs_to_measure X) \ sets (qbs_to_measure X)" using sigma_Mx_def by(auto simp: sets_L space_L) qed (insert qbs_morphism_space[OF h], auto simp: space_L) qed +lemma qbs_morphism_imp_measurable: "f \ X \\<^sub>Q Y \ f \ qbs_to_measure X \\<^sub>M qbs_to_measure Y" + using l_preserves_morphisms by blast abbreviation qbs_borel ("borel\<^sub>Q") where "borel\<^sub>Q \ measure_to_qbs borel" abbreviation qbs_count_space ("count'_space\<^sub>Q") where "qbs_count_space I \ measure_to_qbs (count_space I)" -declare [[coercion measure_to_qbs]] - lemma shows qbs_space_qbs_borel[simp]: "qbs_space borel\<^sub>Q = UNIV" and qbs_space_count_space[simp]: "qbs_space (qbs_count_space I) = I" and qbs_Mx_qbs_borel: "qbs_Mx borel\<^sub>Q = borel_measurable borel" and qbs_Mx_count_space: "qbs_Mx (qbs_count_space I) = borel \\<^sub>M count_space I" by(simp_all add: qbs_space_R qbs_Mx_R) (* Want to remove the following *) lemma shows qbs_space_qbs_borel'[qbs]: "r \ qbs_space borel\<^sub>Q" and qbs_space_count_space_UNIV'[qbs]: "x \ qbs_space (qbs_count_space (UNIV :: (_ :: countable) set))" by simp_all lemma qbs_Mx_is_morphisms: "qbs_Mx X = borel\<^sub>Q \\<^sub>Q X" proof safe fix \ :: "real \ _" assume "\ \ borel\<^sub>Q \\<^sub>Q X" have "id \ qbs_Mx borel\<^sub>Q" by (simp add: qbs_Mx_R) then have "\ \ id \ qbs_Mx X" using qbs_morphism_Mx[OF \\ \ borel\<^sub>Q \\<^sub>Q X\] by blast then show "\ \ qbs_Mx X" by simp qed(auto intro!: qbs_morphismI simp: qbs_Mx_qbs_borel) lemma exp_qbs_Mx': "qbs_Mx (exp_qbs X Y) = {g. case_prod g \ borel\<^sub>Q \\<^sub>Q X \\<^sub>Q Y}" by(auto simp: qbs_Mx_qbs_borel comp_def qbs_Mx_is_morphisms split_beta' intro!:curry_preserves_morphisms) lemma arg_swap_morphism': assumes "(\g. f (\w x. g x w)) \ exp_qbs X (exp_qbs W Y) \\<^sub>Q Z" shows "f \ exp_qbs W (exp_qbs X Y) \\<^sub>Q Z" proof(rule qbs_morphismI) fix \ assume "\ \ qbs_Mx (exp_qbs W (exp_qbs X Y))" then have "(\((r,w),x). \ r w x) \ (borel\<^sub>Q \\<^sub>Q W) \\<^sub>Q X \\<^sub>Q Y" by(auto simp: qbs_Mx_is_morphisms dest: uncurry_preserves_morphisms) hence "(\(r,w,x). \ r w x) \ borel\<^sub>Q \\<^sub>Q W \\<^sub>Q X \\<^sub>Q Y" by(auto intro!: qbs_morphism_cong'[where f="(\((r,w),x). \ r w x) \ (\(x, y, z). ((x, y), z))" and g="\(r,w,x). \ r w x"] qbs_morphism_comp[OF qbs_morphism_pair_assoc2]) hence "(\(r,x,w). \ r w x) \ borel\<^sub>Q \\<^sub>Q X \\<^sub>Q W \\<^sub>Q Y" by(auto intro!: qbs_morphism_cong'[where f="(\(r,w,x). \ r w x) \ map_prod id (\(x,y). (y,x))" and g="(\(r,x,w). \ r w x)"] qbs_morphism_comp qbs_morphism_map_prod qbs_morphism_pair_swap) hence "(\((r,x),w). \ r w x) \ (borel\<^sub>Q \\<^sub>Q X) \\<^sub>Q W \\<^sub>Q Y" by(auto intro!: qbs_morphism_cong'[where f="(\(r,x,w). \ r w x) \ (\((x, y), z). (x, y, z))" and g="\((r,x),w). \ r w x"] qbs_morphism_comp[OF qbs_morphism_pair_assoc1]) hence "(\r x w. \ r w x) \ qbs_Mx (exp_qbs X (exp_qbs W Y))" by(auto simp: qbs_Mx_is_morphisms split_beta') from qbs_morphism_Mx[OF assms this] show "f \ \ \ qbs_Mx Z" by(auto simp: comp_def) qed lemma qbs_Mx_subset_of_measurable: "qbs_Mx X \ borel \\<^sub>M qbs_to_measure X" proof fix \ assume "\ \ qbs_Mx X" show "\ \ borel \\<^sub>M qbs_to_measure X" proof(rule measurableI) fix x show "\ x \ space (qbs_to_measure X)" using qbs_Mx_to_X \\ \ qbs_Mx X\ by(simp add: space_L) next fix A assume "A \ sets (qbs_to_measure X)" then have "\ -`(qbs_space X) = UNIV" using \\ \ qbs_Mx X\ qbs_Mx_to_X by(auto simp: sets_L) then show "\ -` A \ space borel \ sets borel" using \\ \ qbs_Mx X\ \A \ sets (qbs_to_measure X)\ by(auto simp add: sigma_Mx_def sets_L) qed qed lemma L_max_of_measurables: assumes "space M = qbs_space X" and "qbs_Mx X \ borel \\<^sub>M M" shows "sets M \ sets (qbs_to_measure X)" proof fix U assume "U \ sets M" from sets.sets_into_space[OF this] in_mono[OF assms(2)] measurable_sets_borel[OF _ this] show "U \ sets (qbs_to_measure X)" using assms(1) by(auto intro!: exI[where x=U] simp: sigma_Mx_def sets_L) qed - lemma qbs_Mx_are_measurable[simp,measurable]: assumes "\ \ qbs_Mx X" shows "\ \ borel \\<^sub>M qbs_to_measure X" using assms qbs_Mx_subset_of_measurable by auto lemma measure_to_qbs_cong_sets: assumes "sets M = sets N" shows "measure_to_qbs M = measure_to_qbs N" by(rule qbs_eqI) (simp add: qbs_Mx_R measurable_cong_sets[OF _ assms]) lemma lr_sets[simp]: "sets X \ sets (qbs_to_measure (measure_to_qbs X))" unfolding sets_L proof safe fix U assume "U \ sets X" then have "U \ space X = U" by simp moreover have "\\\borel \\<^sub>M X. \ -` U \ sets borel" using \U \ sets X\ by(auto simp add: measurable_def) ultimately show "U \ sigma_Mx (measure_to_qbs X)" by(auto simp add: sigma_Mx_def qbs_Mx_R qbs_space_R) qed lemma(in standard_borel) lr_sets_ident[simp, measurable_cong]: "sets (qbs_to_measure (measure_to_qbs M)) = sets M" unfolding sets_L proof safe fix V assume "V \ sigma_Mx (measure_to_qbs M)" then obtain U where H2: "V = U \ space M" "\\::real \ _. \\borel \\<^sub>M M \ \ -` U \ sets borel" by(auto simp: sigma_Mx_def qbs_Mx_R qbs_space_R) consider "space M = {}" | "space M \ {}" by auto then show "V \ sets M" proof cases case 1 then show ?thesis by(simp add: H2) next case 2 have "from_real -` V = from_real -` (U \ space M)" using H2 by auto also have "... = from_real -` U" using from_real_measurable'[OF 2] by(auto simp add: measurable_def) finally have "to_real -` from_real -` U \ space M \ sets M" by (meson "2" H2(2) from_real_measurable' measurable_sets to_real_measurable) moreover have "to_real -` from_real -` U \ space M = U \ space M" by auto ultimately show ?thesis using H2 by simp qed qed(insert lr_sets, auto simp: sets_L) corollary sets_lr_polish_borel[simp, measurable_cong]: "sets (qbs_to_measure qbs_borel) = sets (borel :: (_ :: polish_space) measure)" by(auto intro!: standard_borel.lr_sets_ident standard_borel_ne.standard_borel) corollary sets_lr_count_space[simp, measurable_cong]: "sets (qbs_to_measure (qbs_count_space (UNIV :: (_ :: countable) set))) = sets (count_space UNIV)" by(rule standard_borel.lr_sets_ident) (auto intro!: standard_borel_ne.standard_borel) +lemma map_qbs_embed_measure1: + assumes "inj_on f (space M)" + shows "map_qbs f (measure_to_qbs M) = measure_to_qbs (embed_measure M f)" +proof(safe intro!: qbs_eqI) + fix \ + show "\ \ qbs_Mx (map_qbs f (measure_to_qbs M)) \ \ \ qbs_Mx (measure_to_qbs (embed_measure M f))" + using measurable_embed_measure2'[OF assms] by(auto intro!: qbs_eqI simp: map_qbs_Mx qbs_Mx_R) +next + fix \ + assume "\ \ qbs_Mx (measure_to_qbs (embed_measure M f))" + hence h[measurable]:"\ \ borel \\<^sub>M embed_measure M f" + by(simp add: qbs_Mx_R) + have [measurable]:"the_inv_into (space M) f \ embed_measure M f \\<^sub>M M" + by (simp add: assms in_sets_embed_measure measurable_def sets.sets_into_space space_embed_measure the_inv_into_into the_inv_into_vimage) + show "\ \ qbs_Mx (map_qbs f (measure_to_qbs M))" + using measurable_space[OF h] f_the_inv_into_f[OF assms] by(auto intro!: exI[where x="the_inv_into (space M) f \ \"] simp: map_qbs_Mx qbs_Mx_R space_embed_measure) +qed + +lemma map_qbs_embed_measure2: + assumes "inj_on f (qbs_space X)" + shows "sets (qbs_to_measure (map_qbs f X)) = sets (embed_measure (qbs_to_measure X) f)" +proof safe + fix A + assume "A \ sets (qbs_to_measure (map_qbs f X))" + then obtain U where "A = U \ f ` qbs_space X" "\\. \ \ qbs_Mx X \ (f \ \) -` U \ sets borel" + by(auto simp: sets_L sigma_Mx_def map_qbs_space map_qbs_Mx) + thus "A \ sets (embed_measure (qbs_to_measure X) f)" + by(auto simp: sets_L sets_embed_measure'[of _ "qbs_to_measure X",simplified space_L,OF assms] sigma_Mx_def qbs_Mx_to_X vimage_def intro!: exI[where x="f -` U \ qbs_space X"]) +next + fix A + assume A:"A \ sets (embed_measure (qbs_to_measure X) f)" + then obtain B U where "A = f ` B" "B = U \ qbs_space X" "\\. \ \ qbs_Mx X \ \ -` U \ sets borel" + by(auto simp: sets_L sigma_Mx_def sets_embed_measure'[of _ "qbs_to_measure X",simplified space_L,OF assms]) + thus "A \ sets (qbs_to_measure (map_qbs f X))" + using measurable_sets_borel[OF measurable_comp[OF qbs_Mx_are_measurable measurable_embed_measure2'[of _ "qbs_to_measure X",simplified space_L,OF assms]]] A + by(auto simp: sets_L sigma_Mx_def map_qbs_space map_qbs_Mx intro!: exI[where x="f ` (U \ qbs_space X)"]) +qed + +lemma(in standard_borel) map_qbs_embed_measure2': + assumes "inj_on f (space M)" + shows "sets (qbs_to_measure (map_qbs f (measure_to_qbs M))) = sets (embed_measure M f)" + by(auto intro!: standard_borel.lr_sets_ident[OF standard_borel_embed_measure] assms simp: map_qbs_embed_measure1[OF assms]) + subsubsection \ The Adjunction \ lemma lr_adjunction_correspondence : "X \\<^sub>Q (measure_to_qbs Y) = (qbs_to_measure X) \\<^sub>M Y" proof safe (* \ *) fix f assume "f \ X \\<^sub>Q (measure_to_qbs Y)" show "f \ qbs_to_measure X \\<^sub>M Y" proof(rule measurableI) fix x assume "x \ space (qbs_to_measure X)" thus "f x \ space Y" using qbs_morphism_space[OF \f \ X \\<^sub>Q (measure_to_qbs Y)\] by (auto simp: qbs_space_R space_L) next fix A assume "A \ sets Y" have "\\ \ qbs_Mx X. f \ \ \ qbs_Mx (measure_to_qbs Y)" using qbs_morphism_Mx[OF \f \ X \\<^sub>Q (measure_to_qbs Y)\] by auto hence "\\ \ qbs_Mx X. f \ \ \ borel \\<^sub>M Y" by (simp add: qbs_Mx_R) hence "\\ \ qbs_Mx X. \ -` (f -` A) \ sets borel" using \A\ sets Y\ measurable_sets_borel vimage_comp by metis thus "f -` A \ space (qbs_to_measure X) \ sets (qbs_to_measure X)" using sigma_Mx_def by (auto simp: space_L sets_L) qed (* \ *) next fix f assume "f \ qbs_to_measure X \\<^sub>M Y" show "f \ X \\<^sub>Q measure_to_qbs Y" proof(rule qbs_morphismI) fix \ assume "\ \ qbs_Mx X" have "f \ \ \ borel \\<^sub>M Y" proof(rule measurableI) fix x :: real from \\ \ qbs_Mx X\ qbs_Mx_to_X have "\ x \ qbs_space X" by auto hence "\ x \ space (qbs_to_measure X)" by (simp add: space_L) thus "(f \ \) x \ space Y" using \f \ qbs_to_measure X \\<^sub>M Y\ by (metis comp_def measurable_space) next fix A assume "A \ sets Y" from \f \ qbs_to_measure X \\<^sub>M Y\ measurable_sets this measurable_def have "f -` A \ space (qbs_to_measure X) \ sets (qbs_to_measure X)" by blast hence "f -` A \ qbs_space X \ sigma_Mx X" by (simp add: sets_L space_L) then have "\V. f -` A \ qbs_space X = V \ qbs_space X \ (\\\ qbs_Mx X. \ -` V \ sets borel)" by (simp add: sigma_Mx_def) then obtain V where h:"f -` A \ qbs_space X = V \ qbs_space X \ (\\\ qbs_Mx X. \ -` V \ sets borel)" by auto have 1:"\ -` (f -` A) = \ -` (f -` A \ qbs_space X)" using \\ \ qbs_Mx X\ qbs_Mx_to_X by blast have 2:"\ -` (V \ qbs_space X) = \ -` V" using \\ \ qbs_Mx X\ qbs_Mx_to_X by blast from 1 2 h have "(f \ \) -` A = \ -` V" by (simp add: vimage_comp) from this h \\ \ qbs_Mx X \show "(f \ \) -` A \ space borel \ sets borel" by simp qed thus "f \ \ \ qbs_Mx (measure_to_qbs Y)" by(simp add:qbs_Mx_R) qed qed lemma(in standard_borel) standard_borel_r_full_faithful: "M \\<^sub>M Y = measure_to_qbs M \\<^sub>Q measure_to_qbs Y" proof have "measure_to_qbs M \\<^sub>Q measure_to_qbs Y \ qbs_to_measure (measure_to_qbs M) \\<^sub>M qbs_to_measure (measure_to_qbs Y)" by (simp add: l_preserves_morphisms) also have "... = M \\<^sub>M qbs_to_measure (measure_to_qbs Y)" using measurable_cong_sets by auto also have "... \ M \\<^sub>M Y" by(rule measurable_mono[OF lr_sets]) (simp_all add: qbs_space_R space_L) finally show "measure_to_qbs M \\<^sub>Q measure_to_qbs Y \ M \\<^sub>M Y" . qed(rule r_preserves_morphisms) lemma qbs_morphism_dest: assumes "f \ X \\<^sub>Q measure_to_qbs Y" shows "f \ qbs_to_measure X \\<^sub>M Y" using assms lr_adjunction_correspondence by auto lemma(in standard_borel) qbs_morphism_dest: assumes "k \ measure_to_qbs M \\<^sub>Q measure_to_qbs Y" shows "k \ M \\<^sub>M Y" using standard_borel_r_full_faithful assms by auto lemma qbs_morphism_measurable_intro: assumes "f \ qbs_to_measure X \\<^sub>M Y" shows "f \ X \\<^sub>Q measure_to_qbs Y" using assms lr_adjunction_correspondence by auto lemma(in standard_borel) qbs_morphism_measurable_intro: assumes "k \ M \\<^sub>M Y" shows "k \ measure_to_qbs M \\<^sub>Q measure_to_qbs Y" using standard_borel_r_full_faithful assms by auto lemma r_preserves_product : "measure_to_qbs (X \\<^sub>M Y) = measure_to_qbs X \\<^sub>Q measure_to_qbs Y" by(auto intro!: qbs_eqI simp: measurable_pair_iff pair_qbs_Mx qbs_Mx_R) lemma l_product_sets: "sets (qbs_to_measure X \\<^sub>M qbs_to_measure Y) \ sets (qbs_to_measure (X \\<^sub>Q Y))" proof(rule sets_pair_in_sets) fix A B assume h:"A \ sets (qbs_to_measure X)" "B \ sets (qbs_to_measure Y)" then obtain Ua Ub where hu: "A = Ua \ qbs_space X" "\\\qbs_Mx X. \ -` Ua \ sets borel" "B = Ub \ qbs_space Y" "\\\qbs_Mx Y. \ -` Ub \ sets borel" by(auto simp add: sigma_Mx_def sets_L) show "A \ B \ sets (qbs_to_measure (X \\<^sub>Q Y))" proof - have "A \ B = Ua \ Ub \ qbs_space (X \\<^sub>Q Y) \ (\\\qbs_Mx (X \\<^sub>Q Y). \ -` (Ua \ Ub) \ sets borel)" using hu by(auto simp add: vimage_Times pair_qbs_space pair_qbs_Mx) thus ?thesis by(auto simp add: sigma_Mx_def sets_L intro!: exI[where x="Ua \ Ub"]) qed qed corollary qbs_borel_prod: "qbs_borel \\<^sub>Q qbs_borel = (qbs_borel :: ('a::second_countable_topology \ 'b::second_countable_topology) quasi_borel)" by(simp add: r_preserves_product[symmetric] borel_prod) corollary qbs_count_space_prod: "qbs_count_space (UNIV :: ('a :: countable) set) \\<^sub>Q qbs_count_space (UNIV :: ('b :: countable) set) = qbs_count_space UNIV" by(auto simp: r_preserves_product[symmetric] count_space_prod) lemma r_preserves_product': "measure_to_qbs (\\<^sub>M i\I. M i) = (\\<^sub>Q i\I. measure_to_qbs (M i))" proof(rule qbs_eqI) show "qbs_Mx (measure_to_qbs (Pi\<^sub>M I M)) = qbs_Mx (\\<^sub>Q i\I. measure_to_qbs (M i))" proof safe fix f :: "real \ _" assume "f \ qbs_Mx (measure_to_qbs (Pi\<^sub>M I M))" with measurable_space[of f borel "Pi\<^sub>M I M"] show "f \ qbs_Mx (\\<^sub>Q i\I. measure_to_qbs (M i))" by(auto simp: qbs_Mx_R PiQ_Mx space_PiM intro!:ext[of "\r. f r _"]) next fix f :: "real \ _" assume "f \ qbs_Mx (\\<^sub>Q i\I. measure_to_qbs (M i))" then have "\i. i \ I \ (\r. f r i) \ borel \\<^sub>M M i" "\i. i \ I \ (\r. f r i) = (\r. undefined)" by (auto simp: qbs_Mx_R PiQ_Mx) with measurable_space[OF this(1)] fun_cong[OF this(2)] show "f \ qbs_Mx (measure_to_qbs (Pi\<^sub>M I M))" by(auto intro!: measurable_PiM_single' simp: qbs_Mx_R) qed qed lemma PiQ_qbs_borel: "(\\<^sub>Q i::('a:: countable)\UNIV. (qbs_borel :: ('b::second_countable_topology quasi_borel))) = qbs_borel" by(simp add: r_preserves_product'[symmetric] measure_to_qbs_cong_sets[OF sets_PiM_equal_borel]) lemma qbs_morphism_from_countable: fixes X :: "'a quasi_borel" assumes "countable (qbs_space X)" "qbs_Mx X \ borel \\<^sub>M count_space (qbs_space X)" and "\i. i \ qbs_space X \ f i \ qbs_space Y" shows "f \ X \\<^sub>Q Y" proof(rule qbs_morphismI) fix \ assume "\ \ qbs_Mx X" then have [measurable]: "\ \ borel \\<^sub>M count_space (qbs_space X)" using assms(2) .. define k :: "'a \ real \ _" where "k \ (\i _. f i)" have "f \ \ = (\r. k (\ r) r)" by(auto simp add: k_def) also have "... \ qbs_Mx Y" by(rule qbs_closed3_dest2[OF assms(1)]) (use assms(3) k_def in simp_all) finally show "f \ \ \ qbs_Mx Y" . qed corollary qbs_morphism_count_space': assumes "\i. i \ I \ f i \ qbs_space Y" "countable I" shows "f \ qbs_count_space I \\<^sub>Q Y" using assms by(auto intro!: qbs_morphism_from_countable simp: qbs_Mx_R) corollary qbs_morphism_count_space: assumes "\i. f i \ qbs_space Y" shows "f \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q Y" using assms by(auto intro!: qbs_morphism_from_countable simp: qbs_Mx_R) lemma [qbs]: shows not_qbs_pred: "Not \ qbs_count_space UNIV \\<^sub>Q qbs_count_space UNIV" and or_qbs_pred: "(\) \ qbs_count_space UNIV \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and and_qbs_pred: "(\) \ qbs_count_space UNIV \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and implies_qbs_pred: "(\) \ qbs_count_space UNIV \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and iff_qbs_pred: "(\) \ qbs_count_space UNIV \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" by(auto intro!: qbs_morphism_count_space) lemma [qbs]: shows less_count_qbs_pred: "(<) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and le_count_qbs_pred: "(\) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and eq_count_qbs_pred: "(=) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and plus_count_qbs_morphism: "(+) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and minus_count_qbs_morphism: "(-) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and mult_count_qbs_morphism: "(*) \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q exp_qbs (qbs_count_space UNIV) (qbs_count_space UNIV)" and Suc_qbs_morphism: "Suc \ qbs_count_space UNIV \\<^sub>Q qbs_count_space UNIV" by(auto intro!: qbs_morphism_count_space) lemma qbs_morphism_product_iff: "f \ X \\<^sub>Q (\\<^sub>Q i :: (_ :: countable)\UNIV. Y) \ f \ X \\<^sub>Q qbs_count_space UNIV \\<^sub>Q Y" proof assume h:"f \ X \\<^sub>Q (\\<^sub>Q i\UNIV. Y)" show "f \ X \\<^sub>Q qbs_count_space UNIV \\<^sub>Q Y" by(rule arg_swap_morphism, rule qbs_morphism_count_space) (simp add: qbs_morphism_component_singleton'[OF h qbs_morphism_ident']) next assume "f \ X \\<^sub>Q qbs_count_space UNIV \\<^sub>Q Y" from qbs_morphism_space[OF arg_swap_morphism[OF this]] show "f \ X \\<^sub>Q (\\<^sub>Q i\UNIV. Y)" by(auto intro!: product_qbs_canonical1[where f="(\i x. f x i)"]) qed lemma qbs_morphism_pair_countable1: assumes "countable (qbs_space X)" "qbs_Mx X \ borel \\<^sub>M count_space (qbs_space X)" and "\i. i \ qbs_space X \ f i \ Y \\<^sub>Q Z" shows "(\(x,y). f x y) \ X \\<^sub>Q Y \\<^sub>Q Z" by(auto intro!: uncurry_preserves_morphisms qbs_morphism_from_countable[OF assms(1,2)] assms(3)) lemma qbs_morphism_pair_countable2: assumes "countable (qbs_space Y)" "qbs_Mx Y \ borel \\<^sub>M count_space (qbs_space Y)" and "\i. i \ qbs_space Y \ (\x. f x i) \ X \\<^sub>Q Z" shows "(\(x,y). f x y) \ X \\<^sub>Q Y \\<^sub>Q Z" by(auto intro!: qbs_morphism_pair_swap[of "case_prod (\x y. f y x)",simplified] qbs_morphism_pair_countable1 assms) corollary qbs_morphism_pair_count_space1: assumes "\i. f i \ Y \\<^sub>Q Z" shows "(\(x,y). f x y) \ qbs_count_space (UNIV :: ('a :: countable) set) \\<^sub>Q Y \\<^sub>Q Z" by(auto intro!: qbs_morphism_pair_countable1 simp: qbs_Mx_R assms) corollary qbs_morphism_pair_count_space2: assumes "\i. (\x. f x i) \ X \\<^sub>Q Z" shows "(\(x,y). f x y) \ X \\<^sub>Q qbs_count_space (UNIV :: ('a :: countable) set) \\<^sub>Q Z" by(auto intro!: qbs_morphism_pair_countable2 simp: qbs_Mx_R assms) lemma qbs_morphism_compose_countable': assumes [qbs]:"\i. i \ I \ (\x. f i x) \ X \\<^sub>Q Y" "g \ X \\<^sub>Q qbs_count_space I" "countable I" shows "(\x. f (g x) x) \ X \\<^sub>Q Y" proof - have [qbs]:"f \ qbs_count_space I \\<^sub>Q X \\<^sub>Q Y" by(auto intro!: qbs_morphism_count_space' simp: assms(3)) show ?thesis by simp qed lemma qbs_morphism_compose_countable: assumes [simp]:"\i::'i::countable. (\x. f i x) \ X \\<^sub>Q Y" "g \ X \\<^sub>Q (qbs_count_space UNIV)" shows "(\x. f (g x) x) \ X \\<^sub>Q Y" by(rule qbs_morphism_compose_countable'[of UNIV f]) simp_all lemma qbs_morphism_op: assumes "case_prod f \ X \\<^sub>M Y \\<^sub>M Z" shows "f \ measure_to_qbs X \\<^sub>Q measure_to_qbs Y \\<^sub>Q measure_to_qbs Z" using r_preserves_morphisms assms by(fastforce simp: r_preserves_product[symmetric] intro!: curry_preserves_morphisms) lemma [qbs]: shows plus_qbs_morphism: "(+) \ (qbs_borel :: (_::{second_countable_topology, topological_monoid_add}) quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and plus_ereal_qbs_morphism: "(+) \ (qbs_borel :: ereal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and diff_qbs_morphism: "(-) \ (qbs_borel :: (_::{second_countable_topology, real_normed_vector}) quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and diff_ennreal_qbs_morphism: "(-) \ (qbs_borel :: ennreal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and diff_ereal_qbs_morphism: "(-) \ (qbs_borel :: ereal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and times_qbs_morphism: "(*) \ (qbs_borel :: (_::{second_countable_topology, real_normed_algebra}) quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and times_ennreal_qbs_morphism: "(*) \ (qbs_borel :: ennreal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and times_ereal_qbs_morphism: "(*) \ (qbs_borel :: ereal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and divide_qbs_morphism: "(/) \ (qbs_borel :: (_::{second_countable_topology, real_normed_div_algebra}) quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and divide_ennreal_qbs_morphism: "(/) \ (qbs_borel :: ennreal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and divide_ereal_qbs_morphism: "(/) \ (qbs_borel :: ereal quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and log_qbs_morphism: "log \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and root_qbs_morphism: "root \ qbs_count_space UNIV \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and scaleR_qbs_morphism: "(*\<^sub>R) \ qbs_borel \\<^sub>Q (qbs_borel :: (_::{second_countable_topology, real_normed_vector}) quasi_borel) \\<^sub>Q qbs_borel" and qbs_morphism_inner: "(\) \ qbs_borel \\<^sub>Q (qbs_borel :: (_::{second_countable_topology, real_inner}) quasi_borel) \\<^sub>Q qbs_borel" and dist_qbs_morphism: "dist \ (qbs_borel :: (_::{second_countable_topology, metric_space}) quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and powr_qbs_morphism: "(powr) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q (qbs_borel :: real quasi_borel)" and max_qbs_morphism: "(max :: (_ :: {second_countable_topology, linorder_topology}) \ _ \ _) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and min_qbs_morphism: "(min :: (_ :: {second_countable_topology, linorder_topology}) \ _ \ _) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and sup_qbs_morphism: "(sup :: (_ :: {lattice,second_countable_topology, linorder_topology}) \ _ \ _) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and inf_qbs_morphism: "(inf :: (_ :: {lattice,second_countable_topology, linorder_topology}) \ _ \ _) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" and less_qbs_pred: "(<) \ (qbs_borel :: _ ::{second_countable_topology, linorder_topology} quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_count_space UNIV" and eq_qbs_pred: "(=) \ (qbs_borel :: _ ::{second_countable_topology, linorder_topology} quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_count_space UNIV" and le_qbs_pred: "(\) \ (qbs_borel :: _ ::{second_countable_topology, linorder_topology} quasi_borel) \\<^sub>Q qbs_borel \\<^sub>Q qbs_count_space UNIV" by(auto intro!: qbs_morphism_op) lemma [qbs]: shows abs_real_qbs_morphism: "abs \ (qbs_borel :: real quasi_borel) \\<^sub>Q qbs_borel" and abs_ereal_qbs_morphism: "abs \ (qbs_borel :: ereal quasi_borel) \\<^sub>Q qbs_borel" and real_floor_qbs_morphism: "(floor :: real \ int) \ qbs_borel \\<^sub>Q qbs_count_space UNIV" and real_ceiling_qbs_morphism: "(ceiling :: real \ int) \ qbs_borel \\<^sub>Q qbs_count_space UNIV" and exp_qbs_morphism: "(exp::'a::{real_normed_field,banach}\'a) \ qbs_borel \\<^sub>Q qbs_borel" and ln_qbs_morphism: "ln \ (qbs_borel :: real quasi_borel) \\<^sub>Q qbs_borel" and sqrt_qbs_morphism: "sqrt \ qbs_borel \\<^sub>Q qbs_borel" and of_real_qbs_morphism: "(of_real :: _ \ (_::real_normed_algebra)) \ qbs_borel \\<^sub>Q qbs_borel" and sin_qbs_morphism: "(sin :: _ \ (_::{real_normed_field,banach})) \ qbs_borel \\<^sub>Q qbs_borel" and cos_qbs_morphism: "(cos :: _ \ (_::{real_normed_field,banach})) \ qbs_borel \\<^sub>Q qbs_borel" and arctan_qbs_morphism: "arctan \ qbs_borel \\<^sub>Q qbs_borel" and Re_qbs_morphism: "Re \ qbs_borel \\<^sub>Q qbs_borel" and Im_qbs_morphism: "Im \ qbs_borel \\<^sub>Q qbs_borel" and sgn_qbs_morphism: "(sgn::_::real_normed_vector \ _) \ qbs_borel \\<^sub>Q qbs_borel" and norm_qbs_morphism: "norm \ qbs_borel \\<^sub>Q qbs_borel" and invers_qbs_morphism: "(inverse :: _ \ (_ ::real_normed_div_algebra)) \ qbs_borel \\<^sub>Q qbs_borel" and invers_ennreal_qbs_morphism: "(inverse :: _ \ ennreal) \ qbs_borel \\<^sub>Q qbs_borel" and invers_ereal_qbs_morphism: "(inverse :: _ \ ereal) \ qbs_borel \\<^sub>Q qbs_borel" and uminus_qbs_morphism: "(uminus :: _ \ (_::{second_countable_topology, real_normed_vector})) \ qbs_borel \\<^sub>Q qbs_borel" and ereal_qbs_morphism: "ereal \ qbs_borel \\<^sub>Q qbs_borel" and real_of_ereal_qbs_morphism: "real_of_ereal \ qbs_borel \\<^sub>Q qbs_borel" and enn2ereal_qbs_morphism: "enn2ereal \ qbs_borel \\<^sub>Q qbs_borel" and e2ennreal_qbs_morphism: "e2ennreal \ qbs_borel \\<^sub>Q qbs_borel" and ennreal_qbs_morphism: "ennreal \ qbs_borel \\<^sub>Q qbs_borel" and qbs_morphism_nth: "(\x::real^'n. x $ i) \ qbs_borel \\<^sub>Q qbs_borel" and qbs_morphism_product_candidate: "\i. (\x. x i) \ qbs_borel \\<^sub>Q qbs_borel" and uminus_ereal_qbs_morphism: "(uminus :: _ \ ereal) \ qbs_borel \\<^sub>Q qbs_borel" by(auto intro!: set_mp[OF r_preserves_morphisms]) lemma qbs_morphism_sum: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, topological_comm_monoid_add}" assumes "\i. i \ S \ f i \ X \\<^sub>Q qbs_borel" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_suminf_order: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}" assumes "\i. f i \ X \\<^sub>Q qbs_borel" shows " (\x. \i. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_prod: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, real_normed_field}" assumes "\i. i \ S \ f i \ X \\<^sub>Q qbs_borel" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_Min: "finite I \ (\i. i \ I \ f i \ X \\<^sub>Q qbs_borel) \ (\x. Min ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ X \\<^sub>Q qbs_borel" by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_Max: "finite I \ (\i. i \ I \ f i \ X \\<^sub>Q qbs_borel) \ (\x. Max ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ X \\<^sub>Q qbs_borel" by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_Max2: fixes f::"_ \ _ \ 'a::{second_countable_topology, dense_linorder, linorder_topology}" shows "finite I \ (\i. f i \ X \\<^sub>Q qbs_borel) \ (\x. Max{f i x |i. i \ I}) \ X \\<^sub>Q qbs_borel" by(simp add: lr_adjunction_correspondence) lemma [qbs]: shows qbs_morphism_liminf: "liminf \ (qbs_count_space UNIV \\<^sub>Q qbs_borel) \\<^sub>Q (qbs_borel :: 'a :: {complete_linorder, second_countable_topology, linorder_topology} quasi_borel)" and qbs_morphism_limsup: "limsup \ (qbs_count_space UNIV \\<^sub>Q qbs_borel) \\<^sub>Q (qbs_borel :: 'a :: {complete_linorder, second_countable_topology, linorder_topology} quasi_borel)" and qbs_morphism_lim: "lim \ (qbs_count_space UNIV \\<^sub>Q qbs_borel) \\<^sub>Q (qbs_borel :: 'a :: {complete_linorder, second_countable_topology, linorder_topology} quasi_borel)" proof(safe intro!: qbs_morphismI) fix f :: "real \ nat \ 'a" assume "f \ qbs_Mx (count_space\<^sub>Q UNIV \\<^sub>Q borel\<^sub>Q)" then have [measurable]:"\i. (\r. f r i) \ borel_measurable borel" by(auto simp: qbs_Mx_is_morphisms) (metis PiQ_qbs_borel measurable_product_then_coordinatewise qbs_Mx_is_morphisms qbs_Mx_qbs_borel qbs_morphism_product_iff) show "liminf \ f \ qbs_Mx borel\<^sub>Q" "limsup \ f \ qbs_Mx borel\<^sub>Q" "lim \ f \ qbs_Mx borel\<^sub>Q" by(auto simp: qbs_Mx_is_morphisms lr_adjunction_correspondence comp_def) qed lemma qbs_morphism_SUP: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes "countable I" "\i. i \ I \ F i \ X \\<^sub>Q qbs_borel" shows "(\x. \ i\I. F i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_INF: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes "countable I" "\i. i \ I \ F i \ X \\<^sub>Q qbs_borel" shows "(\x. \ i\I. F i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_cSUP: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes "countable I" "\i. i \ I \ F i \ X \\<^sub>Q qbs_borel" "\x. x \ qbs_space X \ bdd_above ((\i. F i x) ` I)" shows "(\x. \ i\I. F i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence space_L) lemma qbs_morphism_cINF: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes "countable I" "\i. i \ I \ F i \ X \\<^sub>Q qbs_borel" "\x. x \ qbs_space X \ bdd_below ((\i. F i x) ` I)" shows "(\x. \ i\I. F i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence space_L) lemma qbs_morphism_lim_metric: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes "\i. f i \ X \\<^sub>Q qbs_borel" shows "(\x. lim (\i. f i x)) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_LIMSEQ_metric: fixes f :: "nat \ 'a \ 'b :: metric_space" assumes "\i. f i \ X \\<^sub>Q qbs_borel" "\x. x \ qbs_space X \ (\i. f i x) \ g x" shows "g \ X \\<^sub>Q qbs_borel" using borel_measurable_LIMSEQ_metric[where M="qbs_to_measure X"] assms by(auto simp add: lr_adjunction_correspondence space_L) lemma power_qbs_morphism[qbs]: "(power :: (_ ::{power,real_normed_algebra}) \ nat \ _) \ qbs_borel \\<^sub>Q qbs_count_space UNIV \\<^sub>Q qbs_borel" by(rule arg_swap_morphism) (auto intro!: qbs_morphism_count_space set_mp[OF r_preserves_morphisms]) lemma power_ennreal_qbs_morphism[qbs]: "(power :: ennreal \ nat \ _) \ qbs_borel \\<^sub>Q qbs_count_space UNIV \\<^sub>Q qbs_borel" by(rule arg_swap_morphism) (auto intro!: qbs_morphism_count_space set_mp[OF r_preserves_morphisms]) lemma qbs_morphism_compw: "(^^) \ (X \\<^sub>Q X) \\<^sub>Q qbs_count_space UNIV \\<^sub>Q (X \\<^sub>Q X)" proof(rule arg_swap_morphism,rule qbs_morphism_count_space) fix n show "(\y. y ^^ n) \ X \\<^sub>Q X \\<^sub>Q X \\<^sub>Q X" by(induction n) simp_all qed lemma qbs_morphism_compose_n[qbs]: assumes [qbs]: "f \ X \\<^sub>Q X" shows "(\n. f^^n) \ qbs_count_space UNIV \\<^sub>Q X \\<^sub>Q X" proof(intro qbs_morphism_count_space) fix n show "f ^^ n \ X \\<^sub>Q X" by (induction n) simp_all qed lemma qbs_morphism_compose_n': assumes "f \ X \\<^sub>Q X" shows "f^^n \ X \\<^sub>Q X" using qbs_morphism_space[OF qbs_morphism_compose_n[OF assms]] by(simp add: exp_qbs_space qbs_space_R) lemma qbs_morphism_uminus_eq_ereal[simp]: "(\x. - f x :: ereal) \ X \\<^sub>Q qbs_borel \ f \ X \\<^sub>Q qbs_borel" (is "?l = ?r") by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_ereal_iff: shows "(\x. ereal (f x)) \ X \\<^sub>Q qbs_borel\ f \ X \\<^sub>Q qbs_borel" by(simp add: borel_measurable_ereal_iff lr_adjunction_correspondence) lemma qbs_morphism_ereal_sum: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ X \\<^sub>Q qbs_borel" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_ereal_prod: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ X \\<^sub>Q qbs_borel" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_extreal_suminf: fixes f :: "nat \ 'a \ ereal" assumes "\i. f i \ X \\<^sub>Q qbs_borel" shows "(\x. (\i. f i x)) \ X \\<^sub>Q qbs_borel" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_ennreal_iff: assumes "\x. x \ qbs_space X \ 0 \ f x" shows "(\x. ennreal (f x)) \ X \\<^sub>Q qbs_borel \ f \ X \\<^sub>Q qbs_borel" using borel_measurable_ennreal_iff[where M="qbs_to_measure X"] assms by(simp add: space_L lr_adjunction_correspondence) lemma qbs_morphism_prod_ennreal: fixes f :: "'c \ 'a \ ennreal" assumes "\i. i \ S \ f i \ X \\<^sub>Q qbs_borel" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_borel" using assms by(simp add: space_L lr_adjunction_correspondence) lemma count_space_qbs_morphism: "f \ qbs_count_space (UNIV :: 'a set) \\<^sub>Q qbs_borel" by(auto intro!: set_mp[OF r_preserves_morphisms]) declare count_space_qbs_morphism[where 'a="_ :: countable",qbs] lemma count_space_count_space_qbs_morphism: "f \ qbs_count_space (UNIV :: (_ :: countable) set) \\<^sub>Q qbs_count_space (UNIV :: (_ :: countable) set)" by(auto intro!: set_mp[OF r_preserves_morphisms]) lemma qbs_morphism_case_nat': assumes [qbs]: "i = 0 \ f \ X \\<^sub>Q Y" "\j. i = Suc j \ (\x. g x j) \ X \\<^sub>Q Y" shows "(\x. case_nat (f x) (g x) i) \ X \\<^sub>Q Y" by (cases i) simp_all lemma qbs_morphism_case_nat[qbs]: "case_nat \ X \\<^sub>Q (qbs_count_space UNIV \\<^sub>Q X) \\<^sub>Q qbs_count_space UNIV \\<^sub>Q X" by(rule curry_preserves_morphisms, rule arg_swap_morphism) (auto intro!: qbs_morphism_count_space qbs_morphism_case_nat') lemma qbs_morphism_case_nat'': assumes "f \ X \\<^sub>Q Y" "g \ X \\<^sub>Q (\\<^sub>Q i\UNIV. Y)" shows "(\x. case_nat (f x) (g x)) \ X \\<^sub>Q (\\<^sub>Q i\UNIV. Y)" using assms by (simp add: qbs_morphism_product_iff) lemma qbs_morphism_rec_nat[qbs]: "rec_nat \ X \\<^sub>Q (count_space UNIV \\<^sub>Q X \\<^sub>Q X) \\<^sub>Q count_space UNIV \\<^sub>Q X" proof(rule curry_preserves_morphisms,rule arg_swap_morphism,rule qbs_morphism_count_space) fix n show "(\y. rec_nat (fst y) (snd y) n) \ X \\<^sub>Q (qbs_count_space UNIV \\<^sub>Q X \\<^sub>Q X) \\<^sub>Q X" by (induction n) simp_all qed lemma qbs_morphism_Max_nat: fixes P :: "nat \ 'a \ bool" assumes "\i. P i \ X \\<^sub>Q qbs_count_space UNIV" shows "(\x. Max {i. P i x}) \ X \\<^sub>Q qbs_count_space UNIV" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_Min_nat: fixes P :: "nat \ 'a \ bool" assumes "\i. P i \ X \\<^sub>Q qbs_count_space UNIV" shows "(\x. Min {i. P i x}) \ X \\<^sub>Q qbs_count_space UNIV" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_sum_nat: fixes f :: "'c \ 'a \ nat" assumes "\i. i \ S \ f i \X \\<^sub>Q qbs_count_space UNIV" shows "(\x. \i\S. f i x) \ X \\<^sub>Q qbs_count_space UNIV" using assms by(simp add: lr_adjunction_correspondence) lemma qbs_morphism_case_enat': assumes f[qbs]: "f \ X \\<^sub>Q qbs_count_space UNIV" and [qbs]: "\i. g i \ X \\<^sub>Q Y" "h \ X \\<^sub>Q Y" shows "(\x. case f x of enat i \ g i x | \ \ h x) \ X \\<^sub>Q Y" proof (rule qbs_morphism_compose_countable[OF _ f]) fix i show "(\x. case i of enat i \ g i x | \ \ h x) \ X \\<^sub>Q Y" by (cases i) simp_all qed lemma qbs_morphism_case_enat[qbs]: "case_enat \ qbs_space ((qbs_count_space UNIV \\<^sub>Q X) \\<^sub>Q X \\<^sub>Q qbs_count_space UNIV \\<^sub>Q X)" proof - note qbs_morphism_case_enat'[qbs] show ?thesis by(auto intro!: curry_preserves_morphisms,rule qbs_morphismI) (simp add: qbs_Mx_is_morphisms comp_def, qbs, simp_all) qed lemma qbs_morphism_restrict[qbs]: assumes X: "\i. i \ I \ f i \ X \\<^sub>Q (Y i)" shows "(\x. \i\I. f i x) \ X \\<^sub>Q (\\<^sub>Q i\I. Y i)" using assms by(auto intro!: product_qbs_canonical1) lemma If_qbs_morphism[qbs]: "If \ qbs_count_space UNIV \\<^sub>Q X \\<^sub>Q X \\<^sub>Q X" proof(rule qbs_morphismI) show "\ \ qbs_Mx (count_space\<^sub>Q UNIV) \ If \ \ \ qbs_Mx (X \\<^sub>Q X \\<^sub>Q X)" for \ by(auto intro!: qbs_Mx_indicat[where S="{r. \ (_ (_ r))}",simplified] simp: qbs_Mx_count_space exp_qbs_Mx) qed lemma normal_density_qbs[qbs]: "normal_density \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" proof - have [simp]:"normal_density = (\\ \ x. 1 / sqrt (2 * pi * \\<^sup>2) * exp (-(x - \)\<^sup>2/ (2 * \\<^sup>2)))" by standard+ (auto simp: normal_density_def) show ?thesis by simp qed lemma erlang_density_qbs[qbs]: "erlang_density \ qbs_count_space UNIV \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" proof - have [simp]: "erlang_density = (\k l x. (if x < 0 then 0 else (l^(Suc k) * x^k * exp (- l * x)) / fact k))" by standard+ (auto simp: erlang_density_def) show ?thesis by simp qed lemma list_nil_qbs[qbs]: "[] \ qbs_space (list_qbs X)" by(simp add: list_qbs_space) -lemma list_cons_qbs_morphism: "list_cons \ X \\<^sub>Q (list_of X) \\<^sub>Q (list_of X)" +lemma list_cons_qbs_morphism: "list_cons \ X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. \ assume h:"\ \ qbs_Mx X" - "\ \ qbs_Mx (list_of X)" + "\ \ qbs_Mx (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. f where hf: "\ = (\r. (f r, \ (f r) r))" "f \ borel \\<^sub>M count_space UNIV" "\i. i \ range f \ \ i \ qbs_Mx (\\<^sub>Q j\{..' where "f' \ (\r. Suc (f r))" "\' \ (\i r n. if n = 0 then \ r else \ (i - 1) r (n - 1))" then have "(\r. list_cons (fst (\ r, \ r)) (snd (\ r, \ r))) = (\r. (f' r, \' (f' r) r))" by(auto simp: comp_def hf(1) ext list_cons_def) - also have "... \ qbs_Mx (list_of X)" - unfolding list_of_def - proof(rule coprod_qbs_MxI) + also have "... \ qbs_Mx (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. borel \\<^sub>M count_space UNIV" using hf by(simp add: f'_\'_def(1)) next fix j assume hj:"j \ range f'" then have hj':"j - 1 \ range f" by(auto simp: f'_\'_def(1)) show "\' j \ qbs_Mx (\\<^sub>Q i\{.. {..r. \' j r i) \ qbs_Mx X" proof cases case 1 then show ?thesis by(simp add: h(1) f'_\'_def(2)) next case 2 then have "i - 1 \ {..'_def(2)) qed next fix i assume hi:"i \ {.. 0" "i - Suc 0 \ {..'_def(1) hj by fastforce+ with prod_qbs_MxD(2)[OF hf(3)[OF hj']] show "(\r. \' j r i) = (\r. undefined)" by(simp add: f'_\'_def(2)) qed qed - finally show "(\r. list_cons (fst (\ r, \ r)) (snd (\ r, \ r))) \ qbs_Mx (list_of X)" . + finally show "(\r. list_cons (fst (\ r, \ r)) (snd (\ r, \ r))) \ qbs_Mx (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. X \\<^sub>Q (list_qbs X) \\<^sub>Q list_qbs X" proof(rule arg_swap_morphism) show "(\x y. y # x) \ list_qbs X \\<^sub>Q X \\<^sub>Q list_qbs X" - proof(rule qbs_morphism_cong'[where f="(\l x. x # (to_list l)) \ from_list"]) + proof(rule qbs_morphism_cong') show " (\l x. x # to_list l) \ from_list \ list_qbs X \\<^sub>Q X \\<^sub>Q list_qbs X" - proof(rule qbs_morphism_comp[where Y="list_of X"]) - show " (\l x. x # to_list l) \ list_of X \\<^sub>Q X \\<^sub>Q list_qbs X" + proof(rule qbs_morphism_comp) + show " (\l x. x # to_list l) \ (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q list_qbs X" proof(rule curry_preserves_morphisms) - show "(\lx. snd lx # to_list (fst lx)) \ list_of X \\<^sub>Q X \\<^sub>Q list_qbs X" - proof(rule qbs_morphism_cong'[where f="to_list \ (\(l,x). from_list (x # to_list l))"]) - show "to_list \ (\(l, x). from_list (x # to_list l)) \ list_of X \\<^sub>Q X \\<^sub>Q list_qbs X" - proof(rule qbs_morphism_comp[where Y="list_of X"]) - show "(\(l, x). from_list (x # to_list l)) \ list_of X \\<^sub>Q X \\<^sub>Q list_of X" - by(rule qbs_morphism_cong'[where f="(\(l,x). list_cons x l)",OF _ uncurry_preserves_morphisms[of "\(l,x). list_cons x l",simplified,OF arg_swap_morphism[OF list_cons_qbs_morphism]]]) (auto simp: pair_qbs_space to_list_from_list_ident) + show "(\lx. snd lx # to_list (fst lx)) \ (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q list_qbs X" + proof(rule qbs_morphism_cong') + show "to_list \ (\(l, x). from_list (x # to_list l)) \ (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q list_qbs X" + proof(rule qbs_morphism_comp) + show "(\(l, x). from_list (x # to_list l)) \(\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..(l,x). list_cons x l",simplified,OF arg_swap_morphism[OF list_cons_qbs_morphism]]]) (auto simp: pair_qbs_space to_list_from_list_ident) qed(simp add: list_qbs_def map_qbs_morphism_f) qed(auto simp: pair_qbs_space to_list_from_list_ident to_list_simp2) qed qed(auto simp: list_qbs_def to_list_from_list_ident intro!: map_qbs_morphism_inverse_f) qed(simp add: from_list_to_list_ident) qed lemma rec_list_morphism': - "rec_list' \ qbs_space (Y \\<^sub>Q (X \\<^sub>Q list_of X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q list_of X \\<^sub>Q Y)" - unfolding list_of_def -proof(intro curry_preserves_morphisms[OF arg_swap_morphism] coprod_qbs_canonical1') + "rec_list' \ qbs_space (Y \\<^sub>Q (X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y)" +proof(intro curry_preserves_morphisms[OF arg_swap_morphism] coPiQ_canonical1') fix n show "(\x y. rec_list' (fst y) (snd y) (n, x)) \ (\\<^sub>Q i\{..\<^sub>Q exp_qbs (Y \\<^sub>Q exp_qbs X (exp_qbs (\\<^sub>Q n\UNIV. \\<^sub>Q i\{.. assume h:"\ \ qbs_Mx ((\\<^sub>Q i\{..<0::nat}. X) \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs (\\<^sub>Q n\UNIV. \\<^sub>Q i\{..r. fst (\ r) = (\n. undefined)" proof - fix r have "\i. (\r. fst (\ r) i) = (\r. undefined)" using h by(auto simp: exp_qbs_Mx PiQ_Mx pair_qbs_Mx comp_def split_beta') thus "fst (\ r) = (\n. undefined)" by(fastforce dest: fun_cong) qed hence "(\xy. rec_list' (fst (snd xy)) (snd (snd xy)) (0, fst xy)) \ \ = (\x. fst (snd (\ x)))" by(auto simp: rec_list'_simp1[simplified list_nil_def] comp_def split_beta') also have "... \ qbs_Mx Y" using h by(auto simp: pair_qbs_Mx comp_def) finally show "(\xy. rec_list' (fst (snd xy)) (snd (snd xy)) (0, fst xy)) \ \ \ qbs_Mx Y" . qed next case ih:(Suc n) show ?case proof(rule qbs_morphismI) fix \ assume h:"\ \ qbs_Mx (\\<^sub>Q i\{..' where "\' \ (\r. snd (list_tail (Suc n, \ r)))" define a where "a \ (\r. \ r 0)" then have ha:"a \ qbs_Mx X" using h by(auto simp: PiQ_Mx) have 1:"\' \ qbs_Mx (\\<^sub>Q i\{..'_def) - hence 2: "\r. (n, \' r) \ qbs_space (list_of X)" - using qbs_Mx_to_X[of \'] by (fastforce simp: PiQ_space coprod_qbs_space list_of_def) - have 3: "\r. (Suc n, \ r) \ qbs_space (list_of X)" - using qbs_Mx_to_X[of \] h by (fastforce simp: PiQ_space coprod_qbs_space list_of_def) + hence 2: "\r. (n, \' r) \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..'] by (fastforce simp: PiQ_space coPiQ_space) + have 3: "\r. (Suc n, \ r) \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..] h by (fastforce simp: PiQ_space coPiQ_space) have 4: "\r. (n, \' r) = list_tail (Suc n, \ r)" by(simp add: list_tail_def \'_def) have 5: "\r. (Suc n, \ r) = list_cons (a r) (n, \' r)" unfolding a_def by(simp add: list_simp5[OF 3,simplified 4[symmetric],simplified list_head_def list_cons_def list_nil_def] list_cons_def) auto - have 6: "(\r. (n, \' r)) \ qbs_Mx (list_of X)" - using 1 by(auto intro!: coprod_qbs_MxI simp: PiQ_space coprod_qbs_space list_of_def) + have 6: "(\r. (n, \' r)) \ qbs_Mx (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..x y. rec_list' (fst y) (snd y) (Suc n, x)) \ \ = (\r y. rec_list' (fst y) (snd y) (Suc n, \ r))" by auto also have "... = (\r y. snd y (a r) (n, \' r) (rec_list' (fst y) (snd y) (n, \' r)))" by(simp only: 5 rec_list'_simp2[OF 2]) - also have "... \ qbs_Mx (exp_qbs (Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y))) Y)" + also have "... \ qbs_Mx (exp_qbs (Y \\<^sub>Q exp_qbs X (exp_qbs (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..(r,y). snd y (a r) (n, \' r) (rec_list' (fst y) (snd y) (n, \' r))) = (\(y,x1,x2,x3). y x1 x2 x3) \ (\(r,y). (snd y, a r, (n, \' r), rec_list' (fst y) (snd y) (n, \' r)))" by auto - also have "... \ qbs_borel \\<^sub>Q (Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y))) \\<^sub>Q Y" - proof(rule qbs_morphism_comp[where Y="exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q X \\<^sub>Q list_of X \\<^sub>Q Y"]) - show "(\(r, y). (snd y, a r, (n, \' r), rec_list' (fst y) (snd y) (n, \' r))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q X \\<^sub>Q list_of X \\<^sub>Q Y" + also have "... \ qbs_borel \\<^sub>Q (Y \\<^sub>Q exp_qbs X (exp_qbs (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y" + proof(rule qbs_morphism_comp[where Y="exp_qbs X (exp_qbs (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y"]) + show "(\(r, y). (snd y, a r, (n, \' r), rec_list' (fst y) (snd y) (n, \' r))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y" unfolding split_beta' proof(safe intro!: qbs_morphism_Pair) - show "(\x. a (fst x)) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q X" + show "(\x. a (fst x)) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X" using ha qbs_Mx_is_morphisms[of X] ha by auto next - show "(\x. (n, \' (fst x))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q list_of X" + show "(\x. (n, \' (fst x))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..x. rec_list' (fst (snd x)) (snd (snd x)) (n, \' (fst x))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q Y" - using qbs_morphism_Mx[OF ih 1, simplified comp_def] uncurry_preserves_morphisms[of "(\(x,y). rec_list' (fst y) (snd y) (n, \' x))" qbs_borel "Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y))" Y] qbs_Mx_is_morphisms[of "exp_qbs (Y \\<^sub>Q exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y))) Y"] - by(fastforce simp: split_beta' list_of_def) + show "(\x. rec_list' (fst (snd x)) (snd (snd x)) (n, \' (fst x))) \ qbs_borel \\<^sub>Q Y \\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y" + using qbs_morphism_Mx[OF ih 1, simplified comp_def] uncurry_preserves_morphisms[of "(\(x,y). rec_list' (fst y) (snd y) (n, \' x))" qbs_borel "Y \\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..(y, x1, x2, x3). y x1 x2 x3) \ exp_qbs X (exp_qbs (list_of X) (exp_qbs Y Y)) \\<^sub>Q X \\<^sub>Q list_of X \\<^sub>Q Y \\<^sub>Q Y" + show "(\(y, x1, x2, x3). y x1 x2 x3) \ exp_qbs X (exp_qbs ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q X \\<^sub>Q (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y \\<^sub>Q Y" by simp qed finally show ?thesis by(simp add: exp_qbs_Mx') qed finally show "(\x y. rec_list' (fst y) (snd y) (Suc n, x)) \ \ \ qbs_Mx (exp_qbs (Y \\<^sub>Q exp_qbs X (exp_qbs (\\<^sub>Q n\UNIV. \\<^sub>Q i\{.. qbs_space (Y \\<^sub>Q (X \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q list_qbs X \\<^sub>Q Y)" proof(rule curry_preserves_morphisms[OF arg_swap_morphism]) show "(\l yf. rec_list (fst yf) (snd yf) l) \ list_qbs X \\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q Y" - proof(rule qbs_morphism_cong'[where f="(\l' (y,f). rec_list y f (to_list l')) \ from_list",OF _ qbs_morphism_comp[where Y="list_of X"]]) - show "(\l' (y,f). rec_list y f (to_list l')) \ list_of X \\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q Y" + proof(rule qbs_morphism_cong'[where f="(\l' (y,f). rec_list y f (to_list l')) \ from_list",OF _ qbs_morphism_comp[where Y="(\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..l' (y,f). rec_list y f (to_list l')) \ (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q Y" apply(rule arg_swap_morphism,simp only: split_beta' list_qbs_def) apply(rule uncurry_preserves_morphisms) apply(rule arg_swap_morphism) apply(rule arg_swap_morphism') apply(rule qbs_morphism_cong'[OF _ arg_swap_morphism_map_qbs1[OF arg_swap_morphism'[OF arg_swap_morphism[OF rec_list_morphism']]]]) apply(auto simp: rec_list'_def from_list_to_list_ident) done qed(auto simp: from_list_to_list_ident list_qbs_def to_list_from_list_ident intro!: map_qbs_morphism_inverse_f) qed hide_const (open) list_nil list_cons list_head list_tail from_list rec_list' to_list' hide_fact (open) list_simp1 list_simp2 list_simp3 list_simp4 list_simp5 list_simp6 list_simp7 from_list_in_list_of' list_cons_qbs_morphism rec_list'_simp1 to_list_from_list_ident from_list_in_list_of to_list_set to_list_simp1 to_list_simp2 list_head_def list_tail_def from_list_length list_cons_in_list_of rec_list_morphism' rec_list'_simp2 list_decomp1 list_destruct_rule list_induct_rule from_list_to_list_ident corollary case_list_morphism[qbs]: "case_list \ qbs_space ((Y :: 'b quasi_borel) \\<^sub>Q ((X :: 'a quasi_borel) \\<^sub>Q list_qbs X \\<^sub>Q Y) \\<^sub>Q list_qbs X \\<^sub>Q Y)" proof - have [simp]:"case_list = (\y (f :: 'a \ 'a list \ 'b) l. rec_list y (\x l' y. f x l') l)" proof standard+ fix y :: 'b and f :: "'a \ 'a list \ 'b" and l :: "'a list" show "(case l of [] \ y | x # xa \ f x xa) = rec_list y (\x l' y. f x l') l" by (cases l) auto qed show ?thesis by simp qed lemma fold_qbs_morphism[qbs]: "fold \ qbs_space ((X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y)" proof - have [simp]:"fold = (\f l. rec_list id (\x xs l. l \ f x) l)" apply standard+ subgoal for f l x by(induction l arbitrary: x) simp_all done show ?thesis by simp qed lemma [qbs]: shows foldr_qbs_morphism: "foldr \ qbs_space ((X \\<^sub>Q Y \\<^sub>Q Y) \\<^sub>Q list_qbs X \\<^sub>Q Y \\<^sub>Q Y)" and foldl_qbs_morphism: "foldl \ qbs_space ((X \\<^sub>Q Y \\<^sub>Q X) \\<^sub>Q X \\<^sub>Q list_qbs Y \\<^sub>Q X)" and zip_qbs_morphism: "zip \ qbs_space (list_qbs X \\<^sub>Q list_qbs Y \\<^sub>Q list_qbs (pair_qbs X Y))" and append_qbs_morphism: "append \ qbs_space (list_qbs X \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" and concat_qbs_morphism: "concat \ qbs_space (list_qbs (list_qbs X) \\<^sub>Q list_qbs X)" and drop_qbs_morphism: "drop \ qbs_space (qbs_count_space UNIV \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" and take_qbs_morphism: "take \ qbs_space (qbs_count_space UNIV \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" and rev_qbs_morphism: "rev \ qbs_space (list_qbs X \\<^sub>Q list_qbs X)" by(auto simp: foldr_def foldl_def zip_def append_def concat_def drop_def take_def rev_def) lemma [qbs]: fixes X :: "'a quasi_borel" and Y :: "'b quasi_borel" shows map_qbs_morphism: "map \ qbs_space ((X \\<^sub>Q Y) \\<^sub>Q list_qbs X \\<^sub>Q list_qbs Y)" (is ?map) and fileter_qbs_morphism: "filter \ qbs_space ((X \\<^sub>Q count_space\<^sub>Q UNIV) \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" (is ?filter) and length_qbs_morphism: "length \ qbs_space (list_qbs X \\<^sub>Q qbs_count_space UNIV)" (is ?length) and tl_qbs_morphism: "tl \ qbs_space (list_qbs X \\<^sub>Q list_qbs X)" (is ?tl) and list_all_qbs_morphism: "list_all \ qbs_space ((X \\<^sub>Q qbs_count_space UNIV) \\<^sub>Q list_qbs X \\<^sub>Q qbs_count_space UNIV)" (is ?list_all) and bind_list_qbs_morphism: "(\) \ qbs_space (list_qbs X \\<^sub>Q (X \\<^sub>Q list_qbs Y) \\<^sub>Q list_qbs Y)" (is ?bind) proof - have [simp]: "map = (\f. rec_list [] (\x xs l. f x # l))" apply standard+ subgoal for f l by(induction l) simp_all done have [simp]: "filter = (\P. rec_list [] (\x xs l. if P x then x # l else l))" apply standard+ subgoal for f l by(induction l) simp_all done have [simp]: "length = (\l. foldr (\_ n. Suc n) l 0)" apply standard subgoal for l by (induction l) simp_all done have [simp]: "tl = (\l. case l of [] \ [] | _ # xs \ xs)" by standard (simp add: tl_def) have [simp]: "list_all = (\P xs. foldr (\x b. b \ P x) xs True)" apply (standard,standard) subgoal for P xs by(induction xs arbitrary: P) auto done have [simp]: "List.bind = (\xs f. concat (map f xs))" by standard+ (simp add: List.bind_def) show ?map ?filter ?length ?tl ?list_all ?bind by simp_all qed lemma list_eq_qbs_morphism[qbs]: assumes [qbs]: "(=) \ qbs_space (X \\<^sub>Q X \\<^sub>Q count_space UNIV)" shows "(=) \ qbs_space (list_qbs X \\<^sub>Q list_qbs X \\<^sub>Q count_space UNIV)" proof - have [simp]:"(=) = (\xs ys. length xs = length ys \ list_all (case_prod (=)) (zip xs ys))" using Ball_set list_eq_iff_zip_eq by fastforce show ?thesis by simp qed lemma insort_key_qbs_morphism[qbs]: shows "insort_key \ qbs_space ((X \\<^sub>Q (borel\<^sub>Q ::'b :: {second_countable_topology, linorder_topology} quasi_borel)) \\<^sub>Q X \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" (is ?g1) and "insort_key \ qbs_space ((X \\<^sub>Q count_space\<^sub>Q (UNIV :: (_ :: countable) set)) \\<^sub>Q X \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" (is ?g2) proof - have [simp]:"insort_key = (\f x. rec_list [x] (\y ys l. if f x \ f y then x#y#ys else y#l))" apply standard+ subgoal for f x l by(induction l) simp_all done show ?g1 ?g2 by simp_all qed lemma sort_key_qbs_morphism[qbs]: shows "sort_key \ qbs_space ((X \\<^sub>Q (borel\<^sub>Q ::'b :: {second_countable_topology, linorder_topology} quasi_borel)) \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" and "sort_key \ qbs_space ((X \\<^sub>Q count_space\<^sub>Q (UNIV :: (_ :: countable) set)) \\<^sub>Q list_qbs X \\<^sub>Q list_qbs X)" unfolding sort_key_def by simp_all lemma sort_qbs_morphism[qbs]: shows "sort \ list_qbs (borel\<^sub>Q ::'b :: {second_countable_topology, linorder_topology} quasi_borel) \\<^sub>Q list_qbs borel\<^sub>Q" and "sort \ list_qbs (count_space\<^sub>Q (UNIV :: (_ :: countable) set)) \\<^sub>Q list_qbs (count_space\<^sub>Q UNIV)" by simp_all subsubsection \ Morphism Pred\ abbreviation "qbs_pred X P \ P \ X \\<^sub>Q qbs_count_space (UNIV :: bool set)" lemma qbs_pred_iff_measurable_pred: "qbs_pred X P = Measurable.pred (qbs_to_measure X) P" by(simp add: lr_adjunction_correspondence) lemma(in standard_borel) qbs_pred_iff_measurable_pred: "qbs_pred (measure_to_qbs M) P = Measurable.pred M P" by(simp add: qbs_pred_iff_measurable_pred measurable_cong_sets[OF lr_sets_ident refl]) lemma qbs_pred_iff_sets: "{x \space (qbs_to_measure X). P x} \ sets (qbs_to_measure X) \ qbs_pred X P" by (simp add: Measurable.pred_def lr_adjunction_correspondence space_L) lemma assumes [qbs]:"P \ X \\<^sub>Q Y \\<^sub>Q qbs_count_space UNIV" "f \ X \\<^sub>Q Y" shows indicator_qbs_morphism''': "(\x. indicator {y. P x y} (f x)) \ X \\<^sub>Q qbs_borel" (is ?g1) and indicator_qbs_morphism'': "(\x. indicator {y\qbs_space Y. P x y} (f x)) \ X \\<^sub>Q qbs_borel" (is ?g2) proof - have [simp]:"{x \ qbs_space X. P x (f x)} = {x \ qbs_space X. f x \ qbs_space Y \ P x (f x)}" using qbs_morphism_space[OF assms(2)] by blast show ?g1 ?g2 using qbs_morphism_app[OF assms,simplified qbs_pred_iff_sets[symmetric]] qbs_morphism_space[OF assms(2)] by(auto intro!: borel_measurable_indicator' simp: lr_adjunction_correspondence space_L) qed lemma assumes [qbs]:"P \ X \\<^sub>Q Y \\<^sub>Q qbs_count_space UNIV" shows indicator_qbs_morphism[qbs]:"(\x. indicator {y \ qbs_space Y. P x y}) \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" (is ?g1) and indicator_qbs_morphism':"(\x. indicator {y. P x y}) \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" (is ?g2) proof - note indicator_qbs_morphism''[qbs] indicator_qbs_morphism'''[qbs] show ?g1 ?g2 by(auto intro!: curry_preserves_morphisms[OF pair_qbs_morphismI] simp: qbs_Mx_is_morphisms) qed lemma indicator_qbs[qbs]: assumes "qbs_pred X P" shows "indicator {x. P x} \ X \\<^sub>Q qbs_borel" using assms by(auto simp: lr_adjunction_correspondence) lemma All_qbs_pred[qbs]: "qbs_pred (count_space\<^sub>Q (UNIV :: ('a :: countable) set) \\<^sub>Q count_space\<^sub>Q UNIV) All" proof(rule qbs_morphismI) fix a :: "real \ 'a \ bool" assume "a \ qbs_Mx (count_space\<^sub>Q UNIV \\<^sub>Q count_space\<^sub>Q UNIV)" hence [measurable]: "\f g. f \ borel_measurable borel \ g \ borel \\<^sub>M count_space UNIV \ (\x::real. a (f x) (g x)) \ borel \\<^sub>M count_space UNIV" by(auto simp add: exp_qbs_Mx qbs_Mx_R) show " All \ a \ qbs_Mx (count_space\<^sub>Q UNIV)" by(simp add: comp_def qbs_Mx_R) qed lemma Ex_qbs_pred[qbs]: "qbs_pred (count_space\<^sub>Q (UNIV :: ('a :: countable) set) \\<^sub>Q count_space\<^sub>Q UNIV) Ex" proof(rule qbs_morphismI) fix a :: "real \ 'a \ bool" assume "a \ qbs_Mx (count_space\<^sub>Q UNIV \\<^sub>Q count_space\<^sub>Q UNIV)" hence [measurable]: "\f g. f \ borel_measurable borel \ g \ borel \\<^sub>M count_space UNIV \ (\x::real. a (f x) (g x)) \ borel \\<^sub>M count_space UNIV" by(auto simp add: exp_qbs_Mx qbs_Mx_R) show "Ex \ a \ qbs_Mx (count_space\<^sub>Q UNIV)" by(simp add: comp_def qbs_Mx_R) qed lemma Ball_qbs_pred_countable: assumes "\i::'a :: countable. i \ I \ qbs_pred X (P i)" shows "qbs_pred X (\x. \x\I. P i x)" using assms by(simp add: qbs_pred_iff_measurable_pred) lemma Ball_qbs_pred: assumes "finite I" "\i. i \ I \ qbs_pred X (P i)" shows "qbs_pred X (\x. \x\I. P i x)" using assms by(simp add: qbs_pred_iff_measurable_pred) lemma Bex_qbs_pred_countable: assumes "\i::'a :: countable. i \ I \ qbs_pred X (P i)" shows "qbs_pred X (\x. \x\I. P i x)" using assms by(simp add: qbs_pred_iff_measurable_pred) lemma Bex_qbs_pred: assumes "finite I" "\i. i \ I \ qbs_pred X (P i)" shows "qbs_pred X (\x. \x\I. P i x)" using assms by(simp add: qbs_pred_iff_measurable_pred) lemma qbs_morphism_If_sub_qbs: assumes [qbs]: "qbs_pred X P" and [qbs]: "f \ sub_qbs X {x\qbs_space X. P x} \\<^sub>Q Y" "g \ sub_qbs X {x\qbs_space X. \ P x} \\<^sub>Q Y" shows "(\x. if P x then f x else g x) \ X \\<^sub>Q Y" proof(rule qbs_morphismI) fix \ assume h:"\ \ qbs_Mx X" interpret standard_borel_ne "borel :: real measure" by simp have [measurable]: "Measurable.pred borel (\x. P (\ x))" using h by(simp add: qbs_pred_iff_measurable_pred[symmetric] qbs_Mx_is_morphisms) consider "qbs_space X = {}" | "{x\qbs_space X. \ P x} = qbs_space X" | "{x\qbs_space X. P x} = qbs_space X" | "{x\qbs_space X. P x} \ {}" "{x\qbs_space X. \ P x} \ {}" by blast then show "(\x. if P x then f x else g x) \ \ \ qbs_Mx Y" (is "?f \ _") proof cases case 1 with h show ?thesis by(simp add: qbs_empty_equiv) next case 2 have [simp]:"(\x. if P x then f x else g x) \ \ = g \ \" by standard (use qbs_Mx_to_X[OF h] 2 in auto) show ?thesis using 2 qbs_morphism_Mx[OF assms(3)] h by(simp add: sub_qbs_ident) next case 3 have [simp]:"(\x. if P x then f x else g x) \ \ = f \ \" by standard (use qbs_Mx_to_X[OF h] 3 in auto) show ?thesis using 3 qbs_morphism_Mx[OF assms(2)] h by(simp add: sub_qbs_ident) next case 4 then obtain x0 x1 where x0:"x0 \ qbs_space X" "P x0" and x1:"x1 \ qbs_space X" "\ P x1" by blast define a0 where "a0 = (\r. if P (\ r) then \ r else x0)" define a1 where "a1 = (\r. if \ P (\ r) then \ r else x1)" have "a0 \ qbs_Mx (sub_qbs X {x\qbs_space X. P x})" "a1 \ qbs_Mx (sub_qbs X {x\qbs_space X. \ P x})" using x0 x1 qbs_Mx_to_X[OF h] h by(auto simp: sub_qbs_Mx a0_def a1_def intro!: qbs_closed3_dest2'[of UNIV "\r. P (\ r)" "\b r. if b then \ r else x0"]) (simp_all add: qbs_Mx_is_morphisms) from qbs_morphism_Mx[OF assms(2) this(1)] qbs_morphism_Mx[OF assms(3) this(2)] have h0:"(\r. f (a0 r)) \ qbs_Mx Y" "(\r. g (a1 r)) \ qbs_Mx Y" by (simp_all add: comp_def) have [simp]:"(\x. if P x then f x else g x) \ \ = (\r. if P (\ r) then f (a0 r) else g (a1 r))" by standard (auto simp: comp_def a0_def a1_def) show "(\x. if P x then f x else g x) \ \ \ qbs_Mx Y" using h h0 by(simp add: qbs_Mx_is_morphisms) qed qed subsubsection \ The Adjunction w.r.t. Ordering\ lemma l_mono: "mono qbs_to_measure" proof fix X Y :: "'a quasi_borel" show "X \ Y \ qbs_to_measure X \ qbs_to_measure Y" proof(induction rule: less_eq_quasi_borel.induct) case (1 X Y) then show ?case by(simp add: less_eq_measure.intros(1) space_L) next case (2 X Y) then have "sigma_Mx X \ sigma_Mx Y" by(auto simp add: sigma_Mx_def) then consider "sigma_Mx X \ sigma_Mx Y" | "sigma_Mx X = sigma_Mx Y" by auto then show ?case apply(cases) apply(rule less_eq_measure.intros(2)) apply(simp_all add: 2 space_L sets_L) by(rule less_eq_measure.intros(3),simp_all add: 2 sets_L space_L emeasure_L) qed qed lemma r_mono: "mono measure_to_qbs" proof fix M N :: "'a measure" show "M \ N \ measure_to_qbs M \ measure_to_qbs N" proof(induction rule: less_eq_measure.inducts) case (1 M N) then show ?case by(simp add: less_eq_quasi_borel.intros(1) qbs_space_R) next case (2 M N) then have "(borel :: real measure) \\<^sub>M N \ borel \\<^sub>M M" by(simp add: measurable_mono) then consider "(borel :: real measure) \\<^sub>M N \ borel \\<^sub>M M" | "(borel :: real measure) \\<^sub>M N = borel \\<^sub>M M" by auto then show ?case by cases (rule less_eq_quasi_borel.intros(2),simp_all add: 2 qbs_space_R qbs_Mx_R)+ next case (3 M N) then show ?case apply - by(rule less_eq_quasi_borel.intros(2)) (simp_all add: measurable_mono qbs_space_R qbs_Mx_R) qed qed lemma rl_order_adjunction: "X \ qbs_to_measure Y \ measure_to_qbs X \ Y" proof assume 1: "X \ qbs_to_measure Y" then show "measure_to_qbs X \ Y" proof(induction rule: less_eq_measure.cases) case (1 M N) then have [simp]:"qbs_space Y = space N" by(simp add: 1(2)[symmetric] space_L) show ?case by(rule less_eq_quasi_borel.intros(1),simp add: 1 qbs_space_R) next case (2 M N) then have [simp]:"qbs_space Y = space N" by(simp add: 2(2)[symmetric] space_L) show ?case proof(rule less_eq_quasi_borel.intros(2)) show "qbs_Mx Y \ qbs_Mx (measure_to_qbs X)" unfolding qbs_Mx_R proof fix \ assume h:"\ \ qbs_Mx Y" show "\ \ borel \\<^sub>M X" proof(rule measurableI) show "\x. \ x \ space X" using qbs_Mx_to_X[OF h] by (auto simp add: 2) next fix A assume "A \ sets X" then have "A \ sets (qbs_to_measure Y)" using 2 by auto then obtain U where hu:"A = U \ space N" "(\\\qbs_Mx Y. \ -` U \ sets borel)" by(auto simp add: sigma_Mx_def sets_L) have "\ -` A = \ -` U" using qbs_Mx_to_X[OF h] by(auto simp add: hu) thus "\ -` A \ space borel \ sets borel" using h hu(2) by simp qed qed qed(auto simp: 2 qbs_space_R) next case (3 M N) then have [simp]:"qbs_space Y = space N" by(simp add: 3(2)[symmetric] space_L) show ?case proof(rule less_eq_quasi_borel.intros(2)) show "qbs_Mx Y \ qbs_Mx (measure_to_qbs X)" unfolding qbs_Mx_R proof fix \ assume h:"\ \ qbs_Mx Y" show "\ \ borel \\<^sub>M X" proof(rule measurableI) show "\x. \ x \ space X" using qbs_Mx_to_X[OF h] by(auto simp: 3) next fix A assume "A \ sets X" then have "A \ sets (qbs_to_measure Y)" using 3 by auto then obtain U where hu:"A = U \ space N" "(\\\qbs_Mx Y. \ -` U \ sets borel)" by(auto simp add: sigma_Mx_def sets_L) have "\ -` A = \ -` U" using qbs_Mx_to_X[OF h] by(auto simp add: hu) thus "\ -` A \ space borel \ sets borel" using h hu(2) by simp qed qed qed(auto simp: 3 qbs_space_R) qed next assume "measure_to_qbs X \ Y" then show "X \ qbs_to_measure Y" proof(induction rule: less_eq_quasi_borel.cases) case (1 A B) have [simp]: "space X = qbs_space A" by(simp add: 1(1)[symmetric] qbs_space_R) show ?case by(rule less_eq_measure.intros(1)) (simp add: 1 space_L) next case (2 A B) then have hmy:"qbs_Mx Y \ borel \\<^sub>M X" using qbs_Mx_R by blast have [simp]: "space X = qbs_space A" by(simp add: 2(1)[symmetric] qbs_space_R) have "sets X \ sigma_Mx Y" proof fix U assume hu:"U \ sets X" show "U \ sigma_Mx Y" unfolding sigma_Mx_def proof(safe intro!: exI[where x=U]) show "\x. x \ U \ x \ qbs_space Y" using sets.sets_into_space[OF hu] by(auto simp add: 2) next fix \ assume "\ \ qbs_Mx Y" then have "\ \ borel \\<^sub>M X" using hmy by(auto) thus "\ -` U \ sets borel" using hu by(simp add: measurable_sets_borel) qed qed then consider "sets X = sigma_Mx Y" | "sets X \ sigma_Mx Y" by auto then show ?case proof cases case 1 show ?thesis proof(rule less_eq_measure.intros(3)) show "emeasure X \ emeasure (qbs_to_measure Y)" unfolding emeasure_L proof(rule le_funI) fix U consider "U = {}" | "U \ sigma_Mx Y" | "U \ {} \ U \ sigma_Mx Y" by auto then show "emeasure X U \ (if U = {} \ U \ sigma_Mx Y then 0 else \)" proof cases case 1 then show ?thesis by simp next case h:2 then have "U \ sigma_Mx A" using qbs_Mx_sigma_Mx_contra[OF 2(3)[symmetric] 2(4)] 2(2) by auto hence "U \ sets X" using lr_sets 2(1) sets_L by blast thus ?thesis by(simp add: h emeasure_notin_sets) next case 3 then show ?thesis by simp qed qed qed(simp_all add: 1 2 space_L sets_L) next case h2:2 show ?thesis by(rule less_eq_measure.intros(2)) (simp add: space_L 2, simp add: h2 sets_L) qed qed qed end \ No newline at end of file diff --git a/thys/S_Finite_Measure_Monad/Monad_QuasiBorel.thy b/thys/S_Finite_Measure_Monad/Monad_QuasiBorel.thy --- a/thys/S_Finite_Measure_Monad/Monad_QuasiBorel.thy +++ b/thys/S_Finite_Measure_Monad/Monad_QuasiBorel.thy @@ -1,3501 +1,3512 @@ (* Title: Monad_QuasiBorel.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) section \The S-Finite Measure Monad\ theory Monad_QuasiBorel imports "Measure_QuasiBorel_Adjunction" "Kernels" begin subsection \ The S-Finite Measure Monad\ subsubsection \ Space of S-Finite Measures\ locale in_Mx = fixes X :: "'a quasi_borel" and \ :: "real \ 'a" assumes in_Mx[simp]:"\ \ qbs_Mx X" begin lemma \_measurable[measurable]: "\ \ borel \\<^sub>M qbs_to_measure X" using in_Mx qbs_Mx_subset_of_measurable by blast lemma \_qbs_morphism[qbs]: "\ \ qbs_borel \\<^sub>Q X" using in_Mx by(simp only: qbs_Mx_is_morphisms) lemma X_not_empty: "qbs_space X \ {}" using in_Mx by(auto simp: qbs_empty_equiv simp del: in_Mx) lemma inverse_UNIV[simp]: "\ -` (qbs_space X) = UNIV" by fastforce end locale qbs_s_finite = in_Mx X \ + s_finite_measure \ for X :: "'a quasi_borel" and \ and \ :: "real measure" + assumes mu_sets[measurable_cong]: "sets \ = sets borel" begin lemma mu_not_empty: "space \ \ {}" by(simp add: sets_eq_imp_space_eq[OF mu_sets]) end lemma qbs_s_finite_All: assumes "\ \ qbs_Mx X" "s_finite_kernel M borel k" "x \ space M" shows "qbs_s_finite X \ (k x)" proof - interpret s_finite_kernel M borel k by fact show ?thesis using assms(1,3) image_s_finite_measure[OF assms(3)] by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def kernel_sets) qed locale qbs_prob = in_Mx X \ + real_distribution \ for X :: "'a quasi_borel" and \ \ begin lemma qbs_s_finite: "qbs_s_finite X \ \" by(auto simp: qbs_s_finite_def qbs_s_finite_axioms_def in_Mx_def s_finite_measure_prob) sublocale qbs_s_finite by(rule qbs_s_finite) end lemma(in qbs_s_finite) qbs_probI: "prob_space \ \ qbs_prob X \ \" by(auto simp: qbs_prob_def in_Mx_def real_distribution_def real_distribution_axioms_def mu_sets) locale pair_qbs_s_finites = pq1: qbs_s_finite X \ \ + pq2: qbs_s_finite Y \ \ for X :: "'a quasi_borel" and \ \ and Y :: "'b quasi_borel" and \ \ begin lemma ab_measurable[measurable]: "map_prod \ \ \ borel \\<^sub>M borel \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" proof - have "map_prod \ \ \ qbs_to_measure (measure_to_qbs (borel \\<^sub>M borel)) \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" by(auto intro!: set_mp[OF l_preserves_morphisms] simp: r_preserves_product) moreover have "sets (qbs_to_measure (measure_to_qbs (borel \\<^sub>M borel))) = sets ((borel \\<^sub>M borel) :: (real \ real) measure)" by(auto intro!: standard_borel.lr_sets_ident pair_standard_borel_ne standard_borel_ne.standard_borel) ultimately show ?thesis by simp qed end locale pair_qbs_probs = pq1: qbs_prob X \ \ + pq2: qbs_prob Y \ \ for X :: "'a quasi_borel" and \ \ and Y :: "'b quasi_borel" and \ \ begin sublocale pair_qbs_s_finites by standard end locale pair_qbs_s_finite = pq1: qbs_s_finite X \ \ + pq2: qbs_s_finite X \ \ for X :: "'a quasi_borel" and \ \ and \ \ begin sublocale pair_qbs_s_finites X \ \ X \ \ by standard end locale pair_qbs_prob = pq1: qbs_prob X \ \ + pq2: qbs_prob X \ \ for X :: "'a quasi_borel" and \ \ and \ \ begin sublocale pair_qbs_s_finite X \ \ \ \ by standard sublocale pair_qbs_probs X \ \ X \ \ by standard end type_synonym 'a qbs_s_finite_t = "'a quasi_borel * (real \ 'a) * real measure" definition qbs_s_finite_eq :: "['a qbs_s_finite_t, 'a qbs_s_finite_t] \ bool" where "qbs_s_finite_eq p1 p2 \ (let (X, \, \) = p1; (Y, \, \) = p2 in qbs_s_finite X \ \ \ qbs_s_finite Y \ \ \ X = Y \ distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure Y) \)" definition qbs_s_finite_eq' :: "['a qbs_s_finite_t, 'a qbs_s_finite_t] \ bool" where "qbs_s_finite_eq' p1 p2 \ (let (X, \, \) = p1; (Y, \, \) = p2 in qbs_s_finite X \ \ \ qbs_s_finite Y \ \ \ X = Y \ (\f\X \\<^sub>Q (qbs_borel :: ennreal quasi_borel). (\\<^sup>+x. f (\ x) \\) = (\\<^sup>+x. f (\ x) \\)))" lemma(in qbs_s_finite) shows qbs_s_finite_eq_refl[simp]: "qbs_s_finite_eq (X,\,\) (X,\,\)" and qbs_s_finite_eq'_refl[simp]: "qbs_s_finite_eq' (X,\,\) (X,\,\)" by(simp_all add: qbs_s_finite_eq_def qbs_s_finite_eq'_def qbs_s_finite_axioms) lemma(in pair_qbs_s_finite) shows qbs_s_finite_eq_intro: "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \ \ qbs_s_finite_eq (X,\,\) (X,\,\)" and qbs_s_finite_eq'_intro: "(\f. f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)) \ qbs_s_finite_eq' (X,\,\) (X,\,\)" by(simp_all add: qbs_s_finite_eq_def qbs_s_finite_eq'_def pq1.qbs_s_finite_axioms pq2.qbs_s_finite_axioms) lemma qbs_s_finite_eq_dest: assumes "qbs_s_finite_eq (X,\,\) (Y,\,\)" shows "qbs_s_finite X \ \" "qbs_s_finite Y \ \" "Y = X" "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" using assms by(auto simp: qbs_s_finite_eq_def) lemma qbs_s_finite_eq'_dest: assumes "qbs_s_finite_eq' (X,\,\) (Y,\,\)" shows "qbs_s_finite X \ \" "qbs_s_finite Y \ \" "Y = X" "\f. f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)" using assms by(auto simp: qbs_s_finite_eq'_def) lemma(in qbs_prob) qbs_s_finite_eq_qbs_prob_cong: assumes "qbs_s_finite_eq (X,\,\) (Y,\,\)" shows "qbs_prob Y \ \" proof - interpret qs: pair_qbs_s_finites X \ \ Y \ \ using assms(1) by(auto simp: qbs_s_finite_eq_def pair_qbs_s_finites_def) show ?thesis by(auto intro!: qs.pq2.qbs_probI prob_space_distrD[of \ _ "qbs_to_measure Y"]) (auto simp: qbs_s_finite_eq_dest(3)[OF assms] qbs_s_finite_eq_dest(4)[OF assms,symmetric] intro!: prob_space_distr) qed lemma shows qbs_s_finite_eq_symp: "symp qbs_s_finite_eq" and qbs_s_finite_eq_transp: "transp qbs_s_finite_eq" by(simp_all add: qbs_s_finite_eq_def transp_def symp_def) quotient_type 'a qbs_measure = "'a qbs_s_finite_t" / partial: qbs_s_finite_eq morphisms rep_qbs_measure qbs_measure proof(rule part_equivpI) let ?U = "UNIV :: 'a set" let ?Uf = "UNIV :: (real \ 'a) set" let ?f = "(\_. undefined) :: real \ 'a" have "qbs_s_finite (Abs_quasi_borel (?U, ?Uf)) ?f (return borel 0)" unfolding qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def proof safe have "Rep_quasi_borel (Abs_quasi_borel (?U,?Uf)) = (?U, ?Uf)" using Abs_quasi_borel_inverse by (auto simp add: qbs_closed1_def qbs_closed2_def qbs_closed3_def is_quasi_borel_def) thus "(\_. undefined) \ qbs_Mx (Abs_quasi_borel (?U, ?Uf))" by(simp add: qbs_Mx_def) next show "s_finite_measure (return borel 0)" by(auto intro!: sigma_finite_measure.s_finite_measure prob_space_imp_sigma_finite prob_space_return) qed simp_all thus "\x :: 'a qbs_s_finite_t. qbs_s_finite_eq x x" by(auto simp: qbs_s_finite_eq_def intro!: exI[where x="(Abs_quasi_borel (?U,?Uf), ?f, return borel 0)"]) qed(simp_all add: qbs_s_finite_eq_symp qbs_s_finite_eq_transp) interpretation qbs_measure : quot_type "qbs_s_finite_eq" "Abs_qbs_measure" "Rep_qbs_measure" using Abs_qbs_measure_inverse Rep_qbs_measure by(simp add: quot_type_def equivp_implies_part_equivp qbs_measure_equivp Rep_qbs_measure_inverse Rep_qbs_measure_inject) blast syntax "_qbs_measure" :: "'a quasi_borel \ (real \ 'a) \ real measure \ 'a qbs_measure" ("\_,/ _,/ _\\<^sub>s\<^sub>f\<^sub>i\<^sub>n") translations "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" \ "CONST qbs_measure (X, \, \)" lemma rep_qbs_s_finite_measure': "\X \ \. p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ qbs_s_finite X \ \" by(rule qbs_measure.abs_induct,auto simp add: qbs_s_finite_eq_def) lemma rep_qbs_s_finite_measure: obtains X \ \ where "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" using that rep_qbs_s_finite_measure' by blast definition qbs_null_measure :: "'a quasi_borel \ 'a qbs_measure" where "qbs_null_measure X \ \X, SOME a. a \ qbs_Mx X, null_measure borel\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" lemma qbs_null_measure_s_finite: "qbs_space X \ {} \ qbs_s_finite X (SOME a. a \ qbs_Mx X) (null_measure borel)" by(auto simp: qbs_empty_equiv qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def some_in_eq intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) lemma(in qbs_s_finite) in_Rep_qbs_measure': assumes "qbs_s_finite_eq (X,\,\) (X',\',\')" shows "(X',\',\') \ Rep_qbs_measure \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by (metis assms mem_Collect_eq qbs_s_finite_eq_refl qbs_measure_def qbs_measure.abs_def qbs_measure.abs_inverse) lemmas(in qbs_s_finite) in_Rep_qbs_measure = in_Rep_qbs_measure'[OF qbs_s_finite_eq_refl] lemma(in qbs_s_finite) if_in_Rep_qbs_measure: assumes "(X',\',\') \ Rep_qbs_measure \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" shows "X' = X" "qbs_s_finite X' \' \'" "qbs_s_finite_eq (X,\,\) (X',\',\')" proof - show h:"X' = X" using assms qbs_measure.Rep_qbs_measure[of "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n"] by auto (metis mem_Collect_eq qbs_s_finite_eq_dest(3) qbs_s_finite_eq_refl qbs_measure_def qbs_measure.abs_def qbs_measure.abs_inverse) next show "qbs_s_finite X' \' \'" using assms qbs_measure.Rep_qbs_measure[of "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n"] by (auto simp: qbs_s_finite_eq_dest(2)) next show "qbs_s_finite_eq (X,\,\) (X',\',\')" using assms qbs_measure.Rep_qbs_measure[of "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n"] by auto (metis mem_Collect_eq qbs_s_finite_eq_dest(3) qbs_s_finite_eq_refl qbs_measure_def qbs_measure.abs_def qbs_measure.abs_inverse) qed lemma qbs_s_finite_eq_1_imp_2: assumes "qbs_s_finite_eq (X,\,\) (Y,\,\)" "f \ X \\<^sub>Q (qbs_borel :: (_ :: {banach}) quasi_borel)" shows "(\x. f (\ x) \\) = (\x. f (\ x) \\)" (is "?lhs = ?rhs") proof - interpret pq : pair_qbs_s_finite X \ \ \ \ using assms by(auto intro!: pair_qbs_s_finite.intro simp: qbs_s_finite_eq_def) have [measurable]: "f \ qbs_to_measure X \\<^sub>M borel" using assms by(simp add: lr_adjunction_correspondence) have "?lhs = (\x. f x \(distr \ (qbs_to_measure X) \))" by(simp add: integral_distr) also have "... = (\x. f x \(distr \ (qbs_to_measure X) \))" by(simp add: qbs_s_finite_eq_dest(4)[OF assms(1)]) also have "... = ?rhs" by(simp add: integral_distr) finally show ?thesis . qed lemma qbs_s_finite_eq_equiv: "qbs_s_finite_eq = qbs_s_finite_eq'" proof(rule ext[OF ext]) show "\a b :: 'a qbs_s_finite_t. qbs_s_finite_eq a b = qbs_s_finite_eq' a b" proof safe fix X Y :: "'a quasi_borel" and \ \ \ \ { assume h:"qbs_s_finite_eq (X,\,\) (Y,\,\)" then interpret pq : pair_qbs_s_finite X \ \ \ \ by(auto intro!: pair_qbs_s_finite.intro simp: qbs_s_finite_eq_def) show "qbs_s_finite_eq' (X,\,\) (Y,\,\)" unfolding qbs_s_finite_eq_dest(3)[OF h] proof(rule pq.qbs_s_finite_eq'_intro) fix f :: "'a \ ennreal" assume f:"f \ X \\<^sub>Q qbs_borel" show "(\\<^sup>+ x. f (\ x) \\) = (\\<^sup>+ x. f (\ x) \\)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. f x \(distr \ (qbs_to_measure X) \))" by(rule nn_integral_distr[symmetric]) (use f lr_adjunction_correspondence in auto) also have "... = (\\<^sup>+ x. f x \(distr \ (qbs_to_measure X) \))" by(simp add: qbs_s_finite_eq_dest(4)[OF h]) also have "... = ?rhs" by(rule nn_integral_distr) (use f lr_adjunction_correspondence in auto) finally show ?thesis . qed qed } { assume h:"qbs_s_finite_eq' (X,\,\) (Y,\,\)" then interpret pq : pair_qbs_s_finite X \ \ \ \ by(auto intro!: pair_qbs_s_finite.intro simp: qbs_s_finite_eq'_def) show "qbs_s_finite_eq (X,\,\) (Y,\,\)" unfolding qbs_s_finite_eq'_dest(3)[OF h] proof(rule pq.qbs_s_finite_eq_intro[OF measure_eqI]) fix U assume hu[measurable]:"U \ sets (distr \ (qbs_to_measure X) \)" show "emeasure (distr \ (qbs_to_measure X) \) U = emeasure (distr \ (qbs_to_measure X) \) U" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. indicator U x \ (distr \ (qbs_to_measure X) \))" using hu by simp also have "... = (\\<^sup>+ x. indicator U (\ x) \\)" by(rule nn_integral_distr) (use hu in auto) also have "... = (\\<^sup>+ x. indicator U (\ x) \\)" by(auto intro!: qbs_s_finite_eq'_dest(4)[OF h] simp: lr_adjunction_correspondence) also have "... = (\\<^sup>+ x. indicator U x \ (distr \ (qbs_to_measure X) \))" by(rule nn_integral_distr[symmetric]) (use hu in auto) also have "... = ?rhs" using hu by simp finally show ?thesis . qed qed simp } qed qed lemma qbs_s_finite_measure_eq: "qbs_s_finite_eq (X,\,\) (Y,\,\) \ \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using Quotient3_rel[OF Quotient3_qbs_measure] by blast lemma(in pair_qbs_s_finite) qbs_s_finite_measure_eq: "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \ \ \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto intro!: qbs_s_finite_measure_eq qbs_s_finite_eq_intro) lemma(in pair_qbs_s_finite) qbs_s_finite_measure_eq': "(\f. f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)) \ \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite_eq'_intro[simplified qbs_s_finite_eq_equiv[symmetric]] by(auto intro!: qbs_s_finite_measure_eq simp: qbs_s_finite_eq_def) lemma(in pair_qbs_s_finite) qbs_s_finite_measure_eq_inverse: assumes "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" shows "qbs_s_finite_eq (X,\,\) (X,\,\)" "qbs_s_finite_eq' (X,\,\) (X,\,\)" using Quotient3_rel[OF Quotient3_qbs_measure,of "(X,\,\)" "(X,\,\)",simplified] by(simp_all add: assms qbs_s_finite_eq_equiv) lift_definition qbs_space_of :: "'a qbs_measure \ 'a quasi_borel" is fst by(auto simp: qbs_s_finite_eq_def) lemma(in qbs_s_finite) qbs_space_of[simp]: "qbs_space_of \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = X" by(simp add: qbs_space_of.abs_eq) lemma rep_qbs_space_of: assumes "qbs_space_of s = X" shows "\\ \. s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ qbs_s_finite X \ \" proof - obtain X' \ \ where hs: "s = \X', \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X' \ \" using rep_qbs_s_finite_measure'[of s] by auto then interpret qs:qbs_s_finite X' \ \ by simp show ?thesis using assms hs(2) by(auto simp add: hs(1)) qed corollary qbs_s_space_of_not_empty: "qbs_space (qbs_space_of X) \ {}" by transfer (auto simp: qbs_s_finite_eq_def qbs_s_finite_def in_Mx_def qbs_empty_equiv) subsubsection \ The S-Finite Measure Monad\ definition monadM_qbs :: "'a quasi_borel \ 'a qbs_measure quasi_borel" where "monadM_qbs X \ Abs_quasi_borel ({s. qbs_space_of s = X}, {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k})" lemma shows monadM_qbs_space: "qbs_space (monadM_qbs X) = {s. qbs_space_of s = X}" and monadM_qbs_Mx: "qbs_Mx (monadM_qbs X) = {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" proof - have "{\r::real. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k} \ UNIV \ {s. qbs_space_of s = X}" proof safe fix x \ and k :: "real \ real measure" assume h:"\ \ qbs_Mx X" "s_finite_kernel borel borel k" interpret k:s_finite_kernel borel borel k by fact interpret qbs_s_finite X \ "k x" using k.image_s_finite_measure h(1) by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def k.kernel_sets) show "qbs_space_of \X, \, k x\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = X" by simp qed moreover have "qbs_closed1 {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" proof(safe intro!: qbs_closed1I) fix \ and f :: "real \ real" and k :: "real\ real measure" assume h:"f \ borel_measurable borel" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" then show "\\' ka. (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ f = (\r. \X, \', ka r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \' \ qbs_Mx X \ s_finite_kernel borel borel ka" by(auto intro!: exI[where x=\] exI[where x="\x. k (f x)"] simp: s_finite_kernel.comp_measurable[OF h(3,1)]) qed moreover have "qbs_closed2 {s. qbs_space_of s = X} {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" proof(safe intro!: qbs_closed2I) fix s assume h:"X = qbs_space_of s" from rep_qbs_space_of[OF this[symmetric]] obtain \ \ where s:"s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by auto then interpret qbs_s_finite X \ \ by simp show "\\ k. (\r. s) = (\r. \qbs_space_of s, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx (qbs_space_of s) \ s_finite_kernel borel borel k" by(auto intro!: exI[where x=\] exI[where x="\r. \"] s_finite_kernel_const simp: s(1) s_finite_kernel_cong_sets[OF _ mu_sets[symmetric]] sets_eq_imp_space_eq[OF mu_sets]) qed moreover have "qbs_closed3 {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" proof(safe intro!: qbs_closed3I) fix P :: "real \ nat" and Fi :: "nat \ _" assume P[measurable]: "P \ borel \\<^sub>M count_space UNIV" and "\i. Fi i \ {\r::real. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" then obtain \i ki where Fi: "\i. Fi i = (\r. \X, \i i, ki i r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\i. \i i \ qbs_Mx X" "\i. s_finite_kernel borel borel (ki i)" by auto metis interpret nat_real: standard_borel_ne "count_space (UNIV :: nat set) \\<^sub>M (borel :: real measure)" by(auto intro!: pair_standard_borel_ne) note [simp] = nat_real.from_real_to_real[simplified space_pair_measure, simplified] define \ where "\ \ (\r. case_prod \i (nat_real.from_real r))" define k where "k \ (\r. distr (distr (ki (P r) r) (count_space UNIV \\<^sub>M borel) (\r'. (P r, r'))) borel nat_real.to_real)" have \: "\ \ qbs_Mx X" unfolding \_def qbs_Mx_is_morphisms proof(rule qbs_morphism_compose[where g=nat_real.from_real and Y="qbs_count_space UNIV \\<^sub>Q qbs_borel"]) show "nat_real.from_real \ qbs_borel \\<^sub>Q qbs_count_space UNIV \\<^sub>Q qbs_borel" by(simp add: r_preserves_product[symmetric] standard_borel.standard_borel_r_full_faithful[of "borel :: real measure",simplified,symmetric] standard_borel_ne.standard_borel) next show "case_prod \i \ qbs_count_space UNIV \\<^sub>Q qbs_borel \\<^sub>Q X" using Fi(2) by(auto intro!: qbs_morphism_pair_count_space1 simp: qbs_Mx_is_morphisms) qed have sets_ki[measurable_cong]: "sets (ki i r) = sets borel" "sets (k r) = sets borel" for i r using Fi(3) by(auto simp: s_finite_kernel_def measure_kernel_def k_def) interpret k:s_finite_kernel borel borel k proof - have 1:"k = (\(i,r). distr (ki i r) borel (\r'. nat_real.to_real (i, r'))) \ (\r. (P r, r))" by standard (auto simp: k_def distr_distr comp_def) have "s_finite_kernel borel borel ..." unfolding comp_def by(rule s_finite_kernel.comp_measurable[where X="count_space UNIV \\<^sub>M borel"],rule s_finite_kernel_pair_countble1, auto intro!: s_finite_kernel.distr_s_finite_kernel[OF Fi(3)]) thus "s_finite_kernel borel borel k" by(simp add: 1) qed have "(\r. Fi (P r) r) = (\r. \X, \, k r \\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" unfolding Fi(1) proof fix r interpret pq:pair_qbs_s_finite X "\i (P r)" "ki (P r) r" \ "k r" by(auto simp: pair_qbs_s_finite_def qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def k.image_s_finite_measure s_finite_kernel.image_s_finite_measure[OF Fi(3)] sets_ki \ Fi(2)) show "\X, \i (P r), ki (P r) r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(rule pq.qbs_s_finite_measure_eq, simp add: k_def distr_distr comp_def,simp add: \_def) qed thus "\\ k. (\r. Fi (P r) r) = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx X \ s_finite_kernel borel borel k" by(auto intro!: exI[where x=\] exI[where x=k] simp: \ k.s_finite_kernel_axioms) qed ultimately have "Rep_quasi_borel (monadM_qbs X) = ({s. qbs_space_of s = X}, {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k})" by(auto intro!: Abs_quasi_borel_inverse simp: monadM_qbs_def is_quasi_borel_def) thus "qbs_space (monadM_qbs X) = {s. qbs_space_of s = X}" "qbs_Mx (monadM_qbs X) = {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ s_finite_kernel borel borel k}" by(simp_all add: qbs_space_def qbs_Mx_def) qed lemma monadM_qbs_empty_iff: "qbs_space X = {} \ qbs_space (monadM_qbs X) = {}" by(auto simp: monadM_qbs_space qbs_s_space_of_not_empty) (meson in_Mx.intro qbs_closed2_dest qbs_s_finite.qbs_space_of qbs_s_finite_def rep_qbs_s_finite_measure') lemma(in qbs_s_finite) in_space_monadM[qbs]: "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ qbs_space (monadM_qbs X)" by(simp add: monadM_qbs_space) lemma rep_qbs_space_monadM: assumes "s \ qbs_space (monadM_qbs X)" obtains \ \ where "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" using rep_qbs_space_of assms that by(auto simp: monadM_qbs_space) lemma rep_qbs_space_monadM_sigma_finite: assumes "s \ qbs_space (monadM_qbs X)" obtains \ \ where "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" "sigma_finite_measure \" proof - obtain \ \ where s:"s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by(metis rep_qbs_space_monadM assms) hence "standard_borel_ne \""s_finite_measure \" by(auto intro!: standard_borel_ne_sets[of borel \] simp: qbs_s_finite_def qbs_s_finite_axioms_def) from exists_push_forward[OF this] obtain \' f where f: "f \ (borel :: real measure) \\<^sub>M \" "sets \' = sets borel" "sigma_finite_measure \'" "distr \' \ f = \" by metis hence [measurable]: "f \ borel_measurable borel" using s(2) by(auto simp: qbs_s_finite_def qbs_s_finite_axioms_def cong: measurable_cong_sets) interpret pair_qbs_s_finite X \ \ "\ \ f" \' proof - have "qbs_s_finite X (\ \ f) \'" using s(2) by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def[of \'] f(2,3) sigma_finite_measure.s_finite_measure) thus "pair_qbs_s_finite X \ \ (\ \ f) \'" by(auto simp: pair_qbs_s_finite_def s(2)) qed have "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \ \ f, \'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - have [simp]:" distr \ (qbs_to_measure X) \ = distr (distr \' \ f) (qbs_to_measure X) \" by(simp add: f(4)) show ?thesis by(auto intro!: qbs_s_finite_measure_eq simp: distr_distr) qed with s(1) pq2.qbs_s_finite_axioms f(3) that show ?thesis by metis qed lemma qbs_space_of_in: "s \ qbs_space (monadM_qbs X) \ qbs_space_of s = X" by(simp add: monadM_qbs_space) lemma in_qbs_space_of: "s \ qbs_space (monadM_qbs (qbs_space_of s))" by(simp add: monadM_qbs_space) subsubsection \ $l$ \ lift_definition qbs_l :: "'a qbs_measure \ 'a measure" is "\p. distr (snd (snd p)) (qbs_to_measure (fst p)) (fst (snd p))" by(auto simp: qbs_s_finite_eq_def) lemma(in qbs_s_finite) qbs_l: "qbs_l \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = distr \ (qbs_to_measure X) \" by(simp add: qbs_l.abs_eq) interpretation qbs_l_s_finite: s_finite_measure "qbs_l (s :: 'a qbs_measure)" proof(transfer) show "\s:: 'a qbs_s_finite_t. qbs_s_finite_eq s s \ s_finite_measure (distr (snd (snd s)) (qbs_to_measure (fst s)) (fst (snd s)))" proof safe fix X :: "'a quasi_borel" fix \ \ assume "qbs_s_finite_eq (X,\,\) (X,\,\)" then interpret qbs_s_finite X \ \ by(simp add: qbs_s_finite_eq_def) show "s_finite_measure (distr (snd (snd (X,\,\))) (qbs_to_measure (fst (X,\,\))) (fst (snd (X,\,\))))" by(auto intro!: s_finite_measure.s_finite_measure_distr simp: s_finite_measure_axioms) qed qed lemma space_qbs_l: "qbs_space (qbs_space_of s) = space (qbs_l s)" by(transfer, auto simp: space_L) lemma space_qbs_l_ne: "space (qbs_l s) \ {}" by transfer (auto simp: qbs_s_finite_eq_def qbs_s_finite_def in_Mx_def space_L qbs_empty_equiv) lemma qbs_l_sets: "sets (qbs_to_measure (qbs_space_of s)) = sets (qbs_l s)" by(transfer,simp) lemma qbs_null_measure_in_Mx: "qbs_space X \ {} \ qbs_null_measure X \ qbs_space (monadM_qbs X)" by(simp add: qbs_s_finite.in_space_monadM[OF qbs_null_measure_s_finite] qbs_null_measure_def) lemma qbs_null_measure_null_measure:"qbs_space X \ {} \ qbs_l (qbs_null_measure X) = null_measure (qbs_to_measure X)" by(auto simp: qbs_null_measure_def qbs_s_finite.qbs_l[OF qbs_null_measure_s_finite] null_measure_distr) lemma space_qbs_l_in: assumes "s \ qbs_space (monadM_qbs X)" shows "space (qbs_l s) = qbs_space X" by (metis assms qbs_s_finite.qbs_space_of rep_qbs_space_monadM space_qbs_l) lemma sets_qbs_l: assumes "s \ qbs_space (monadM_qbs X)" shows "sets (qbs_l s) = sets (qbs_to_measure X)" using assms qbs_l_sets qbs_space_of_in by blast lemma measurable_qbs_l: assumes "s \ qbs_space (monadM_qbs X)" shows "qbs_l s \\<^sub>M M = X \\<^sub>Q measure_to_qbs M" by(auto simp: measurable_cong_sets[OF qbs_l_sets[of s,simplified qbs_space_of_in[OF assms(1)],symmetric] refl] lr_adjunction_correspondence) lemma measurable_qbs_l': assumes "s \ qbs_space (monadM_qbs X)" shows "qbs_l s \\<^sub>M M = qbs_to_measure X \\<^sub>M M" by(simp add: measurable_qbs_l[OF assms] lr_adjunction_correspondence) lemma rep_qbs_Mx_monadM: assumes "\ \ qbs_Mx (monadM_qbs X)" obtains \ k where "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" proof - have "\\ r k. \ \ qbs_Mx X \ s_finite_kernel borel borel k \ qbs_s_finite X \ (k r)" by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def s_finite_kernel.image_s_finite_measure) (auto simp: s_finite_kernel_def measure_kernel_def) thus ?thesis using that assms by(fastforce simp: monadM_qbs_Mx) qed lemma qbs_l_measurable[measurable]:"qbs_l \ qbs_to_measure (monadM_qbs X) \\<^sub>M s_finite_measure_algebra (qbs_to_measure X)" proof(rule qbs_morphism_dest[OF qbs_morphismI]) fix \ assume "\ \ qbs_Mx (monadM_qbs X)" from rep_qbs_Mx_monadM[OF this] obtain \ k where h: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis show "qbs_l \ \ \ qbs_Mx (measure_to_qbs (s_finite_measure_algebra (qbs_to_measure X)))" by(auto simp add: qbs_Mx_R comp_def h(1) qbs_s_finite.qbs_l[OF h(4)] h(2,3) intro!: s_finite_kernel.kernel_measurable_s_finite s_finite_kernel.distr_s_finite_kernel[where Y=borel]) qed lemma qbs_l_measure_kernel: "measure_kernel (qbs_to_measure (monadM_qbs X)) (qbs_to_measure X) qbs_l" proof(cases "qbs_space X = {}") case True with monadM_qbs_empty_iff[of X,simplified this] show ?thesis by(auto intro!: measure_kernel_empty_trivial simp: space_L) next case 1:False show ?thesis proof show "\x. x \ space (qbs_to_measure (monadM_qbs X)) \ sets (qbs_l x) = sets (qbs_to_measure X)" using qbs_l_sets by(auto simp: space_L monadM_qbs_space) next show "space (qbs_to_measure X) \ {}" by(simp add: space_L 1) qed (rule measurable_emeasure_kernel_s_finite_measure_algebra[OF qbs_l_measurable]) qed lemma qbs_l_inj: "inj_on qbs_l (qbs_space (monadM_qbs X))" by standard (auto simp: monadM_qbs_space, transfer,auto simp: qbs_s_finite_eq_def) lemma qbs_l_morphism: assumes [measurable]:"A \ sets (qbs_to_measure X)" shows "(\s. qbs_l s A) \ monadM_qbs X \\<^sub>Q qbs_borel" proof(rule qbs_morphismI) fix \ assume h:"\ \ qbs_Mx (monadM_qbs X)" hence [qbs]: "\ \ qbs_borel \\<^sub>Q monadM_qbs X" by(simp_all add: qbs_Mx_is_morphisms) from rep_qbs_Mx_monadM[OF h(1)] obtain \ k where hk: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis then interpret a: in_Mx X \ by(simp add: in_Mx_def) have k[measurable_cong]:"sets (k r) = sets borel" for r using hk(3) by(auto simp: s_finite_kernel_def measure_kernel_def) show "(\s. emeasure (qbs_l s) A) \ \ \ qbs_Mx qbs_borel" by(auto simp: hk(1) qbs_s_finite.qbs_l[OF hk(4)] comp_def qbs_Mx_qbs_borel emeasure_distr sets_eq_imp_space_eq[OF k] intro!: s_finite_kernel.emeasure_measurable'[OF hk(3)] measurable_sets_borel[OF _ assms]) qed lemma qbs_l_finite_pred: "qbs_pred (monadM_qbs X) (\s. finite_measure (qbs_l s))" proof - have "qbs_space X \ sets (qbs_to_measure X)" by (metis sets.top space_L) note qbs_l_morphism[OF this,qbs] have [simp]:"finite_measure (qbs_l s) \ qbs_l s X \ \" if "s \ monadM_qbs X" for s by(auto intro!: finite_measureI dest: finite_measure.emeasure_finite simp: space_qbs_l_in[OF that]) show ?thesis by(simp cong: qbs_morphism_cong) qed lemma qbs_l_subprob_pred: "qbs_pred (monadM_qbs X) (\s. subprob_space (qbs_l s))" proof - have "qbs_space X \ sets (qbs_to_measure X)" by (metis sets.top space_L) note qbs_l_morphism[OF this,qbs] have [simp]:"subprob_space (qbs_l s) \ qbs_l s X \ 1" if "s \ monadM_qbs X" for s by(auto intro!: subprob_spaceI dest: subprob_space.subprob_emeasure_le_1 simp: space_qbs_l_ne) (simp add: space_qbs_l_in[OF that]) show ?thesis by(simp cong: qbs_morphism_cong) qed lemma qbs_l_prob_pred: "qbs_pred (monadM_qbs X) (\s. prob_space (qbs_l s))" proof - have "qbs_space X \ sets (qbs_to_measure X)" by (metis sets.top space_L) note qbs_l_morphism[OF this,qbs] have [simp]:"prob_space (qbs_l s) \ qbs_l s X = 1" if "s \ monadM_qbs X" for s by(auto intro!: prob_spaceI simp: space_qbs_l_ne) (auto simp add: space_qbs_l_in[OF that] dest: prob_space.emeasure_space_1) show ?thesis by(simp cong: qbs_morphism_cong) qed subsubsection \ Return \ definition return_qbs :: "'a quasi_borel \ 'a \ 'a qbs_measure" where "return_qbs X x \ \X, \r. x, SOME \. real_distribution \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" lemma(in real_distribution) assumes "x \ qbs_space X" shows return_qbs:"return_qbs X x = \X, \r. x, M\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and return_qbs_prob:"qbs_prob X (\r. x) M" and return_qbs_s_finite:"qbs_s_finite X (\r. x) M" proof - interpret qs1: qbs_prob X "\r. x" M by(auto simp: qbs_prob_def in_Mx_def real_distribution_axioms intro!: qbs_closed2_dest assms) show "return_qbs X x = \X, \r. x, M\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" unfolding return_qbs_def proof(rule someI2) show "real_distribution (return borel 0)" by (auto simp: real_distribution_def real_distribution_axioms_def,rule prob_space_return) simp next fix N assume "real_distribution N" then interpret qs2: qbs_s_finite X "\r. x" N by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def real_distribution_def real_distribution_axioms_def intro!: qbs_closed2_dest assms sigma_finite_measure.s_finite_measure prob_space_imp_sigma_finite) interpret pair_qbs_s_finite X "\r. x" N "\r. x" M by standard show "\X, \r. x, N\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X, \r. x, M\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto intro!: qbs_s_finite_measure_eq measure_eqI simp: emeasure_distr) (metis \real_distribution N\ emeasure_space_1 prob_space.emeasure_space_1 qs2.mu_sets real_distribution.axioms(1) sets_eq_imp_space_eq space_borel space_eq_univ) qed show "qbs_prob X (\r. x) M" "qbs_s_finite X (\r. x) M" by(simp_all add: qs1.qbs_prob_axioms qs1.qbs_s_finite_axioms) qed lemma return_qbs_comp: assumes "\ \ qbs_Mx X" shows "(return_qbs X \ \) = (\r. \X, \, return borel r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" proof fix r interpret pqp: pair_qbs_prob X "\k. \ r" "return borel 0" \ "return borel r" by(simp add: assms qbs_Mx_to_X[OF assms] pair_qbs_prob_def qbs_prob_def in_Mx_def real_distribution_def real_distribution_axioms_def prob_space_return) show "(return_qbs X \ \) r = \X, \, return borel r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto simp: pqp.pq1.return_qbs[OF qbs_Mx_to_X[OF assms]] distr_return intro!: pqp.qbs_s_finite_measure_eq) qed corollary return_qbs_morphism[qbs]: "return_qbs X \ X \\<^sub>Q monadM_qbs X" proof(rule qbs_morphismI) interpret rr : real_distribution "return borel 0" by(simp add: real_distribution_def real_distribution_axioms_def prob_space_return) fix \ assume h:"\ \ qbs_Mx X" then have 1:"return_qbs X \ \ = (\r. \X, \, return borel r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(rule return_qbs_comp) show "return_qbs X \ \ \ qbs_Mx (monadM_qbs X)" by(auto simp: 1 monadM_qbs_Mx h prob_kernel_def' intro!: exI[where x=\] exI[where x="return borel"] prob_kernel.s_finite_kernel_prob_kernel) qed subsubsection \Bind\ definition bind_qbs :: "['a qbs_measure, 'a \ 'b qbs_measure] \ 'b qbs_measure" where "bind_qbs s f \ (let (X, \, \) = rep_qbs_measure s; Y = qbs_space_of (f (\ undefined)); (\, k) = (SOME (\, k). f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx Y \ s_finite_kernel borel borel k) in \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" adhoc_overloading Monad_Syntax.bind bind_qbs lemma(in qbs_s_finite) assumes "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "f \ X \\<^sub>Q monadM_qbs Y" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" and "(f \ \) = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows bind_qbs_s_finite:"qbs_s_finite Y \ (\ \\<^sub>k k)" and bind_qbs: "s \ f = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - interpret k: s_finite_kernel borel borel k by fact interpret s_fin: qbs_s_finite Y \ "\ \\<^sub>k k" by(auto simp: qbs_s_finite_def in_Mx_def assms(3) mu_sets qbs_s_finite_axioms_def k.sets_bind_kernel[OF _ mu_sets] intro!:k.comp_s_finite_measure s_finite_measure_axioms) show "s \ f = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - { fix X' \' \' assume "(X',\',\') \ Rep_qbs_measure \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" then have h: "X' = X" "qbs_s_finite X' \' \'" "qbs_s_finite_eq (X,\,\) (X',\',\')" by(simp_all add: if_in_Rep_qbs_measure) then interpret s_fin_pq1: pair_qbs_s_finite X \ \ \' \' by(auto simp: pair_qbs_s_finite_def qbs_s_finite_axioms) have [simp]: "qbs_space_of (f (\' r)) = Y" for r using qbs_Mx_to_X[OF qbs_morphism_Mx[OF assms(2) s_fin_pq1.pq2.in_Mx],of r] by(auto simp: monadM_qbs_space) have "(let Y = qbs_space_of (f (\' undefined)) in case SOME (\, k). (\r. f (\' r)) = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx Y \ s_finite_kernel borel borel k of (\, k) \ \Y, \, \' \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - have "(case SOME (\, k). (\r. f (\' r)) = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx Y \ s_finite_kernel borel borel k of (\, k) \ \Y, \, \' \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof(rule someI2_ex) show "\a. case a of (\, k) \ (\r. f (\' r)) = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx Y \ s_finite_kernel borel borel k" using qbs_morphism_Mx[OF assms(2) s_fin_pq1.pq2.in_Mx] by(auto simp: comp_def monadM_qbs_Mx) next show "\x. (case x of (\, k) \ (\r. f (\' r)) = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx Y \ s_finite_kernel borel borel k) \ (case x of (\, k) \ \Y, \, \' \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof safe fix \' k' assume h':"(\r. f (\' r)) = (\r. \Y, \', k' r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\' \ qbs_Mx Y" "s_finite_kernel borel borel k'" interpret k': s_finite_kernel borel borel k' by fact have "qbs_s_finite Y \' (\' \\<^sub>k k')" by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def h'(2) k'.sets_bind_kernel[OF _ s_fin_pq1.pq2.mu_sets] s_fin_pq1.pq2.mu_sets intro!:k'.comp_s_finite_measure s_fin_pq1.pq2.s_finite_measure_axioms) then interpret s_fin_pq2: pair_qbs_s_finite Y \' "\' \\<^sub>k k'" \ "\ \\<^sub>k k" by(auto simp: pair_qbs_s_finite_def s_fin.qbs_s_finite_axioms) show "\Y, \', \' \\<^sub>k k'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \Y, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof(rule s_fin_pq2.qbs_s_finite_measure_eq) show "distr (\' \\<^sub>k k') (qbs_to_measure Y) \' = distr (\ \\<^sub>k k) (qbs_to_measure Y) \" (is "?lhs = ?rhs") proof - have "?lhs = \' \\<^sub>k (\r. distr (k' r) (qbs_to_measure Y) \')" by(simp add: k'.distr_bind_kernel[OF _ s_fin_pq1.pq2.mu_sets]) also have "... = \' \\<^sub>k (\r. qbs_l \Y, \', k' r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(rule bind_kernel_cong_All,rule qbs_s_finite.qbs_l[symmetric,OF qbs_s_finite_All[where k=k' and M=borel]]) (auto simp: k'.s_finite_kernel_axioms) also have "... = \' \\<^sub>k (\r. qbs_l (f (\' r)))" by(auto simp: fun_cong[OF h'(1)]) also have "... = distr \' (qbs_to_measure X) \' \\<^sub>k (\x. qbs_l (f x))" by(simp add: measure_kernel.bind_kernel_distr[OF measure_kernel.measure_kernel_comp[OF qbs_l_measure_kernel set_mp[OF l_preserves_morphisms assms(2)]]] sets_eq_imp_space_eq[OF s_fin_pq1.pq2.mu_sets]) also have "... = distr \ (qbs_to_measure X) \ \\<^sub>k (\x. qbs_l (f x))" by(simp add: qbs_s_finite_eq_dest(4)[OF h(3)]) also have "... = \ \\<^sub>k (\r. qbs_l (f (\ r)))" by(simp add: measure_kernel.bind_kernel_distr[OF measure_kernel.measure_kernel_comp[OF qbs_l_measure_kernel set_mp[OF l_preserves_morphisms assms(2)]]] sets_eq_imp_space_eq[OF mu_sets]) also have "... = \ \\<^sub>k (\r. qbs_l \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(simp add: fun_cong[OF assms(5),simplified comp_def]) also have "... = \ \\<^sub>k (\r. distr (k r) (qbs_to_measure Y) \)" by(rule bind_kernel_cong_All,rule qbs_s_finite.qbs_l[OF qbs_s_finite_All[where k=k and M=borel]]) (auto simp: k.s_finite_kernel_axioms) also have "... = ?rhs" by(simp add: k.distr_bind_kernel[OF _ mu_sets]) finally show ?thesis . qed qed qed qed thus ?thesis by simp qed } show ?thesis unfolding bind_qbs_def rep_qbs_measure_def qbs_measure.rep_def assms(1) by(rule someI2, rule in_Rep_qbs_measure, auto) fact qed show "qbs_s_finite Y \ (\ \\<^sub>k k)" by(rule s_fin.qbs_s_finite_axioms) qed lemma bind_qbs_morphism': assumes "f \ X \\<^sub>Q monadM_qbs Y" shows "(\x. x \ f) \ monadM_qbs X \\<^sub>Q monadM_qbs Y" proof(rule qbs_morphismI) fix \ assume "\ \ qbs_Mx (monadM_qbs X)" from rep_qbs_Mx_monadM[OF this] obtain \ k where h: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms this(2)]] obtain \' k' where h': "f \ \ = (\r. \Y, \', k' r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\' \ qbs_Mx Y" "s_finite_kernel borel borel k'" "\r. qbs_s_finite Y \' (k' r)" by metis have [simp]:"(\x. x \ f) \ \ = (\r. \Y, \', k r \\<^sub>k k'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by standard (simp add: h(1) qbs_s_finite.bind_qbs[OF h(4) _ assms h'(2,3,1)]) show "(\x. x \ f) \ \ \ qbs_Mx (monadM_qbs Y)" using h'(2) by(auto simp: s_finite_kernel.bind_kernel_s_finite_kernel[OF h(3) h'(3)] monadM_qbs_Mx intro!: exI[where x=\']) qed lemma bind_qbs_return': assumes "x \ qbs_space (monadM_qbs X)" shows "x \ return_qbs X = x" proof - obtain \ \ where h:"qbs_s_finite X \ \" "x = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using rep_qbs_space_monadM[OF assms] by blast then interpret qs: qbs_s_finite X \ \ by simp interpret prob_kernel borel borel "return borel" by(simp add: prob_kernel_def') show ?thesis by(simp add: qs.bind_qbs[OF h(2) return_qbs_morphism _ _ return_qbs_comp] s_finite_kernel_axioms bind_kernel_return''[OF qs.mu_sets] h(2)[symmetric]) qed lemma bind_qbs_return: assumes "f \ X \\<^sub>Q monadM_qbs Y" and "x \ qbs_space X" shows "return_qbs X x \ f = f x" proof - from rep_qbs_space_monadM[OF qbs_morphism_space[OF assms]] obtain \ \ where h: "f x = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" by auto then interpret qs:qbs_s_finite Y \ \ by simp interpret sk: s_finite_kernel borel borel "\r. \" by(auto intro!: s_finite_measure.s_finite_kernel_const simp: s_finite_kernel_cong_sets[OF refl qs.mu_sets[symmetric]] qs.s_finite_measure_axioms qs.mu_not_empty) interpret rd: real_distribution "return borel 0" by(simp add: real_distribution_def prob_space_return real_distribution_axioms_def) interpret qbs_prob X "\r. x" "return borel 0" by(rule rd.return_qbs_prob[OF assms(2)]) show ?thesis using bind_qbs[OF rd.return_qbs[OF assms(2)] assms(1) qs.in_Mx sk.s_finite_kernel_axioms] by(simp add: h(1) sk.bind_kernel_return) qed lemma bind_qbs_assoc: assumes "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q monadM_qbs Y" and "g \ Y \\<^sub>Q monadM_qbs Z" shows "s \ (\x. f x \ g) = (s \ f) \ g" (is "?lhs = ?rhs") proof - obtain \ \ where h:"s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" using rep_qbs_space_monadM[OF assms(1)] by blast then interpret qs: qbs_s_finite X \ \ by simp from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(2) qs.in_Mx]] obtain \ k where h': "f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" "\r. qbs_s_finite Y \ (k r)" by metis from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(3) h'(2)]] obtain \ k' where h'': "g \ \ = (\r. \Z, \, k' r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Z" "s_finite_kernel borel borel k'" "\r. qbs_s_finite Z \ (k' r)" by metis have 1:"(\x. f x \ g) \ \ = (\r. \Z, \, k r \\<^sub>k k'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by standard (simp add: qbs_s_finite.bind_qbs[OF h'(4) fun_cong[OF h'(1),simplified] assms(3) h''(2,3,1)]) have "?lhs = \Z, \, \ \\<^sub>k (\r. k r \\<^sub>k k')\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(rule qs.bind_qbs[OF h(1) qbs_morphism_compose[OF assms(2) bind_qbs_morphism'[OF assms(3)]] h''(2) s_finite_kernel.bind_kernel_s_finite_kernel[OF h'(3) h''(3)] 1]) also have "... = \Z, \, \ \\<^sub>k k \\<^sub>k k'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(simp add: s_finite_kernel.bind_kernel_assoc[OF h'(3) h''(3) qs.mu_sets]) also have "... = ?rhs" by(simp add: qbs_s_finite.bind_qbs[OF qs.bind_qbs_s_finite[OF h(1) assms(2) h'(2,3,1)] qs.bind_qbs[OF h(1) assms(2) h'(2,3,1)] assms(3) h''(2,3,1)]) finally show ?thesis . qed lemma bind_qbs_cong: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ f x = g x" and [qbs]:"f \ X \\<^sub>Q monadM_qbs Y" shows "s \ f = s \ g" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where h: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by auto interpret qbs_s_finite X \ \ by fact from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(3) in_Mx]] obtain \ k where h': "f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" by metis have g: "g \ X \\<^sub>Q monadM_qbs Y" "g \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" using qbs_Mx_to_X[OF in_Mx] assms(2) fun_cong[OF h'(1)] by(auto simp: assms(2)[symmetric] cong: qbs_morphism_cong) show ?thesis by(simp add: bind_qbs[OF h(1) assms(3) h'(2,3,1)] bind_qbs[OF h(1) g(1) h'(2,3) g(2)]) qed subsubsection \ The Functorial Action \ definition distr_qbs :: "['a quasi_borel, 'b quasi_borel,'a \ 'b,'a qbs_measure] \ 'b qbs_measure" where "distr_qbs _ Y f sx \ sx \ return_qbs Y \ f" lemma distr_qbs_morphism': assumes "f \ X \\<^sub>Q Y" shows "distr_qbs X Y f \ monadM_qbs X \\<^sub>Q monadM_qbs Y" unfolding distr_qbs_def by(rule bind_qbs_morphism'[OF qbs_morphism_comp[OF assms return_qbs_morphism]]) lemma(in qbs_s_finite) assumes "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and "f \ X \\<^sub>Q Y" shows distr_qbs_s_finite:"qbs_s_finite Y (f \ \) \" and distr_qbs: "distr_qbs X Y f s = \Y, f \ \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto intro!: bind_qbs[OF assms(1) qbs_morphism_comp[OF assms(2) return_qbs_morphism],of "f \ \" "return borel" ,simplified bind_kernel_return''[OF mu_sets]] bind_qbs_s_finite[OF assms(1) qbs_morphism_comp[OF assms(2) return_qbs_morphism],of "f \ \" "return borel" ,simplified bind_kernel_return''[OF mu_sets]] simp: distr_qbs_def return_qbs_comp[OF qbs_morphism_Mx[OF assms(2) in_Mx],simplified comp_assoc[symmetric]] qbs_morphism_Mx[OF assms(2) in_Mx] prob_kernel.s_finite_kernel_prob_kernel[of borel borel "return borel",simplified prob_kernel_def']) lemma(in qbs_prob) assumes "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and "f \ X \\<^sub>Q Y" shows distr_qbs_prob:"qbs_prob Y (f \ \) \" by(auto simp: distr_qbs_def prob_space_axioms intro!: qbs_s_finite.qbs_probI[OF distr_qbs_s_finite[OF assms]]) text \ We show that $M$ is a functor i.e. $M$ preserve identity and composition.\ lemma distr_qbs_id: assumes "s \ qbs_space (monadM_qbs X)" shows "distr_qbs X X id s = s" using bind_qbs_return'[OF assms] by(simp add: distr_qbs_def) lemma distr_qbs_comp: assumes "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q Y" and "g \ Y \\<^sub>Q Z" shows "((distr_qbs Y Z g) \ (distr_qbs X Y f)) s = distr_qbs X Z (g \ f) s" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where h: "qbs_s_finite X \ \" "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by metis have "qbs_s_finite Y (f \ \) \" "distr_qbs X Y f s = \Y, f \ \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(simp_all add: qbs_s_finite.distr_qbs_s_finite[OF h assms(2)] qbs_s_finite.distr_qbs[OF h assms(2)]) from qbs_s_finite.distr_qbs[OF this assms(3)] qbs_s_finite.distr_qbs[OF h qbs_morphism_comp[OF assms(2,3)]] show ?thesis by(simp add: comp_assoc) qed subsubsection \ Join \ definition join_qbs :: "'a qbs_measure qbs_measure \ 'a qbs_measure" where "join_qbs \ (\sst. sst \ id)" lemma join_qbs_morphism[qbs]: "join_qbs \ monadM_qbs (monadM_qbs X) \\<^sub>Q monadM_qbs X" by(simp add: join_qbs_def bind_qbs_morphism'[OF qbs_morphism_ident]) lemma assumes "qbs_s_finite (monadM_qbs X) \ \" "ssx = \monadM_qbs X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" and "\ =(\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows qbs_s_finite_join_qbs_s_finite: "qbs_s_finite X \ (\ \\<^sub>k k)" and qbs_s_finite_join_qbs: "join_qbs ssx = \X, \, \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite.bind_qbs[OF assms(1,2) qbs_morphism_ident assms(3,4)] qbs_s_finite.bind_qbs_s_finite[OF assms(1,2) qbs_morphism_ident assms(3,4)] by(auto simp: assms(5) join_qbs_def) subsubsection \ Strength \ definition strength_qbs :: "['a quasi_borel,'b quasi_borel,'a \ 'b qbs_measure] \ ('a \ 'b) qbs_measure" where "strength_qbs W X = (\(w,sx). let (_,\,\) = rep_qbs_measure sx in \W \\<^sub>Q X, \r. (w,\ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" lemma(in qbs_s_finite) assumes "w \ qbs_space W" and "sx = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" shows strength_qbs_s_finite: "qbs_s_finite (W \\<^sub>Q X) (\r. (w,\ r)) \" and strength_qbs: "strength_qbs W X (w,sx) = \W \\<^sub>Q X, \r. (w,\ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - interpret qs: qbs_s_finite "W \\<^sub>Q X" "\r. (w,\ r)" \ by(auto simp: qbs_s_finite_def s_finite_measure_axioms qbs_s_finite_axioms_def mu_sets in_Mx_def assms(1) intro!: pair_qbs_MxI) show "qbs_s_finite (W \\<^sub>Q X) (\r. (w,\ r)) \" by (rule qs.qbs_s_finite_axioms) show "strength_qbs W X (w,sx) = \W \\<^sub>Q X, \r. (w,\ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - { fix X' \' \' assume "(X',\',\') \ Rep_qbs_measure \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" then have h: "X' = X" "qbs_s_finite X' \' \'" "qbs_s_finite_eq (X,\,\) (X',\',\')" by(simp_all add: if_in_Rep_qbs_measure) then interpret qs': qbs_s_finite "W \\<^sub>Q X" "\r. (w,\' r)" \' by(auto simp: qbs_s_finite_def in_Mx_def assms(1) intro!: pair_qbs_MxI) interpret pq: pair_qbs_s_finite "W \\<^sub>Q X" "\r. (w,\ r)" \ "\r. (w,\' r)" \' by(auto simp: qs.qbs_s_finite_axioms qs'.qbs_s_finite_axioms pair_qbs_s_finite_def) have "\W \\<^sub>Q X, \r. (w, \' r), \'\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \W \\<^sub>Q X, \r. (w, \ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof(rule pq.qbs_s_finite_measure_eq'[symmetric]) fix f :: "_ \ ennreal" assume "f \ W \\<^sub>Q X \\<^sub>Q qbs_borel" then have f: "curry f w \ X \\<^sub>Q qbs_borel" by (metis assms(1) qbs_morphism_curry qbs_morphism_space) show "(\\<^sup>+ x. f (w, \ x) \\) = (\\<^sup>+ x. f (w, \' x) \\')" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. curry f w (\ x) \\)" by simp also have "... = (\\<^sup>+ x. curry f w (\' x) \\')" using h(3) f by(auto simp: qbs_s_finite_eq_equiv qbs_s_finite_eq'_def h(1)) also have "... = ?rhs" by simp finally show ?thesis . qed qed } show ?thesis by(simp add: strength_qbs_def rep_qbs_measure_def qbs_measure.rep_def assms(2)) (rule someI2, rule in_Rep_qbs_measure, auto, fact) qed qed lemma(in qbs_prob) assumes "w \ qbs_space W" and "sx = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" shows strength_qbs_prob: "qbs_prob (W \\<^sub>Q X) (\r. (w,\ r)) \" by(auto intro!: qbs_s_finite.qbs_probI[OF strength_qbs_s_finite[OF assms]] prob_space_axioms) lemma strength_qbs_natural: assumes "f \ X \\<^sub>Q X'" "g \ Y \\<^sub>Q Y'" "x \ qbs_space X" and "sy \ qbs_space (monadM_qbs Y)" shows "(distr_qbs (X \\<^sub>Q Y) (X' \\<^sub>Q Y') (map_prod f g) \ strength_qbs X Y) (x,sy) = (strength_qbs X' Y' \ map_prod f (distr_qbs Y Y' g)) (x,sy)" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms(4)] obtain \ \ where h:"sy = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" by metis have "?lhs = (distr_qbs (X \\<^sub>Q Y) (X' \\<^sub>Q Y') (map_prod f g)) (\X \\<^sub>Q Y, \r. (x,\ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(simp add: qbs_s_finite.strength_qbs[OF h(2) assms(3) h(1)]) also have "... = \X' \\<^sub>Q Y', map_prod f g \ (\r. (x, \ r)), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using assms by(simp add: qbs_s_finite.distr_qbs[OF qbs_s_finite.strength_qbs_s_finite[OF h(2) assms(3) h(1)] refl ]) also have "... = \X' \\<^sub>Q Y',\r. (f x, (g \ \) r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by (simp add: comp_def) also have "... = ?rhs" by(simp add: qbs_s_finite.strength_qbs[OF qbs_s_finite.distr_qbs_s_finite[OF h(2,1) assms(2)] qbs_morphism_space[OF assms(1,3)] qbs_s_finite.distr_qbs[OF h(2,1) assms(2)]]) finally show ?thesis . qed context begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) declare rr.from_real_to_real[simplified space_pair_measure,simplified,simp] lemma rr_from_real_to_real_id[simp]: "rr.from_real \ rr.to_real = id" by(auto simp: comp_def) lemma assumes "\ \ qbs_Mx X" "\ \ qbs_Mx (monadM_qbs Y)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" and "\ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows strength_qbs_ab_r_s_finite: "qbs_s_finite (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (return borel r \\<^sub>M k r) borel rr.to_real)" and strength_qbs_ab_r: "strength_qbs X Y (\ r, \ r) = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, distr (return borel r \\<^sub>M k r) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is ?goal2) proof - interpret k: s_finite_kernel borel borel k by fact note 1[measurable_cong] = sets_return[of borel r] k.kernel_sets[of r,simplified] show "qbs_s_finite (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (return borel r \\<^sub>M k r) borel rr.to_real)" using assms(1,3) by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def qbs_Mx_is_morphisms r_preserves_product[symmetric] standard_borel_ne.standard_borel intro!: s_finite_measure.s_finite_measure_distr[OF pair_measure_s_finite_measure[OF prob_space.s_finite_measure_prob[OF prob_space_return[of r borel]] k.image_s_finite_measure[of r]]] qbs_morphism_comp[where Y="qbs_borel \\<^sub>Q qbs_borel"] qbs_morphism_space[OF qbs_morphism_space[OF qbs_morphism_map_prod]] standard_borel.qbs_morphism_measurable_intro[of "borel :: real measure"]) then interpret qs: qbs_s_finite "X \\<^sub>Q Y" "map_prod \ \ \ rr.from_real" "distr (return borel r \\<^sub>M k r) borel rr.to_real" . interpret qs2: qbs_s_finite Y \ "k r" by(auto simp: qbs_s_finite_def k.image_s_finite_measure in_Mx_def assms qbs_s_finite_axioms_def k.kernel_sets) interpret pq: pair_qbs_s_finite "X \\<^sub>Q Y" "\l. (\ r, \ l)" "k r" "map_prod \ \ \ rr.from_real" "distr (return borel r \\<^sub>M k r) borel rr.to_real" by (auto simp: pair_qbs_s_finite_def qs.qbs_s_finite_axioms qs2.strength_qbs_s_finite[OF qbs_Mx_to_X[OF assms(1),of r] fun_cong[OF assms(5)]]) have [measurable]: "map_prod \ \ \ borel \\<^sub>M borel \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" proof - have "map_prod \ \ \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q X \\<^sub>Q Y" using assms by(auto intro!: qbs_morphism_map_prod simp: qbs_Mx_is_morphisms) also have "... \ qbs_to_measure (qbs_borel \\<^sub>Q qbs_borel) \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" by(rule l_preserves_morphisms) also have "... = borel \\<^sub>M borel \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" using rr.lr_sets_ident l_preserves_morphisms by(auto simp add: r_preserves_product[symmetric]) finally show ?thesis . qed show ?goal2 unfolding qs2.strength_qbs[OF qbs_Mx_to_X[OF assms(1),of r] fun_cong[OF assms(5)]] proof(rule pq.qbs_s_finite_measure_eq) show "distr (k r) (qbs_to_measure (X \\<^sub>Q Y)) (\l. (\ r, \ l)) = distr (distr (return borel r \\<^sub>M k r) borel rr.to_real) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \ \ rr.from_real)" (is "?lhs = ?rhs") proof - have "?lhs = distr (k r) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \ \ Pair r)" by(simp add: comp_def) also have "... = distr (distr (k r) (borel \\<^sub>M borel) (Pair r)) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \)" by(auto intro!: distr_distr[symmetric]) also have "... = distr (return borel r \\<^sub>M k r) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \)" proof - have "return borel r \\<^sub>M k r = distr (k r) (borel \\<^sub>M borel) (\l. (r,l))" by(auto intro!: measure_eqI simp: sets_pair_measure_cong[OF refl 1(2)] qs2.emeasure_pair_measure_alt' emeasure_distr nn_integral_return[OF _ qs2.measurable_emeasure_Pair']) thus ?thesis by simp qed also have "... = ?rhs" by(simp add: distr_distr comp_def) finally show ?thesis . qed qed qed lemma strength_qbs_morphism[qbs]: "strength_qbs X Y \ X \\<^sub>Q monadM_qbs Y \\<^sub>Q monadM_qbs (X \\<^sub>Q Y)" proof(rule pair_qbs_morphismI) fix \ \ assume h:"\ \ qbs_Mx X" "\ \ qbs_Mx (monadM_qbs Y)" from rep_qbs_Mx_monadM[OF this(2)] obtain \ k where hb: "\ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" by metis have "s_finite_kernel borel borel (\r. distr (return borel r \\<^sub>M k r) borel rr.to_real)" by(auto intro!: s_finite_kernel.distr_s_finite_kernel[where Y="borel \\<^sub>M borel"] s_finite_kernel_pair_measure[OF prob_kernel.s_finite_kernel_prob_kernel] simp:hb prob_kernel_def') thus "(\r. strength_qbs X Y (\ r, \ r)) \ qbs_Mx (monadM_qbs (X \\<^sub>Q Y))" using strength_qbs_ab_r[OF h hb(2,3,1)] strength_qbs_ab_r_s_finite[OF h hb(2,3,1)] by(auto simp: monadM_qbs_Mx qbs_s_finite_def in_Mx_def intro!: exI[where x="map_prod \ \ \ rr.from_real"] exI[where x="\r. distr (return borel r \\<^sub>M k r) borel rr.to_real"]) qed lemma bind_qbs_morphism[qbs]: "(\) \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q monadM_qbs Y) \\<^sub>Q monadM_qbs Y" proof - { fix f s assume h:"f \ X \\<^sub>Q monadM_qbs Y" "s \ qbs_space (monadM_qbs X)" from rep_qbs_space_monadM[OF this(2)] obtain \ \ where h': "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qbs_s_finite X \ \ by simp from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF h(1) in_Mx]] obtain \ k where hb:"f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" by metis have "join_qbs (distr_qbs ((X \\<^sub>Q monadM_qbs Y) \\<^sub>Q X) (monadM_qbs Y) (\fx. fst fx (snd fx)) (strength_qbs (X \\<^sub>Q monadM_qbs Y) X (f, s))) = s \ f" using qbs_s_finite_join_qbs[OF qbs_s_finite.distr_qbs_s_finite[OF strength_qbs_s_finite[of f "X \\<^sub>Q monadM_qbs Y",OF h(1) h'(1)] strength_qbs[of f "X \\<^sub>Q monadM_qbs Y",OF h(1) h'(1)] qbs_morphism_eval] qbs_s_finite.distr_qbs[OF strength_qbs_s_finite[of f "X \\<^sub>Q monadM_qbs Y",OF h(1) h'(1)] strength_qbs[of f "X \\<^sub>Q monadM_qbs Y",OF h(1) h'(1)] qbs_morphism_eval] hb(2,3)] hb(1) by(simp add: bind_qbs[OF h'(1) h(1) hb(2,3,1)] comp_def) } thus ?thesis by(auto intro!: arg_swap_morphism[OF curry_preserves_morphisms[OF qbs_morphism_cong'[of _ "join_qbs \ (distr_qbs (exp_qbs X (monadM_qbs Y) \\<^sub>Q X) (monadM_qbs Y) (\fx. (fst fx) (snd fx))) \ (strength_qbs (exp_qbs X (monadM_qbs Y)) X)"]]] qbs_morphism_comp distr_qbs_morphism' strength_qbs_morphism join_qbs_morphism qbs_morphism_eval simp: pair_qbs_space) qed lemma strength_qbs_law1: assumes "x \ qbs_space (unit_quasi_borel \\<^sub>Q monadM_qbs X)" shows "snd x = (distr_qbs (unit_quasi_borel \\<^sub>Q X) X snd \ strength_qbs unit_quasi_borel X) x" proof - obtain \ \ where h: "qbs_s_finite X \ \" "(snd x) = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using rep_qbs_space_monadM[of "snd x" X] assms by (auto simp: pair_qbs_space) metis have [simp]: "((),snd x) = x" using SigmaE assms by (auto simp: pair_qbs_space) show ?thesis using qbs_s_finite.distr_qbs[OF qbs_s_finite.strength_qbs_s_finite[OF h(1) _ h(2),of "fst x" unit_quasi_borel] qbs_s_finite.strength_qbs[OF h(1) _ h(2)] snd_qbs_morphism] by(auto simp: comp_def,simp add: h(2)) qed lemma strength_qbs_law2: assumes "x \ qbs_space ((X \\<^sub>Q Y) \\<^sub>Q monadM_qbs Z)" shows "(strength_qbs X (Y \\<^sub>Q Z) \ (map_prod id (strength_qbs Y Z)) \ (\((x,y),z). (x,(y,z)))) x = (distr_qbs ((X \\<^sub>Q Y) \\<^sub>Q Z) (X \\<^sub>Q (Y \\<^sub>Q Z)) (\((x,y),z). (x,(y,z))) \ strength_qbs (X \\<^sub>Q Y) Z) x" (is "?lhs = ?rhs") proof - obtain \ \ where h: "qbs_s_finite Z \ \" "snd x = \Z, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using rep_qbs_space_monadM[of "snd x" Z] assms by (auto simp: pair_qbs_space) metis then have "?lhs = \X \\<^sub>Q Y \\<^sub>Q Z, \r. (fst (fst x), snd (fst x), \ r), \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using assms qbs_s_finite.strength_qbs_s_finite[OF h(1) _ h(2),of "snd (fst x)" Y] by(auto intro!: qbs_s_finite.strength_qbs simp: pair_qbs_space) also have "... = ?rhs" using qbs_s_finite.distr_qbs[OF qbs_s_finite.strength_qbs_s_finite[OF h(1) _ h(2),of "fst x" "X \\<^sub>Q Y"] qbs_s_finite.strength_qbs[OF h(1) _ h(2),of "fst x" "X \\<^sub>Q Y"] qbs_morphism_pair_assoc1] assms by(auto simp: comp_def pair_qbs_space) finally show ?thesis . qed lemma strength_qbs_law3: assumes "x \ qbs_space (X \\<^sub>Q Y)" shows "return_qbs (X \\<^sub>Q Y) x = (strength_qbs X Y \ (map_prod id (return_qbs Y))) x" proof - interpret qp: qbs_prob Y "\r. snd x" "return borel 0" using assms by(auto simp: prob_space_return pair_qbs_space qbs_prob_def in_Mx_def real_distribution_def real_distribution_axioms_def) show ?thesis using qp.strength_qbs[OF _ qp.return_qbs[of "snd x" Y],of "fst x" X] qp.return_qbs[OF assms] assms by(auto simp: pair_qbs_space) qed lemma strength_qbs_law4: assumes "x \ qbs_space (X \\<^sub>Q monadM_qbs (monadM_qbs Y))" shows "(strength_qbs X Y \ map_prod id join_qbs) x = (join_qbs \ distr_qbs (X \\<^sub>Q monadM_qbs Y) (monadM_qbs (X \\<^sub>Q Y)) (strength_qbs X Y) \ strength_qbs X (monadM_qbs Y)) x" (is "?lhs = ?rhs") proof - from assms rep_qbs_space_monadM[of "snd x" "monadM_qbs Y"] obtain \ \ where h:"qbs_s_finite (monadM_qbs Y) \ \" "snd x = \monadM_qbs Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by (auto simp: pair_qbs_space) metis with rep_qbs_Mx_monadM[of \ Y] obtain \ k where h': "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" "\ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" and h'': "\r. qbs_s_finite Y \ (k r)" by(auto simp: qbs_s_finite_def in_Mx_def) metis have "?lhs = \X \\<^sub>Q Y, \r. (fst x, \ r), \ \\<^sub>k k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite.strength_qbs[OF qbs_s_finite_join_qbs_s_finite[OF h h'] _ qbs_s_finite_join_qbs[OF h h'],of "fst x" X] assms by(auto simp: pair_qbs_space) also have "... = ?rhs" using qbs_s_finite_join_qbs[OF qbs_s_finite.distr_qbs_s_finite[OF qbs_s_finite.strength_qbs_s_finite[OF h(1) _ h(2),of "fst x" X] qbs_s_finite.strength_qbs[OF h(1) _ h(2),of "fst x"] strength_qbs_morphism] qbs_s_finite.distr_qbs[OF qbs_s_finite.strength_qbs_s_finite[OF h(1) _ h(2),of "fst x" X] qbs_s_finite.strength_qbs[OF h(1) _ h(2),of "fst x"] strength_qbs_morphism] pair_qbs_MxI h'(2),of "\r. (fst x, \ r)",simplified comp_def qbs_s_finite.strength_qbs[OF h'' _ fun_cong[OF h'(3)],of "fst x" X]] assms h'(1) by(auto simp: pair_qbs_space qbs_s_finite_def in_Mx_def) finally show ?thesis . qed lemma distr_qbs_morphism[qbs]: "distr_qbs X Y \ (X \\<^sub>Q Y) \\<^sub>Q (monadM_qbs X \\<^sub>Q monadM_qbs Y)" proof - have [simp]: "distr_qbs X Y = (\f sx. sx \ return_qbs Y \ f)" by standard+ (auto simp: distr_qbs_def) show ?thesis by simp qed lemma assumes "\ \ qbs_Mx X" "\ \ qbs_Mx Y" shows return_qbs_pair_Mx: "return_qbs (X \\<^sub>Q Y) (\ r, \ k) = \X \\<^sub>Q Y,map_prod \ \ \ rr.from_real, distr (return borel r \\<^sub>M return borel k) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and return_qbs_pair_Mx_prob: "qbs_prob (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (return borel r \\<^sub>M return borel k) borel rr.to_real)" proof - note [measurable_cong] = sets_return[of borel] interpret qp: qbs_prob "X \\<^sub>Q Y" "map_prod \ \ \ rr.from_real" "distr (return borel r \\<^sub>M return borel k) borel rr.to_real" using qbs_closed1_dest[OF assms(1)] qbs_closed1_dest[OF assms(2)] by(auto intro!: prob_space.prob_space_distr prob_space_pair simp: comp_def prob_space_return pair_qbs_Mx qbs_prob_def in_Mx_def real_distribution_def real_distribution_axioms_def) show "qbs_prob (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (return borel r \\<^sub>M return borel k) borel rr.to_real)" by standard show "return_qbs (X \\<^sub>Q Y) (\ r, \ k) = \X \\<^sub>Q Y,map_prod \ \ \ rr.from_real, distr (return borel r \\<^sub>M return borel k) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is "?lhs = ?rhs") proof - have "?lhs = (strength_qbs X Y \ map_prod id (return_qbs Y)) (\ r, \ k)" by(rule strength_qbs_law3[of "(\ r, \ k)" X Y], insert assms) (auto simp: qbs_Mx_to_X pair_qbs_space) also have "... = strength_qbs X Y (\ r, \Y, \, return borel k\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" using fun_cong[OF return_qbs_comp[OF assms(2)]] by simp also have "... = ?rhs" by(rule strength_qbs_ab_r[OF assms(1) _ assms(2)]) (auto intro!: qbs_closed2_dest qbs_s_finite.in_space_monadM s_finite_measure.s_finite_kernel_const[of "return borel k",simplified s_finite_kernel_cong_sets[OF _ sets_return]] prob_space.s_finite_measure_prob simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def assms(2) prob_space_return) finally show ?thesis . qed qed lemma bind_bind_return_distr: assumes "s_finite_measure \" and "s_finite_measure \" and [measurable_cong]: "sets \ = sets borel" "sets \ = sets borel" shows "\ \\<^sub>k (\r. \ \\<^sub>k (\l. distr (return borel r \\<^sub>M return borel l) borel rr.to_real)) = distr (\ \\<^sub>M \) borel rr.to_real" (is "?lhs = ?rhs") proof - interpret rd1: s_finite_measure \ by fact interpret rd2: s_finite_measure \ by fact have ne: "space \ \ {}" "space \ \ {}" by(auto simp: sets_eq_imp_space_eq assms(3,4)) have "?lhs = \ \\<^sub>k (\r. \ \\<^sub>k (\l. distr (return (borel \\<^sub>M borel) (r,l)) borel rr.to_real))" by(simp add: pair_measure_return) also have "... = \ \\<^sub>k (\r. \ \\<^sub>k (\l. distr (return (\ \\<^sub>M \) (r, l)) borel rr.to_real))" proof - have "return (borel \\<^sub>M borel) = return (\ \\<^sub>M \)" by(auto intro!: return_sets_cong sets_pair_measure_cong simp: assms(3,4)) thus ?thesis by simp qed also have "... = \ \\<^sub>k (\r. distr (\ \\<^sub>k (\l. (return (\ \\<^sub>M \) (r, l)))) borel rr.to_real)" by(auto intro!: bind_kernel_cong_All measure_kernel.distr_bind_kernel[of \ "\ \\<^sub>M \",symmetric] simp: ne measure_kernel_def space_pair_measure) also have "... = distr (\ \\<^sub>k (\r. \ \\<^sub>k (\l. return (\ \\<^sub>M \) (r, l)))) borel rr.to_real" by(auto intro!: measure_kernel.distr_bind_kernel[of \ "\ \\<^sub>M \",symmetric] s_finite_kernel.axioms(1) s_finite_kernel.bind_kernel_s_finite_kernel'[where Y=\] s_finite_measure.s_finite_kernel_const[OF assms(2)] prob_kernel.s_finite_kernel_prob_kernel[of "\ \\<^sub>M \"] simp: ne prob_kernel_def') also have "... = ?rhs" by(simp add: pair_measure_eq_bind_s_finite[OF assms(1,2),symmetric]) finally show ?thesis . qed end context begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) lemma from_real_rr_qbs_morphism[qbs]: "rr.from_real \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" by (metis borel_prod qbs_Mx_R qbs_Mx_is_morphisms qbs_borel_prod rr.from_real_measurable) end context pair_qbs_s_finites begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) sublocale qbs_s_finite "X \\<^sub>Q Y" "map_prod \ \ \ rr.from_real" "distr (\ \\<^sub>M \) borel rr.to_real" by(auto simp: qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def qbs_Mx_is_morphisms pq1.s_finite_measure_axioms pq2.s_finite_measure_axioms intro!: s_finite_measure.s_finite_measure_distr[OF pair_measure_s_finite_measure]) lemma qbs_bind_bind_return_qp: "\Y,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\y. \X,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\x. return_qbs (X \\<^sub>Q Y) (x,y))) = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, distr (\ \\<^sub>M \) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is "?lhs = ?rhs") proof - have "?lhs = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, \ \\<^sub>k (\l. \ \\<^sub>k (\r. distr (return borel r \\<^sub>M return borel l) borel rr.to_real))\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto intro!: pq2.bind_qbs[OF refl] s_finite_kernel.bind_kernel_s_finite_kernel'[where Y=\] s_finite_measure.s_finite_kernel_const s_finite_kernel.distr_s_finite_kernel[where Y="borel \\<^sub>M borel"] prob_kernel.s_finite_kernel_prob_kernel[of "borel \\<^sub>M \"] simp: sets_eq_imp_space_eq[OF pq1.mu_sets] pq1.s_finite_measure_axioms split_beta' pair_measure_return[of _ "snd _"] prob_kernel_def') (auto intro!: pq1.bind_qbs prob_kernel.s_finite_kernel_prob_kernel simp: comp_def return_qbs_pair_Mx qbs_Mx_is_morphisms prob_kernel_def') also have "... = ?rhs" proof - have "\ \\<^sub>k (\l. \ \\<^sub>k (\r. distr (return borel r \\<^sub>M return borel l) borel rr.to_real)) = distr (\ \\<^sub>M \) borel rr.to_real" by(auto simp: bind_bind_return_distr[OF pq1.s_finite_measure_axioms pq2.s_finite_measure_axioms pq1.mu_sets pq2.mu_sets,symmetric] pq1.s_finite_measure_axioms pq2.s_finite_measure_axioms prob_kernel_def' intro!: bind_kernel_rotate[where Z=borel] prob_kernel.s_finite_kernel_prob_kernel) thus ?thesis by simp qed finally show ?thesis . qed lemma qbs_bind_bind_return_pq: "\X,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\x. \Y,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\y. return_qbs (X \\<^sub>Q Y) (x,y))) = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, distr (\ \\<^sub>M \) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is "?lhs = ?rhs") proof - have "?lhs = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, \ \\<^sub>k (\r. \ \\<^sub>k (\l. distr (return borel r \\<^sub>M return borel l) borel rr.to_real))\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto intro!: pq1.bind_qbs[OF refl]s_finite_kernel.bind_kernel_s_finite_kernel'[where Y=\] s_finite_measure.s_finite_kernel_const s_finite_kernel.distr_s_finite_kernel[where Y="borel \\<^sub>M borel"] prob_kernel.s_finite_kernel_prob_kernel[of "borel \\<^sub>M \"] simp: sets_eq_imp_space_eq[OF pq2.mu_sets] pq2.s_finite_measure_axioms split_beta' pair_measure_return[of _ "fst _"] prob_kernel_def') (auto intro!: pq2.bind_qbs prob_kernel.s_finite_kernel_prob_kernel simp: comp_def return_qbs_pair_Mx qbs_Mx_is_morphisms prob_kernel_def') also have "... = ?rhs" by(simp add: bind_bind_return_distr[OF pq1.s_finite_measure_axioms pq2.s_finite_measure_axioms pq1.mu_sets pq2.mu_sets]) finally show ?thesis . qed end lemma bind_qbs_return_rotate: assumes "p \ qbs_space (monadM_qbs X)" and "q \ qbs_space (monadM_qbs Y)" shows "q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y))) = p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))" proof - from rep_qbs_space_monadM[OF assms(1)] rep_qbs_space_monadM[OF assms(2)] obtain \ \ \ \ where h: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" "qbs_s_finite Y \ \" by metis then interpret pair_qbs_s_finites X \ \ Y \ \ by(simp add: pair_qbs_s_finites_def) show ?thesis by(simp add: h(1,2) qbs_bind_bind_return_pq qbs_bind_bind_return_qp) qed lemma qbs_bind_bind_return1: assumes [qbs]: "f \ X \\<^sub>Q Y \\<^sub>Q monadM_qbs Z" "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" shows "q \ (\y. p \ (\x. f (x,y))) = (q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y)))) \ f" (is "?lhs = ?rhs") proof - have "?lhs = q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y) \ f))" by(auto intro!: bind_qbs_cong[OF assms(3),where Y=Z] bind_qbs_cong[OF assms(2),where Y=Z] simp: bind_qbs_return[OF assms(1),simplified pair_qbs_space]) also have "... = q \ (\y. (p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y))) \ f)" by(auto intro!: bind_qbs_cong[OF assms(3),where Y=Z] bind_qbs_assoc[OF assms(2) _ assms(1)] simp: ) also have "... = ?rhs" by(simp add: bind_qbs_assoc[OF assms(3) _ assms(1)]) finally show ?thesis . qed lemma qbs_bind_bind_return2: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q monadM_qbs Z" "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" shows "p \ (\x. q \ (\y. f (x,y))) = (p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))) \ f" (is "?lhs = ?rhs") proof - have "?lhs = p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y) \ f))" by(auto intro!: bind_qbs_cong[OF assms(2),where Y=Z] bind_qbs_cong[OF assms(3),where Y=Z] simp: bind_qbs_return[OF assms(1),simplified pair_qbs_space]) also have "... = p \ (\x. (q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y))) \ f)" by(auto intro!: bind_qbs_cong[OF assms(2),where Y=Z] bind_qbs_assoc[OF assms(3) _ assms(1)]) also have "... = ?rhs" by(simp add: bind_qbs_assoc[OF assms(2) _ assms(1)]) finally show ?thesis . qed corollary bind_qbs_rotate: assumes "f \ X \\<^sub>Q Y \\<^sub>Q monadM_qbs Z" "p \ qbs_space (monadM_qbs X)" and "q \ qbs_space (monadM_qbs Y)" shows "q \ (\y. p \ (\x. f (x,y))) = p \ (\x. q \ (\y. f (x,y)))" by(simp add: qbs_bind_bind_return1[OF assms] qbs_bind_bind_return2[OF assms] bind_qbs_return_rotate assms) context pair_qbs_s_finites begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) lemma assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q Z" shows qbs_bind_bind_return:"\X,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\x. \Y,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\y. return_qbs Z (f (x,y)))) = \Z, f \ (map_prod \ \ \ rr.from_real), distr (\ \\<^sub>M \) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is "?lhs = ?rhs") and qbs_bind_bind_return_s_finite: "qbs_s_finite Z (f \ (map_prod \ \ \ rr.from_real)) (distr (\ \\<^sub>M \) borel rr.to_real)" proof - show "qbs_s_finite Z (f \ (map_prod \ \ \ rr.from_real)) (distr (\ \\<^sub>M \) borel rr.to_real)" using qbs_s_finite_axioms by(auto simp: qbs_s_finite_def in_Mx_def qbs_Mx_is_morphisms) have "?lhs = \X,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\x. \Y,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ (\y. return_qbs (X \\<^sub>Q Y) (x,y))) \ return_qbs Z \ f" by(auto simp: comp_def intro!: qbs_bind_bind_return2[of "return_qbs Z \ f" _ _ Z,simplified comp_def]) also have "... = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, distr (\ \\<^sub>M \) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ return_qbs Z \ f" by(simp add: qbs_bind_bind_return_pq) also have "... = ?rhs" by(rule distr_qbs[OF refl assms,simplified distr_qbs_def]) finally show "?lhs = ?rhs" . qed end subsubsection \The Probability Monad\ definition "monadP_qbs X \ sub_qbs (monadM_qbs X) {s. prob_space (qbs_l s)}" lemma shows qbs_space_monadPM: "s \ qbs_space (monadP_qbs X) \ s \ qbs_space (monadM_qbs X)" and qbs_Mx_monadPM: "f \ qbs_Mx (monadP_qbs X) \ f \ qbs_Mx (monadM_qbs X)" by(simp_all add: monadP_qbs_def sub_qbs_space sub_qbs_Mx) lemma monadP_qbs_space: "qbs_space (monadP_qbs X) = {s. qbs_space_of s = X \ prob_space (qbs_l s)}" by(auto simp: monadP_qbs_def sub_qbs_space monadM_qbs_space) lemma rep_qbs_space_monadP: assumes "s \ qbs_space (monadP_qbs X)" obtains \ \ where "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_prob X \ \" proof - obtain \ \ where h:"s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" using assms rep_qbs_space_monadM[of s X] by(auto simp: monadP_qbs_def sub_qbs_space) interpret qbs_s_finite X \ \ by fact have "prob_space \" by(rule prob_space_distrD[of \ _ "qbs_to_measure X"]) (insert assms, auto simp: qbs_l[symmetric] h(1)[symmetric] monadP_qbs_space) thus ?thesis by (simp add: h(1) in_Mx_axioms mu_sets qbs_prob.intro real_distribution_axioms_def real_distribution_def that) qed lemma qbs_l_prob_space: "s \ qbs_space (monadP_qbs X) \ prob_space (qbs_l s)" by(auto simp: monadP_qbs_space) lemma monadP_qbs_empty_iff: "(qbs_space X = {}) = (qbs_space (monadP_qbs X) = {})" proof show "qbs_space X = {} \ qbs_space (monadP_qbs X) = {}" using qbs_s_space_of_not_empty by(auto simp add: monadP_qbs_space) next assume "qbs_space (monadP_qbs X) = {}" then have h:"\s. qbs_space_of s = X \ \ prob_space (qbs_l s)" by(simp add: monadP_qbs_space) show "qbs_space X = {}" proof(rule ccontr) assume "qbs_space X \ {}" then obtain a where a:"a \ qbs_Mx X" by (auto simp: qbs_empty_equiv) then interpret qbs_prob X a "return borel 0" by(auto simp: qbs_prob_def in_Mx_def real_distribution_axioms_def real_distribution_def prob_space_return) have "qbs_space_of \X, a, return borel 0\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = X" "prob_space (qbs_l \X, a, return borel 0\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(auto simp: qbs_l intro!: prob_space_distr) with h show False by simp qed qed lemma in_space_monadP_qbs_pred: "qbs_pred (monadM_qbs X) (\s. s \ monadP_qbs X)" by(rule qbs_morphism_cong'[where f="\s. prob_space (qbs_l s)"],auto simp: qbs_l_prob_pred) (auto simp: monadP_qbs_def sub_qbs_space) lemma(in qbs_prob) in_space_monadP[qbs]: "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \ qbs_space (monadP_qbs X)" by(auto simp: monadP_qbs_space qbs_l prob_space_distr) lemma qbs_morphism_monadPD: "f \ X \\<^sub>Q monadP_qbs Y \ f \ X \\<^sub>Q monadM_qbs Y" unfolding monadP_qbs_def by(rule qbs_morphism_subD) lemma qbs_morphism_monadPD': "f \ monadM_qbs X \\<^sub>Q Y \ f \ monadP_qbs X \\<^sub>Q Y" unfolding monadP_qbs_def by(rule qbs_morphism_subI2) lemma qbs_morphism_monadPI: assumes "\x. x \ qbs_space X \ prob_space (qbs_l (f x))" "f \ X \\<^sub>Q monadM_qbs Y" shows "f \ X \\<^sub>Q monadP_qbs Y" using assms by(auto simp: monadP_qbs_def intro!:qbs_morphism_subI1) lemma qbs_morphism_monadPI': assumes "\x. x \ qbs_space X \ f x \ qbs_space (monadP_qbs Y)" "f \ X \\<^sub>Q monadM_qbs Y" shows "f \ X \\<^sub>Q monadP_qbs Y" using assms by(auto intro!: qbs_morphism_monadPI simp: monadP_qbs_space) lemma qbs_morphism_monadPI'': assumes "f \ monadM_qbs X \\<^sub>Q monadM_qbs Y" "\s. s \ qbs_space (monadP_qbs X) \ f s \ qbs_space (monadP_qbs Y)" shows "f \ monadP_qbs X \\<^sub>Q monadP_qbs Y" proof - have 1:"\X. monadP_qbs X = sub_qbs (monadM_qbs X) {s. qbs_space_of s = X \ prob_space (qbs_l s)}" (is "\X. ?l X = ?r X") proof - fix X have "?l X = sub_qbs (sub_qbs (monadM_qbs X) (qbs_space (monadM_qbs X))) {s. prob_space (qbs_l s)}" by(simp add: sub_qbs_ident monadP_qbs_def) also have "... = ?r X" by(auto simp: sub_qbs_sub_qbs monadM_qbs_space Collect_conj_eq) finally show "?l X = ?r X" . qed show ?thesis unfolding 1 using assms(2) by(auto intro!: qbs_morphism_subsubI[OF assms(1),of " {s. qbs_space_of s = X \ prob_space (qbs_l s)}" " {s. qbs_space_of s = Y \ prob_space (qbs_l s)}"] simp: 1 sub_qbs_space monadM_qbs_space) qed lemma monadP_qbs_Mx: "qbs_Mx (monadP_qbs X) = {\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n |\ k. \ \ qbs_Mx X \ k \ borel \\<^sub>M prob_algebra borel}" proof safe fix \ assume h:"\ \ qbs_Mx (monadP_qbs X)" then obtain \ k where h1: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" using rep_qbs_Mx_monadM[of \ X] by(simp add: monadP_qbs_def sub_qbs_Mx) metis interpret s_finite_kernel borel borel k by fact have "\ \ UNIV \ {s. qbs_space_of s = X \ prob_space (qbs_l s)}" using h qbs_Mx_to_X[OF h] by(auto simp: monadP_qbs_def sub_qbs_Mx monadM_qbs_space sub_qbs_space) hence "\r. prob_space (k r)" using h1(2) by(auto simp add: h1(1) Pi_iff qbs_s_finite.qbs_l[OF h1(4)] intro!: prob_space_distrD[of \ _ "qbs_to_measure X"]) hence "prob_kernel borel borel k" by(auto simp: prob_kernel_def prob_kernel_axioms_def measure_kernel_axioms) with h1(1,2) show "\\ k. \ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx X \ k \ borel \\<^sub>M prob_algebra borel" by(auto intro!: exI[where x=\] exI[where x=k] simp: prob_kernel_def') next fix \ and k :: "real \ real measure" assume h:"\ \ qbs_Mx X" "k \ borel \\<^sub>M prob_algebra borel" then interpret pk: prob_kernel borel borel k by(simp add: prob_kernel_def'[symmetric]) have qp: "qbs_prob X \ (k r)" for r using h by(auto simp: qbs_prob_def in_Mx_def pk.kernel_sets pk.prob_spaces real_distribution_axioms_def real_distribution_def) show "(\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ qbs_Mx (monadP_qbs X)" using h(1) qp by(auto simp: monadP_qbs_def sub_qbs_Mx monadM_qbs_space qbs_s_finite.qbs_l[OF qbs_prob.qbs_s_finite[OF qp]] qbs_s_finite.qbs_space_of[OF qbs_prob.qbs_s_finite[OF qp]] monadM_qbs_Mx qbs_prob_def real_distribution_def intro!: exI[where x=\] exI[where x=k] h pk.s_finite_kernel_axioms prob_space.prob_space_distr) qed lemma rep_qbs_Mx_monadP: assumes "\ \ qbs_Mx (monadP_qbs X)" obtains \ k where "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "k \ borel \\<^sub>M prob_algebra borel" "\r. qbs_prob X \ (k r)" proof - have "\\ r k. \ \ qbs_Mx X \ k \ borel \\<^sub>M prob_algebra borel \ qbs_prob X \ (k r)" by(auto simp: qbs_prob_def in_Mx_def real_distribution_def real_distribution_axioms_def prob_kernel_def'[symmetric] prob_kernel_def prob_kernel_axioms_def measure_kernel_def) thus ?thesis using assms that by(fastforce simp: monadP_qbs_Mx) qed lemma qbs_l_monadP_le1:"s \ qbs_space (monadP_qbs X) \ qbs_l s A \ 1" by(auto simp: monadP_qbs_space intro!: prob_space.emeasure_le_1) lemma qbs_l_inj_P: "inj_on qbs_l (qbs_space (monadP_qbs X))" by(auto intro!: inj_on_subset[OF qbs_l_inj] simp: monadP_qbs_def sub_qbs_space) lemma qbs_l_measurable_prob[measurable]:"qbs_l \ qbs_to_measure (monadP_qbs X) \\<^sub>M prob_algebra (qbs_to_measure X)" proof(rule qbs_morphism_dest[OF qbs_morphismI]) fix \ assume "\ \ qbs_Mx (monadP_qbs X)" from rep_qbs_Mx_monadP[OF this] obtain \ k where h[measurable]: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "k \ borel \\<^sub>M prob_algebra borel" "\r. qbs_prob X \ (k r)" by metis show "qbs_l \ \ \ qbs_Mx (measure_to_qbs (prob_algebra (qbs_to_measure X)))" by(auto simp: qbs_Mx_R comp_def h(1) qbs_s_finite.qbs_l[OF qbs_prob.qbs_s_finite[OF h(4)]]) qed lemma return_qbs_morphismP: "return_qbs X \ X \\<^sub>Q monadP_qbs X" proof(rule qbs_morphismI) interpret rr : real_distribution "return borel 0" by(simp add: real_distribution_def real_distribution_axioms_def prob_space_return) fix \ assume h:"\ \ qbs_Mx X" then have 1:"return_qbs X \ \ = (\r. \X, \, return borel r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(rule return_qbs_comp) show "return_qbs X \ \ \ qbs_Mx (monadP_qbs X)" by(auto simp: 1 monadP_qbs_Mx h intro!: exI[where x=\] exI[where x="return borel"]) qed lemma(in qbs_prob) assumes "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "f \ X \\<^sub>Q monadP_qbs Y" "\ \ qbs_Mx Y" and g[measurable]:"g \ borel \\<^sub>M prob_algebra borel" and "(f \ \) = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows bind_qbs_prob:"qbs_prob Y \ (\ \ g)" and bind_qbs': "s \ f = \Y, \, \ \ g\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - interpret prob_kernel borel borel g using assms(4) by(simp add: prob_kernel_def') have "prob_space (\ \ g)" by(auto intro!: prob_space_bind'[OF _ g] simp: space_prob_algebra prob_space_axioms) thus "qbs_prob Y \ (\ \ g)" "s \ f = \Y, \, \ \ g\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite.qbs_probI[OF bind_qbs_s_finite[OF assms(1) qbs_morphism_monadPD[OF assms(2)] assms(3) s_finite_kernel_axioms assms(5)]] by(simp_all add: bind_qbs[OF assms(1) qbs_morphism_monadPD[OF assms(2)] assms(3) s_finite_kernel_axioms assms(5)] bind_kernel_bind[of g \ borel]) qed lemma bind_qbs_morphism'P: assumes "f \ X \\<^sub>Q monadP_qbs Y" shows "(\x. x \ f) \ monadP_qbs X \\<^sub>Q monadP_qbs Y" proof(safe intro!: qbs_morphism_monadPI') fix x assume "x \ qbs_space (monadP_qbs X)" from rep_qbs_space_monadP[OF this] obtain \ \ where h:"x = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_prob X \ \" by metis then interpret qbs_prob X \ \ by simp from rep_qbs_Mx_monadP[OF qbs_morphism_Mx[OF assms in_Mx]] obtain \ g where h'[measurable]: "f \ \ = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "g \ borel \\<^sub>M prob_algebra borel" by metis show "x \ f \ qbs_space (monadP_qbs Y)" using sets_bind[of \ g] measurable_space[OF h'(3),simplified space_prob_algebra] by(auto simp: qbs_prob.bind_qbs'[OF h(2,1) assms h'(2,3,1)] qbs_prob_def in_Mx_def h'(2) real_distribution_def real_distribution_axioms_def intro!: qbs_prob.in_space_monadP prob_space_bind[where S=borel] measurable_prob_algebraD) qed(auto intro!: qbs_morphism_monadPD' bind_qbs_morphism'[OF qbs_morphism_monadPD[OF assms]]) lemma distr_qbs_morphismP': assumes "f \ X \\<^sub>Q Y" shows "distr_qbs X Y f \ monadP_qbs X \\<^sub>Q monadP_qbs Y" unfolding distr_qbs_def by(rule bind_qbs_morphism'P[OF qbs_morphism_comp[OF assms return_qbs_morphismP]]) lemma join_qbs_morphismP: "join_qbs \ monadP_qbs (monadP_qbs X) \\<^sub>Q monadP_qbs X" by(simp add: join_qbs_def bind_qbs_morphism'P[OF qbs_morphism_ident]) lemma assumes "qbs_prob (monadP_qbs X) \ \" "ssx = \monadP_qbs X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "\ \ qbs_Mx X" "g \ borel \\<^sub>M prob_algebra borel" and "\ =(\r. \X, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows qbs_prob_join_qbs_s_finite: "qbs_prob X \ (\ \ g)" and qbs_prob_join_qbs: "join_qbs ssx = \X, \, \ \ g\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_prob.bind_qbs'[OF assms(1,2) qbs_morphism_ident assms(3,4)] qbs_prob.bind_qbs_prob[OF assms(1,2) qbs_morphism_ident assms(3,4)] by(auto simp: assms(5) join_qbs_def) context begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) lemma strength_qbs_ab_r_prob: assumes "\ \ qbs_Mx X" "\ \ qbs_Mx (monadP_qbs Y)" "\ \ qbs_Mx Y" and [measurable]:"g \ borel \\<^sub>M prob_algebra borel" and "\ = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" shows "qbs_prob (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (return borel r \\<^sub>M g r) borel rr.to_real)" using measurable_space[OF assms(4),of r] sets_return[of borel r] by(auto intro!: qbs_s_finite.qbs_probI strength_qbs_ab_r_s_finite[OF assms(1) qbs_Mx_monadPM[OF assms(2)] assms(3) prob_kernel.s_finite_kernel_prob_kernel assms(5),simplified prob_kernel_def',OF assms(4)] prob_space.prob_space_distr prob_space_pair prob_space_return simp: space_prob_algebra simp del: sets_return) lemma strength_qbs_morphismP: "strength_qbs X Y \ X \\<^sub>Q monadP_qbs Y \\<^sub>Q monadP_qbs (X \\<^sub>Q Y)" proof(rule pair_qbs_morphismI) fix \ \ assume h:"\ \ qbs_Mx X" "\ \ qbs_Mx (monadP_qbs Y)" from rep_qbs_Mx_monadP[OF this(2)] obtain \ g where hb[measurable]: "\ = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "g \ borel \\<^sub>M prob_algebra borel" by metis show "(\r. strength_qbs X Y (\ r, \ r)) \ qbs_Mx (monadP_qbs (X \\<^sub>Q Y))" using strength_qbs_ab_r_prob[OF h hb(2,3,1)] strength_qbs_ab_r[OF h(1) qbs_Mx_monadPM[OF h(2)] hb(2) prob_kernel.s_finite_kernel_prob_kernel hb(1),simplified prob_kernel_def',OF hb(3)] by(auto simp: monadP_qbs_Mx qbs_prob_def in_Mx_def intro!: exI[where x="map_prod \ \ \ rr.from_real"] exI[where x="\r. distr (return borel r \\<^sub>M g r) borel rr.to_real"]) qed end lemma bind_qbs_morphismP: "(\) \ monadP_qbs X \\<^sub>Q (X \\<^sub>Q monadP_qbs Y) \\<^sub>Q monadP_qbs Y" proof - { fix f s assume h:"f \ X \\<^sub>Q monadP_qbs Y" "s \ qbs_space (monadP_qbs X)" from rep_qbs_space_monadP[OF this(2)] obtain \ \ where h': "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_prob X \ \" by metis then interpret qbs_prob X \ \ by simp from rep_qbs_Mx_monadP[OF qbs_morphism_Mx[OF h(1) in_Mx]] obtain \ g where hb[measurable]:"f \ \ = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "g \ borel \\<^sub>M prob_algebra borel" by metis have "join_qbs (distr_qbs ((X \\<^sub>Q monadP_qbs Y) \\<^sub>Q X) (monadP_qbs Y) (\fx. fst fx (snd fx)) (strength_qbs (X \\<^sub>Q monadP_qbs Y) X (f, s))) = s \ f" using qbs_prob_join_qbs[OF qbs_prob.distr_qbs_prob[OF strength_qbs_prob[of f "X \\<^sub>Q monadP_qbs Y",OF h(1) h'(1)] strength_qbs[of f "X \\<^sub>Q monadP_qbs Y",OF h(1) h'(1)] qbs_morphism_eval] qbs_s_finite.distr_qbs[OF strength_qbs_s_finite[of f "X \\<^sub>Q monadP_qbs Y",OF h(1) h'(1)] strength_qbs[of f "X \\<^sub>Q monadP_qbs Y",OF h(1) h'(1)] qbs_morphism_eval] hb(2,3)] hb(1) by(simp add: bind_qbs[OF h'(1) qbs_morphism_monadPD[OF h(1)] hb(2) prob_kernel.s_finite_kernel_prob_kernel hb(1),simplified prob_kernel_def',OF hb(3)] comp_def bind_kernel_bind[of g \ borel,OF measurable_prob_algebraD]) } thus ?thesis by(auto intro!: arg_swap_morphism[OF curry_preserves_morphisms [OF qbs_morphism_cong'[of _ "join_qbs \ (distr_qbs (exp_qbs X (monadP_qbs Y) \\<^sub>Q X) (monadP_qbs Y) (\fx. (fst fx) (snd fx))) \ (strength_qbs (exp_qbs X (monadP_qbs Y)) X)"]]] qbs_morphism_comp distr_qbs_morphismP' strength_qbs_morphismP join_qbs_morphismP qbs_morphism_eval simp: pair_qbs_space) qed corollary strength_qbs_law1P: assumes "x \ qbs_space (unit_quasi_borel \\<^sub>Q monadP_qbs X)" shows "snd x = (distr_qbs (unit_quasi_borel \\<^sub>Q X) X snd \ strength_qbs unit_quasi_borel X) x" by(rule strength_qbs_law1, insert assms) (auto simp: pair_qbs_space qbs_space_monadPM) corollary strength_qbs_law2P: assumes "x \ qbs_space ((X \\<^sub>Q Y) \\<^sub>Q monadP_qbs Z)" shows "(strength_qbs X (Y \\<^sub>Q Z) \ (map_prod id (strength_qbs Y Z)) \ (\((x,y),z). (x,(y,z)))) x = (distr_qbs ((X \\<^sub>Q Y) \\<^sub>Q Z) (X \\<^sub>Q (Y \\<^sub>Q Z)) (\((x,y),z). (x,(y,z))) \ strength_qbs (X \\<^sub>Q Y) Z) x" by(rule strength_qbs_law2, insert assms) (auto simp: pair_qbs_space qbs_space_monadPM) lemma strength_qbs_law4P: assumes "x \ qbs_space (X \\<^sub>Q monadP_qbs (monadP_qbs Y))" shows "(strength_qbs X Y \ map_prod id join_qbs) x = (join_qbs \ distr_qbs (X \\<^sub>Q monadP_qbs Y) (monadP_qbs (X \\<^sub>Q Y)) (strength_qbs X Y) \ strength_qbs X (monadP_qbs Y)) x" (is "?lhs = ?rhs") proof - from assms rep_qbs_space_monadP[of "snd x" "monadP_qbs Y"] obtain \ \ where h:"qbs_prob (monadP_qbs Y) \ \" "snd x = \monadP_qbs Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by (auto simp: pair_qbs_space) metis then interpret qp: qbs_prob "monadP_qbs Y" \ \ by simp from rep_qbs_Mx_monadP[OF qp.in_Mx] obtain \ g where h': "\ \ qbs_Mx Y" "g \ borel \\<^sub>M prob_algebra borel" "\ = (\r. \Y, \, g r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" and h'': "\r. qbs_prob Y \ (g r)" by metis have "?lhs = \X \\<^sub>Q Y, \r. (fst x, \ r), \ \ g\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite.strength_qbs[OF qbs_prob.qbs_s_finite[OF qbs_prob_join_qbs_s_finite[OF h h']] _ qbs_prob_join_qbs[OF h h'],of "fst x" X] assms by (auto simp: pair_qbs_space) also have "... = ?rhs" using qbs_prob_join_qbs[OF qbs_prob.distr_qbs_prob[OF qp.strength_qbs_prob[OF _ h(2),of "fst x" X] qp.strength_qbs[OF _ h(2)] strength_qbs_morphismP] qbs_s_finite.distr_qbs[OF qp.strength_qbs_s_finite[OF _ h(2),of "fst x" X] qp.strength_qbs[OF _ h(2)] strength_qbs_morphismP] pair_qbs_MxI h'(2),of "\r. (fst x, \ r)",simplified comp_def qbs_s_finite.strength_qbs[OF qbs_prob.qbs_s_finite[OF h''] _ fun_cong[OF h'(3)]]] assms by(auto simp: pair_qbs_space h') finally show ?thesis . qed lemma distr_qbs_morphismP: "distr_qbs X Y \ X \\<^sub>Q Y \\<^sub>Q monadP_qbs X \\<^sub>Q monadP_qbs Y" proof - note [qbs] = bind_qbs_morphismP return_qbs_morphismP have [simp]: "distr_qbs X Y = (\f sx. sx \ return_qbs Y \ f)" by standard+ (auto simp: distr_qbs_def) show ?thesis by simp qed lemma bind_qbs_return_rotateP: assumes "p \ qbs_space (monadP_qbs X)" and "q \ qbs_space (monadP_qbs Y)" shows "q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y))) = p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))" by(auto intro!: bind_qbs_return_rotate qbs_space_monadPM assms) lemma qbs_bind_bind_return1P: assumes "f \ X \\<^sub>Q Y \\<^sub>Q monadP_qbs Z" "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" shows "q \ (\y. p \ (\x. f (x,y))) = (q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y)))) \ f" by(auto intro!: qbs_bind_bind_return1 assms qbs_space_monadPM qbs_morphism_monadPD) corollary qbs_bind_bind_return1P': assumes [qbs]:"f \ qbs_space (X \\<^sub>Q Y \\<^sub>Q monadP_qbs Z)" "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" shows "q \ (\y. p \ (\x. f x y)) = (q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y)))) \ (case_prod f)" by(auto intro!: qbs_bind_bind_return1P[where f="case_prod f" and Z=Z,simplified]) lemma qbs_bind_bind_return2P: assumes "f \ X \\<^sub>Q Y \\<^sub>Q monadP_qbs Z" "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" shows "p \ (\x. q \ (\y. f (x,y))) = (p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))) \ f" by(auto intro!: qbs_bind_bind_return2 assms qbs_space_monadPM qbs_morphism_monadPD) corollary qbs_bind_bind_return2P': assumes [qbs]:"f \ qbs_space (X \\<^sub>Q Y \\<^sub>Q monadP_qbs Z)" "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" shows "p \ (\x. q \ (\y. f x y)) = (p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))) \ (case_prod f)" by(auto intro!: qbs_bind_bind_return2P[where f="case_prod f" and Z=Z,simplified]) corollary bind_qbs_rotateP: assumes "f \ X \\<^sub>Q Y \\<^sub>Q monadP_qbs Z" "p \ qbs_space (monadP_qbs X)" and "q \ qbs_space (monadP_qbs Y)" shows "q \ (\y. p \ (\x. f (x,y))) = p \ (\x. q \ (\y. f (x,y)))" by(auto intro!: bind_qbs_rotate assms qbs_space_monadPM qbs_morphism_monadPD) context pair_qbs_probs begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) sublocale qbs_prob "X \\<^sub>Q Y" "map_prod \ \ \ rr.from_real" "distr (\ \\<^sub>M \) borel rr.to_real" by(auto simp: qbs_prob_def in_Mx_def real_distribution_def qbs_Mx_is_morphisms real_distribution_axioms_def pq1.prob_space_axioms pq2.prob_space_axioms intro!: prob_space.prob_space_distr prob_space_pair) lemma qbs_bind_bind_return_prob: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q Z" shows "qbs_prob Z (f \ (map_prod \ \ \ rr.from_real)) (distr (\ \\<^sub>M \) borel rr.to_real)" using qbs_prob_axioms by(auto simp: qbs_prob_def in_Mx_def qbs_Mx_is_morphisms) end subsubsection \ Almost Everywhere \ lift_definition qbs_almost_everywhere :: "['a qbs_measure, 'a \ bool] \ bool" is "\(X,\,\). almost_everywhere (distr \ (qbs_to_measure X) \)" by(auto simp: qbs_s_finite_eq_def) metis syntax "_qbs_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE\<^sub>Q _ in _. _" [0,0,10] 10) translations "AE\<^sub>Q x in p. P" \ "CONST qbs_almost_everywhere p (\x. P)" lemma AEq_qbs_l: "(AE\<^sub>Q x in p. P x) = (AE x in qbs_l p. P x)" by transfer (simp add: case_prod_beta') lemma(in qbs_s_finite) AEq_def: "(AE\<^sub>Q x in \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n . P x) = (AE x in (distr \ (qbs_to_measure X) \). P x)" by(simp add: qbs_almost_everywhere.abs_eq) lemma(in qbs_s_finite) AEq_AE: "(AE\<^sub>Q x in \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n . P x) \ (AE x in \. P (\ x))" by(auto simp: AEq_def intro!:AE_distrD[of \]) lemma(in qbs_s_finite) AEq_AE_iff: assumes [qbs]:"qbs_pred X P" shows "(AE\<^sub>Q x in \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n . P x) \ (AE x in \. P (\ x))" by(auto simp: AEq_AE AEq_def qbs_pred_iff_sets intro!: AE_distr_iff[THEN iffD2]) lemma AEq_qbs_pred[qbs]: "qbs_almost_everywhere \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q qbs_count_space UNIV) \\<^sub>Q qbs_count_space UNIV" proof(rule curry_preserves_morphisms[OF pair_qbs_morphismI]) fix \ \ assume h:"\ \ qbs_Mx (monadM_qbs X)" "\ \ qbs_Mx (X \\<^sub>Q qbs_count_space (UNIV :: bool set))" from rep_qbs_Mx_monadM[OF h(1)] obtain \ k where hk: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis interpret s:standard_borel_ne "borel :: real measure" by simp interpret s2: standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(simp add: borel_prod) have [measurable]:"Measurable.pred (borel \\<^sub>M borel) (\(x, y). \ x (\ y))" using h(2) hk(2) by(simp add: s2.qbs_pred_iff_measurable_pred[symmetric] r_preserves_product qbs_Mx_is_morphisms) show "(\r. qbs_almost_everywhere (fst (\ r, \ r)) (snd (\ r, \ r))) \ qbs_Mx (qbs_count_space UNIV)" using h(2) hk(2) by(simp add: hk(1) qbs_Mx_is_morphisms qbs_s_finite.AEq_AE_iff[OF hk(4)]) (auto simp add: s.qbs_pred_iff_measurable_pred intro!: s_finite_kernel.AE_pred[OF hk(3)]) qed lemma AEq_I2[simp]: assumes "p \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ P x" shows "AE\<^sub>Q x in p. P x" by(auto simp: space_qbs_l_in[OF assms(1)] assms(2) AEq_qbs_l) lemma AEq_mp[elim!]: assumes "AE\<^sub>Q x in s. P x" "AE\<^sub>Q x in s. P x \ Q x" shows "AE\<^sub>Q x in s. Q x" using assms by(auto simp: AEq_qbs_l) lemma shows AEq_iffI: "AE\<^sub>Q x in s. P x \ AE\<^sub>Q x in s. P x \ Q x \ AE\<^sub>Q x in s. Q x" and AEq_disjI1: "AE\<^sub>Q x in s. P x \ AE\<^sub>Q x in s. P x \ Q x" and AEq_disjI2: "AE\<^sub>Q x in s. Q x \ AE\<^sub>Q x in s. P x \ Q x" and AEq_conjI: "AE\<^sub>Q x in s. P x \ AE\<^sub>Q x in s. Q x \ AE\<^sub>Q x in s. P x \ Q x" and AEq_conj_iff[simp]: "(AE\<^sub>Q x in s. P x \ Q x) \ (AE\<^sub>Q x in s. P x) \ (AE\<^sub>Q x in s. Q x)" by(auto simp: AEq_qbs_l) lemma AEq_symmetric: assumes "AE\<^sub>Q x in s. P x = Q x" shows "AE\<^sub>Q x in s. Q x = P x" using assms by(auto simp: AEq_qbs_l) lemma AEq_impI: "(P \ AE\<^sub>Q x in M. Q x) \ AE\<^sub>Q x in M. P \ Q x" by(auto simp: AEq_qbs_l AE_impI) lemma AEq_Ball_mp: "s \ qbs_space (monadM_qbs X) \ (\x. x\qbs_space X \ P x) \ AE\<^sub>Q x in s. P x \ Q x \ AE\<^sub>Q x in s. Q x" by auto lemma AEq_cong: "s \ qbs_space (monadM_qbs X) \ (\x. x \ qbs_space X \ P x \ Q x) \ (AE\<^sub>Q x in s. P x) \ (AE\<^sub>Q x in s. Q x)" by auto lemma AEq_cong_simp: "s \ qbs_space (monadM_qbs X) \ (\x. x \ qbs_space X =simp=> P x = Q x) \ (AE\<^sub>Q x in s. P x) \ (AE\<^sub>Q x in s. Q x)" by (auto simp: simp_implies_def) lemma AEq_all_countable: "(AE\<^sub>Q x in s. \i. P i x) \ (\i::'i::countable. AE\<^sub>Q x in s. P i x)" by(simp add: AEq_qbs_l AE_all_countable) lemma AEq_ball_countable: "countable X \ (AE\<^sub>Q x in s. \y\X. P x y) \ (\y\X. AE\<^sub>Q x in s. P x y)" by(simp add: AEq_qbs_l AE_ball_countable) lemma AEq_ball_countable': "(\N. N \ I \ AE\<^sub>Q x in s. P N x) \ countable I \ AE\<^sub>Q x in s. \N \ I. P N x" unfolding AEq_ball_countable by simp lemma AEq_pairwise: "countable F \ pairwise (\A B. AE\<^sub>Q x in s. R x A B) F \ (AE\<^sub>Q x in s. pairwise (R x) F)" unfolding pairwise_alt by (simp add: AEq_ball_countable) lemma AEq_finite_all: "finite S \ (AE\<^sub>Q x in s. \i\S. P i x) \ (\i\S. AE\<^sub>Q x in s. P i x)" by(simp add: AEq_qbs_l AE_finite_all) lemma AE_finite_allI:"finite S \ (\s. s \ S \ AE\<^sub>Q x in M. Q s x) \ AE\<^sub>Q x in M. \s\S. Q s x" by(simp add: AEq_qbs_l AE_finite_all) subsubsection \ Integral \ lift_definition qbs_nn_integral :: "['a qbs_measure, 'a \ ennreal] \ ennreal" is "\(X,\,\) f.(\\<^sup>+x. f x \distr \ (qbs_to_measure X) \)" by(auto simp: qbs_s_finite_eq_def) lift_definition qbs_integral :: "['a qbs_measure, 'a \ ('b :: {banach,second_countable_topology})] \ 'b" is "\p f. if f \ (fst p) \\<^sub>Q qbs_borel then (\x. f (fst (snd p) x) \ (snd (snd p))) else 0" using qbs_s_finite_eq_dest(3) qbs_s_finite_eq_1_imp_2 by fastforce syntax "_qbs_nn_integral" :: "pttrn \ ennreal \ 'a qbs_measure \ ennreal" ("\\<^sup>+\<^sub>Q((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sup>+\<^sub>Q x. f \p" \ "CONST qbs_nn_integral p (\x. f)" syntax "_qbs_integral" :: "pttrn \ _ \ 'a qbs_measure \ _" ("\\<^sub>Q((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sub>Q x. f \p" \ "CONST qbs_integral p (\x. f)" lemma(in qbs_s_finite) shows qbs_nn_integral_def: "f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+\<^sub>Q x. f x \\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = (\\<^sup>+x. f (\ x) \ \)" and qbs_nn_integral_def2:"(\\<^sup>+\<^sub>Q x. f x \\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = (\\<^sup>+x. f x \(distr \ (qbs_to_measure X) \))" by(simp_all add: qbs_nn_integral.abs_eq nn_integral_distr lr_adjunction_correspondence) lemma(in qbs_s_finite) qbs_integral_def: "f \ X \\<^sub>Q qbs_borel \ (\\<^sub>Q x. f x \\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = (\x. f (\ x) \ \)" by(simp add: qbs_integral.abs_eq) lemma(in qbs_s_finite) qbs_integral_def2: "(\\<^sub>Q x. f x \\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = (\x. f x \(distr \ (qbs_to_measure X) \))" proof - consider "f \ X \\<^sub>Q qbs_borel" | "f \ X \\<^sub>Q qbs_borel" by auto thus ?thesis proof cases case h:2 then have "\ integrable (qbs_l \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) f" by (metis borel_measurable_integrable measurable_distr_eq1 qbs_l qbs_morphism_measurable_intro) thus ?thesis using h by(simp add: qbs_l qbs_integral.abs_eq lr_adjunction_correspondence not_integrable_integral_eq) qed(simp add: qbs_integral.abs_eq lr_adjunction_correspondence integral_distr) qed lemma qbs_measure_eqI: assumes [qbs]:"p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs X)" and "\f. f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+\<^sub>Q x. f x \p) = (\\<^sup>+\<^sub>Q x. f x \q)" shows "p = q" proof - obtain \ \ \ \ where h:"p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "q = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" "qbs_s_finite X \ \" by (metis rep_qbs_space_monadM assms(1,2)) then interpret pq:pair_qbs_s_finite X \ \ \ \ by(auto simp: pair_qbs_s_finite_def) show ?thesis using assms(3) by(auto simp: h(1,2) pq.pq1.qbs_nn_integral_def pq.pq2.qbs_nn_integral_def intro!: pq.qbs_s_finite_measure_eq') qed lemma qbs_nn_integral_def2_l: "qbs_nn_integral s f = integral\<^sup>N (qbs_l s) f" by transfer auto lemma qbs_integral_def2_l: "qbs_integral s f = integral\<^sup>L (qbs_l s) f" by (metis in_qbs_space_of qbs_s_finite.qbs_integral_def2 qbs_s_finite.qbs_l rep_qbs_space_monadM) lift_definition qbs_integrable :: "'a qbs_measure \ ('a \ 'b::{second_countable_topology,banach}) \ bool" is "\p f. f \ fst p \\<^sub>Q qbs_borel \ integrable (snd (snd p)) (f \ (fst (snd p)))" proof safe have 0:"f \ Y \\<^sub>Q qbs_borel" "integrable \ (\x. f (\ x))" if "qbs_s_finite_eq (X,\,\) (Y,\,\)" "f \ X \\<^sub>Q qbs_borel" "integrable \ (\x. f (\ x))" for X :: "'a quasi_borel" and Y \ \ \ \ and f :: "_ \ 'b" proof - interpret pair_qbs_s_finite X \ \ \ \ using qbs_s_finite_eq_dest[OF that(1)] by(auto simp: pair_qbs_s_finite_def) show "f \ Y \\<^sub>Q qbs_borel" "integrable \ (\x. f (\ x))" using that qbs_s_finite_eq_dest(3)[OF that(1)] by(simp_all add: integrable_distr_eq[symmetric,of \ \ "qbs_to_measure X" f] integrable_distr_eq[symmetric,of \ \ "qbs_to_measure X" f] lr_adjunction_correspondence qbs_s_finite_eq_dest(4)[OF that(1)]) qed { fix X Y :: "'a quasi_borel" fix \ \ \ \ and f :: "_ \ 'b" assume 1:"qbs_s_finite_eq (X, \, \) (Y, \, \)" then have 2:"qbs_s_finite_eq (Y, \, \) (X, \, \)" by(auto simp: qbs_s_finite_eq_def) have "f \ X \\<^sub>Q qbs_borel \ integrable \ (f \ \) \ f \ Y \\<^sub>Q qbs_borel \ integrable \ (f \ \)" unfolding comp_def using 0[OF 1,of f] 0[OF 2,of f] by blast } thus "\prod1 prod2 :: 'a qbs_s_finite_t. qbs_s_finite_eq prod1 prod2 \ (\f:: _ \ 'b. f \ fst prod1 \\<^sub>Q borel\<^sub>Q \ integrable (snd (snd prod1)) (f \ fst (snd prod1))) = (\f. f \ fst prod2 \\<^sub>Q borel\<^sub>Q \ integrable (snd (snd prod2)) (f \ fst (snd prod2)))" by fastforce qed lemma(in qbs_s_finite) qbs_integrable_def: "qbs_integrable \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n f \ f \ X \\<^sub>Q qbs_borel \ integrable \ (\x. f (\ x))" by(simp add: qbs_integrable.abs_eq comp_def) lemma qbs_integrable_morphism_dest: assumes "s \ qbs_space (monadM_qbs X)" and "qbs_integrable s f" shows "f \ X \\<^sub>Q qbs_borel" by (metis assms qbs_s_finite.qbs_integrable_def rep_qbs_space_monadM) lemma qbs_integrable_morphismP: assumes "s \ qbs_space (monadP_qbs X)" and "qbs_integrable s f" shows "f \ X \\<^sub>Q qbs_borel" by(auto intro!: qbs_integrable_morphism_dest assms qbs_space_monadPM) lemma(in qbs_s_finite) qbs_integrable_measurable[simp]: assumes "qbs_integrable \X,\,\\\<^sub>s\<^sub>f\<^sub>i\<^sub>n f" shows "f \ qbs_to_measure X \\<^sub>M borel" by(auto intro!: qbs_integrable_morphism_dest assms simp: lr_adjunction_correspondence[symmetric]) lemma qbs_integrable_iff_integrable: "(qbs_integrable (s::'a qbs_measure) (f :: 'a \ 'b::{second_countable_topology,banach})) = (integrable (qbs_l s) f)" proof transfer fix f ::" 'a \ 'b::{second_countable_topology,banach}" show "qbs_s_finite_eq s s \ (f \ fst s \\<^sub>Q borel\<^sub>Q \ integrable (snd (snd s)) (f \ fst (snd s))) = integrable (distr (snd (snd s)) (qbs_to_measure (fst s)) (fst (snd s))) f" for s proof(rule prod_cases3[of s]) fix X :: "'a quasi_borel" fix \ \ assume "qbs_s_finite_eq s s" and s: "s = (X,\,\)" then interpret qbs_s_finite X \ \ by(simp add: qbs_s_finite_eq_def) show "f \ fst s \\<^sub>Q qbs_borel \ integrable (snd (snd s)) (\x. (f \ fst (snd s)) x) \ integrable (distr (snd (snd s)) (qbs_to_measure (fst s)) (fst (snd s))) f" using integrable_distr_eq[of \ \ "qbs_to_measure X" f,simplified] by(auto simp add: lr_adjunction_correspondence s) qed qed corollary(in qbs_s_finite) qbs_integrable_distr: "qbs_integrable \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n f = integrable (distr \ (qbs_to_measure X) \) f" by(simp add: qbs_integrable_iff_integrable qbs_l) lemma qbs_integrable_morphism[qbs]: "qbs_integrable \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q (qbs_borel :: ('a :: {banach, second_countable_topology}) quasi_borel)) \\<^sub>Q qbs_count_space UNIV" proof(rule curry_preserves_morphisms[OF pair_qbs_morphismI]) fix \ \ assume h:"\ \ qbs_Mx (monadM_qbs X)" "\ \ qbs_Mx (X \\<^sub>Q (qbs_borel :: 'a quasi_borel))" from rep_qbs_Mx_monadM[OF this(1)] obtain \ k where hk:"\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis then interpret ina: in_Mx X \ by (simp add: in_Mx_def) interpret standard_borel_ne "borel :: real measure" by simp have [measurable]: "\ r \ qbs_to_measure X \\<^sub>M borel" for r using h(2) by(simp add: qbs_Mx_is_morphisms lr_adjunction_correspondence[symmetric]) have [measurable_cong]: "sets (k r) = sets borel" for r using hk(4) qbs_s_finite.mu_sets by blast have 1: "borel_measurable (borel \\<^sub>M borel) = (qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel :: (real \ real \ 'a) set)" by (metis borel_prod pair_standard_borel qbs_borel_prod standard_borel.standard_borel_r_full_faithful standard_borel_axioms) show "(\r. qbs_integrable (fst (\ r, \ r)) (snd (\ r, \ r))) \ qbs_Mx (qbs_count_space UNIV)" by(auto simp: fun_cong[OF hk(1)] qbs_s_finite.qbs_integrable_distr[OF hk(4)] integrable_distr_eq qbs_Mx_is_morphisms qbs_pred_iff_measurable_pred intro!: s_finite_kernel.integrable_measurable_pred[OF hk(3)]) (insert h(2), simp add: 1 qbs_Mx_is_morphisms split_beta') qed lemma(in qbs_s_finite) qbs_integrable_iff_integrable: assumes "f \ qbs_to_measure X \\<^sub>M borel" shows "qbs_integrable \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n f = integrable \ (\x. f (\ x))" by(auto intro!: integrable_distr_eq[of \ \ "qbs_to_measure X" f] simp: assms qbs_integrable_distr) lemma qbs_integrable_iff_bounded: assumes "s \ qbs_space (monadM_qbs X)" shows "qbs_integrable s f \ f \ X \\<^sub>Q qbs_borel \ (\\<^sup>+\<^sub>Q x. ennreal (norm (f x)) \s) < \" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms] obtain \ \ where hs: "qbs_s_finite X \ \" "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by metis then interpret qs:qbs_s_finite X \ \ by simp have "?lhs = integrable (distr \ (qbs_to_measure X) \) f" by (simp add: hs(2) qbs_integrable_iff_integrable qs.qbs_l) also have "... = (f \ borel_measurable (distr \ (qbs_to_measure X) \) \ ((\\<^sup>+ x. ennreal (norm (f x)) \(distr \ (qbs_to_measure X) \)) < \))" by(rule integrable_iff_bounded) also have "... = ?rhs" by(auto simp add: hs(2) qs.qbs_nn_integral_def2 lr_adjunction_correspondence) finally show ?thesis . qed lemma not_qbs_integrable_qbs_integral: "\ qbs_integrable s f \ qbs_integral s f = 0" by(simp add: qbs_integral_def2_l qbs_integrable_iff_integrable not_integrable_integral_eq) lemma qbs_integrable_cong_AE: assumes "s \ qbs_space (monadM_qbs X)" "AE\<^sub>Q x in s. f x = g x" and "qbs_integrable s f" "g \ X \\<^sub>Q qbs_borel" shows "qbs_integrable s g" using assms(2,4) by(auto intro!: qbs_integrable_iff_integrable[THEN iffD2] Bochner_Integration.integrable_cong_AE[of g _ f,THEN iffD2] qbs_integrable_iff_integrable[THEN iffD1,OF assms(3)] qbs_integrable_morphism_dest[OF assms(1),of f] simp: AEq_qbs_l measurable_qbs_l[OF assms(1)]) lemma qbs_integrable_cong: assumes "s \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ f x = g x" and "qbs_integrable s f" shows "qbs_integrable s g" by(auto intro!: qbs_integrable_iff_integrable[THEN iffD2] Bochner_Integration.integrable_cong[OF refl,of _ g f,THEN iffD2] qbs_integrable_iff_integrable[THEN iffD1,OF assms(3)] simp: space_qbs_l_in[OF assms(1)] assms(2)) lemma qbs_integrable_zero[simp, intro]: "qbs_integrable s (\x. 0)" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_const: assumes "s \ qbs_space (monadP_qbs X)" shows "qbs_integrable s (\x. c)" using assms by(auto intro!: qbs_integrable_iff_integrable[THEN iffD2] finite_measure.integrable_const simp: monadP_qbs_space prob_space_def) lemma qbs_integrable_add[simp,intro!]: assumes "qbs_integrable s f" and "qbs_integrable s g" shows "qbs_integrable s (\x. f x + g x)" by(rule qbs_integrable_iff_integrable[THEN iffD2,OF Bochner_Integration.integrable_add[OF qbs_integrable_iff_integrable[THEN iffD1,OF assms(1)] qbs_integrable_iff_integrable[THEN iffD1,OF assms(2)]]]) lemma qbs_integrable_diff[simp,intro!]: assumes "qbs_integrable s f" and "qbs_integrable s g" shows "qbs_integrable s (\x. f x - g x)" by(rule qbs_integrable_iff_integrable[THEN iffD2,OF Bochner_Integration.integrable_diff[OF qbs_integrable_iff_integrable[THEN iffD1,OF assms(1)] qbs_integrable_iff_integrable[THEN iffD1,OF assms(2)]]]) lemma qbs_integrable_sum[simp, intro!]: "(\i. i \ I \ qbs_integrable s (f i)) \ qbs_integrable s (\x. \i\I. f i x)" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_scaleR_left[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. f x *\<^sub>R (c :: 'a :: {second_countable_topology,banach}))" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_scaleR_right[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. c *\<^sub>R (f x :: 'a :: {second_countable_topology,banach}) )" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_mult_iff: fixes f :: "'a \ real" shows "(qbs_integrable s (\x. c * f x)) = (c = 0 \ qbs_integrable s f)" using qbs_integrable_iff_integrable[of s "\x. c * f x"] integrable_mult_left_iff[of _ c f] qbs_integrable_iff_integrable[of s f] by simp lemma fixes c :: "_::{real_normed_algebra,second_countable_topology}" assumes "qbs_integrable s f" shows qbs_integrable_mult_right:"qbs_integrable s (\x. c * f x)" and qbs_integrable_mult_left: "qbs_integrable s (\x. f x * c)" using assms by(auto simp: qbs_integrable_iff_integrable) lemma qbs_integrable_divide_zero[simp, intro!]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "qbs_integrable s f \ qbs_integrable s (\x. f x / c)" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_inner_left[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. f x \ c)" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_inner_right[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. c \ f x)" by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_minus[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. - f x)" by(simp add: qbs_integrable_iff_integrable) lemma [simp, intro]: assumes "qbs_integrable s f" shows qbs_integrable_Re: "qbs_integrable s (\x. Re (f x))" and qbs_integrable_Im: "qbs_integrable s (\x. Im (f x))" and qbs_integrable_cnj: "qbs_integrable s (\x. cnj (f x))" using assms by(simp_all add: qbs_integrable_iff_integrable) lemma qbs_integrable_of_real[simp, intro!]: "qbs_integrable s f \ qbs_integrable s (\x. of_real (f x))" by(simp_all add: qbs_integrable_iff_integrable) lemma [simp, intro]: assumes "qbs_integrable s f" shows qbs_integrable_fst: "qbs_integrable s (\x. fst (f x))" and qbs_integrable_snd: "qbs_integrable s (\x. snd (f x))" using assms by(simp_all add: qbs_integrable_iff_integrable) lemma qbs_integrable_norm: assumes "qbs_integrable s f" shows "qbs_integrable s (\x. norm (f x))" by(rule qbs_integrable_iff_integrable[THEN iffD2,OF integrable_norm[OF qbs_integrable_iff_integrable[THEN iffD1,OF assms]]]) lemma qbs_integrable_abs: fixes f :: "_ \ real" assumes "qbs_integrable s f" shows "qbs_integrable s (\x. \f x\)" by(rule qbs_integrable_iff_integrable[THEN iffD2,OF integrable_abs[OF qbs_integrable_iff_integrable[THEN iffD1,OF assms]]]) lemma qbs_integrable_sq: fixes c :: "_::{real_normed_field,second_countable_topology}" assumes "qbs_integrable s (\x. c)" "qbs_integrable s f" and "qbs_integrable s (\x. (f x)\<^sup>2)" shows "qbs_integrable s (\x. (f x - c)\<^sup>2)" by(simp add: comm_ring_1_class.power2_diff,rule qbs_integrable_diff,rule qbs_integrable_add) (simp_all add: comm_semiring_1_class.semiring_normalization_rules(16)[of 2] assms qbs_integrable_mult_right power2_eq_square[of c]) lemma qbs_nn_integral_eq_integral_AEq: assumes "qbs_integrable s f" "AE\<^sub>Q x in s. 0 \ f x" shows "(\\<^sup>+\<^sub>Q x. ennreal (f x) \s) = ennreal (\\<^sub>Q x. f x \s)" using nn_integral_eq_integral[OF qbs_integrable_iff_integrable[THEN iffD1,OF assms(1)] ] qbs_integrable_morphism_dest[OF in_qbs_space_of assms(1)] assms(2) by(simp add: qbs_integral_def2_l qbs_nn_integral_def2_l AEq_qbs_l) lemma qbs_nn_integral_eq_integral: assumes "s \ qbs_space (monadM_qbs X)" "qbs_integrable s f" and "\x. x \ qbs_space X \ 0 \ f x" shows "(\\<^sup>+\<^sub>Q x. ennreal (f x) \s) = ennreal (\\<^sub>Q x. f x \s)" using qbs_nn_integral_eq_integral_AEq[OF assms(2) AEq_I2[OF assms(1,3)]] by simp lemma qbs_nn_integral_cong_AEq: assumes "s \ qbs_space (monadM_qbs X)" "AE\<^sub>Q x in s. f x = g x" shows "qbs_nn_integral s f = qbs_nn_integral s g" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qs: qbs_s_finite X \ \ by simp show ?thesis using assms(2) by(auto simp: qs.qbs_nn_integral_def2 hs(1) qs.AEq_def intro!: nn_integral_cong_AE) qed lemma qbs_nn_integral_cong: assumes "s \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ f x = g x" shows "qbs_nn_integral s f = qbs_nn_integral s g" using qbs_nn_integral_cong_AEq[OF assms(1) AEq_I2[OF assms]] by simp lemma qbs_nn_integral_const: "(\\<^sup>+\<^sub>Q x. c \s) = c * qbs_l s (qbs_space (qbs_space_of s))" by(simp add: qbs_nn_integral_def2_l space_qbs_l) lemma qbs_nn_integral_const_prob: assumes "s \ qbs_space (monadP_qbs X)" shows "(\\<^sup>+\<^sub>Q x. c \s) = c" using assms by(simp add: qbs_nn_integral_const prob_space.emeasure_space_1 qbs_l_prob_space space_qbs_l) lemma qbs_nn_integral_add: assumes "s \ qbs_space (monadM_qbs X)" and [qbs]:"f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q x. f x + g x \s) = (\\<^sup>+\<^sub>Q x. f x \s) + (\\<^sup>+\<^sub>Q x. g x \s)" by(auto simp: qbs_nn_integral_def2_l measurable_qbs_l[OF assms(1)] intro!: nn_integral_add measurable_qbs_l) lemma qbs_nn_integral_cmult: assumes "s \ qbs_space (monadM_qbs X)" and [qbs]:"f \ X \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q x. c * f x \s) = c * (\\<^sup>+\<^sub>Q x. f x \s)" by(auto simp: qbs_nn_integral_def2_l measurable_qbs_l[OF assms(1)] intro!: nn_integral_cmult) lemma qbs_integral_cong_AEq: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" and "AE\<^sub>Q x in s. f x = g x" shows "qbs_integral s f = qbs_integral s g" using assms(4) by(auto simp: qbs_integral_def2_l AEq_qbs_l measurable_qbs_l[OF assms(1)] intro!: integral_cong_AE ) lemma qbs_integral_cong: assumes "s \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ f x = g x" shows "qbs_integral s f = qbs_integral s g" by(auto simp: qbs_integral_def2_l space_qbs_l_in[OF assms(1)] assms(2) intro!: Bochner_Integration.integral_cong) lemma qbs_integral_nonneg_AEq: fixes f :: "_ \ real" shows "AE\<^sub>Q x in s. 0 \ f x \ 0 \ qbs_integral s f" by(auto simp: qbs_integral_def2_l AEq_qbs_l intro!: integral_nonneg_AE ) lemma qbs_integral_nonneg: fixes f :: "_ \ real" assumes "s \ qbs_space (monadM_qbs X)" "\x. x \ qbs_space X \ 0 \ f x" shows "0 \ qbs_integral s f" by(auto simp: qbs_integral_def2_l space_qbs_l_in[OF assms(1)] assms(2) intro!: Bochner_Integration.integral_nonneg) lemma qbs_integral_mono_AEq: fixes f :: "_ \ real" assumes "qbs_integrable s f" "qbs_integrable s g" "AE\<^sub>Q x in s. f x \ g x" shows "qbs_integral s f \ qbs_integral s g" using assms by(auto simp: qbs_integral_def2_l AEq_qbs_l qbs_integrable_iff_integrable intro!: integral_mono_AE) lemma qbs_integral_mono: fixes f :: "_ \ real" assumes "s \ qbs_space (monadM_qbs X)" and "qbs_integrable s f" "qbs_integrable s g" "\x. x \ qbs_space X \ f x \ g x" shows "qbs_integral s f \ qbs_integral s g" by(auto simp: qbs_integral_def2_l space_qbs_l_in[OF assms(1)] qbs_integrable_iff_integrable[symmetric] assms intro!: integral_mono) lemma qbs_integral_const_prob: assumes "s \ qbs_space (monadP_qbs X)" shows "(\\<^sub>Q x. c \s) = c" using assms by(simp add: qbs_integral_def2_l monadP_qbs_space prob_space.prob_space) lemma assumes "qbs_integrable s f" "qbs_integrable s g" shows qbs_integral_add:"(\\<^sub>Q x. f x + g x \s) = (\\<^sub>Q x. f x \s) + (\\<^sub>Q x. g x \s)" and qbs_integral_diff: "(\\<^sub>Q x. f x - g x \s) = (\\<^sub>Q x. f x \s) - (\\<^sub>Q x. g x \s)" using assms by(auto simp: qbs_integral_def2_l qbs_integrable_iff_integrable[symmetric] intro!: Bochner_Integration.integral_add Bochner_Integration.integral_diff) lemma [simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows qbs_integral_mult_right_zero:"(\\<^sub>Q x. c * f x \s) = c * (\\<^sub>Q x. f x \s)" and qbs_integral_mult_left_zero:"(\\<^sub>Q x. f x * c \s) = (\\<^sub>Q x. f x \s) * c" and qbs_integral_divide_zero: "(\\<^sub>Q x. f x / c \s) = (\\<^sub>Q x. f x \s) / c" by(auto simp: qbs_integral_def2_l) lemma qbs_integral_minus[simp]: "(\\<^sub>Q x. - f x \s) = - (\\<^sub>Q x. f x \s)" by(auto simp: qbs_integral_def2_l) lemma [simp]: shows qbs_integral_scaleR_right:"(\\<^sub>Q x. c *\<^sub>R f x \s) = c *\<^sub>R (\\<^sub>Q x. f x \s)" and qbs_integral_scaleR_left: "(\\<^sub>Q x. f x *\<^sub>R c \s) = (\\<^sub>Q x. f x \s) *\<^sub>R c" by(auto simp: qbs_integral_def2_l) lemma [simp]: shows qbs_integral_inner_left: "qbs_integrable s f \ (\\<^sub>Q x. f x \ c \s) = (\\<^sub>Q x. f x \s) \ c" and qbs_integral_inner_right: "qbs_integrable s f \ (\\<^sub>Q x. c \ f x \s) = c \ (\\<^sub>Q x. f x \s) " by(auto simp: qbs_integral_def2_l qbs_integrable_iff_integrable) lemma integral_complex_of_real[simp]: "(\\<^sub>Q x. complex_of_real (f x) \s)= of_real (\\<^sub>Q x. f x \s)" by(simp add: qbs_integral_def2_l) lemma integral_cnj[simp]: "(\\<^sub>Q x. cnj (f x) \s) = cnj (\\<^sub>Q x. f x \s)" by(simp add: qbs_integral_def2_l) lemma [simp]: assumes "qbs_integrable s f" shows qbs_integral_Im: "(\\<^sub>Q x. Im (f x) \s) = Im (\\<^sub>Q x. f x \s)" and qbs_integral_Re: "(\\<^sub>Q x. Re (f x) \s) = Re (\\<^sub>Q x. f x \s)" using assms by(auto simp: qbs_integral_def2_l qbs_integrable_iff_integrable) lemma qbs_integral_of_real[simp]:"qbs_integrable s f \ (\\<^sub>Q x. of_real (f x) \s) = of_real (\\<^sub>Q x. f x \s)" by(auto simp: qbs_integral_def2_l qbs_integrable_iff_integrable) lemma [simp]: assumes "qbs_integrable s f" shows qbs_integral_fst: "(\\<^sub>Q x. fst (f x) \s) = fst (\\<^sub>Q x. f x \s)" and qbs_integral_snd: "(\\<^sub>Q x. snd (f x) \s) = snd (\\<^sub>Q x. f x \s)" using assms by(auto simp: qbs_integral_def2_l qbs_integrable_iff_integrable) lemma real_qbs_integral_def: assumes "qbs_integrable s f" shows "qbs_integral s f = enn2real (\\<^sup>+\<^sub>Q x. ennreal (f x) \s) - enn2real (\\<^sup>+\<^sub>Q x. ennreal (- f x) \s)" using qbs_integrable_morphism_dest[OF in_qbs_space_of assms] assms by(auto simp: qbs_integral_def2_l qbs_nn_integral_def2_l qbs_integrable_iff_integrable[symmetric] intro!: real_lebesgue_integral_def) lemma Markov_inequality_qbs_prob: "qbs_integrable s f \ AE\<^sub>Q x in s. 0 \ f x \ 0 < c \ \

(x in qbs_l s. c \ f x) \ (\\<^sub>Q x. f x \s) / c" by(auto simp: qbs_integral_def2_l AEq_qbs_l qbs_integrable_iff_integrable intro!: integral_Markov_inequality_measure[where A="{}"]) lemma Chebyshev_inequality_qbs_prob: assumes "s \ qbs_space (monadP_qbs X)" and "f \ X \\<^sub>Q qbs_borel" "qbs_integrable s (\x. (f x)\<^sup>2)" and "0 < e" shows "\

(x in qbs_l s. e \ \f x - (\\<^sub>Q x. f x \s)\) \ (\\<^sub>Q x. (f x - (\\<^sub>Q x. f x \s))\<^sup>2 \s) / e\<^sup>2" using prob_space.Chebyshev_inequality[OF qbs_l_prob_space[OF assms(1)] _ _ assms(4),of f] assms(2,3) by(simp add: qbs_integral_def2_l qbs_integrable_iff_integrable lr_adjunction_correspondence measurable_qbs_l'[OF qbs_space_monadPM[OF assms(1)]]) lemma qbs_l_return_qbs: assumes "x \ qbs_space X" shows "qbs_l (return_qbs X x) = return (qbs_to_measure X) x" proof - interpret qp: qbs_prob X "\r. x" "return borel 0" by(auto simp: qbs_prob_def prob_space_return assms in_Mx_def real_distribution_def real_distribution_axioms_def) show ?thesis by(simp add: qp.return_qbs[OF assms] distr_return qp.qbs_l) qed lemma qbs_l_bind_qbs: assumes [qbs]: "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q monadM_qbs Y" shows "qbs_l (s \ f) = qbs_l s \\<^sub>k qbs_l \ f" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qs: qbs_s_finite X \ \ by simp from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(2) qs.in_Mx]] obtain \ k where hk: "f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" "\r. qbs_s_finite Y \ (k r )" by metis then interpret sk: s_finite_kernel borel borel k by simp interpret im: in_Mx Y \ using hk(2) by(simp add: in_Mx_def) have "?lhs = distr (\ \\<^sub>k k) (qbs_to_measure Y) \" by(simp add: qs.bind_qbs[OF hs(1) assms(2) hk(2,3,1)] qbs_s_finite.qbs_l[OF qs.bind_qbs_s_finite[OF hs(1) assms(2) hk(2,3,1)]]) also have "... = \ \\<^sub>k (\x. distr (k x) (qbs_to_measure Y) \)" by(auto intro!: sk.distr_bind_kernel simp: qs.mu_sets) also have "... = \ \\<^sub>k (\r. qbs_l ((f \ \) r))" by(simp add: qbs_s_finite.qbs_l[OF hk(4)] hk(1)) also have "... = \ \\<^sub>k (\r. (\x. qbs_l (f x)) (\ r))" by simp also have "... = distr \ (qbs_to_measure X) \ \\<^sub>k (\x. qbs_l (f x))" using l_preserves_morphisms[of X "monadM_qbs Y"] assms(2) by(auto intro!: measure_kernel.bind_kernel_distr[OF measure_kernel.measure_kernel_comp[OF qbs_l_measure_kernel],symmetric] simp: sets_eq_imp_space_eq[OF qs.mu_sets]) also have "... = ?rhs" by(simp add: hs(1) qs.qbs_l comp_def) finally show ?thesis . qed +lemma qbs_l_bind_qbsP: + assumes [qbs]: "s \ qbs_space (monadP_qbs X)" "f \ X \\<^sub>Q monadP_qbs Y" + shows "qbs_l (s \ f) = qbs_l s \ qbs_l \ f" +proof - + have "qbs_l (s \ f) = qbs_l s \\<^sub>k qbs_l \ f" + by(auto intro!: qbs_l_bind_qbs[where X=X and Y=Y] qbs_space_monadPM qbs_morphism_monadPD) + also have "... = qbs_l s \ qbs_l \ f" + using qbs_l_measurable_prob qbs_morphism_imp_measurable[OF assms(2)] + by(auto intro!: bind_kernel_bind[where N="qbs_to_measure Y"] measurable_prob_algebraD simp: measurable_qbs_l'[OF qbs_space_monadPM[OF assms(1)]]) + finally show ?thesis . +qed + lemma qbs_integrable_return[simp, intro]: assumes "x \ qbs_space X" "f \ X \\<^sub>Q qbs_borel" shows "qbs_integrable (return_qbs X x) f" using assms by(auto simp: qbs_integrable_iff_integrable qbs_l_return_qbs[OF assms(1)] lr_adjunction_correspondence nn_integral_return space_L intro!: integrableI_bounded) lemma qbs_integrable_bind_return: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ Y \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q Y" shows "qbs_integrable (s \ (\x. return_qbs Y (g x))) f = qbs_integrable s (f \ g)" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qs: qbs_s_finite X \ \ by simp have 1:"return_qbs Y \ (g \ \) = (\r. \Y, g \ \, return borel r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" by(auto intro!: return_qbs_comp qbs_morphism_Mx[OF assms(3)]) have hb: "qbs_s_finite Y (g \ \) \" "s \ (\x. return_qbs Y (g x)) = \Y, g \ \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qbs_s_finite.bind_qbs[OF hs(2,1) qbs_morphism_comp[OF assms(3) return_qbs_morphism] qbs_morphism_Mx[OF assms(3)] prob_kernel.s_finite_kernel_prob_kernel 1[simplified comp_assoc[symmetric]]] qbs_s_finite.bind_qbs_s_finite[OF hs(2,1) qbs_morphism_comp[OF assms(3) return_qbs_morphism] qbs_morphism_Mx[OF assms(3)] prob_kernel.s_finite_kernel_prob_kernel 1[simplified comp_assoc[symmetric]]] by(auto simp: prob_kernel_def' comp_def bind_kernel_return''[OF qs.mu_sets]) have "?lhs = integrable \ (f \ (g \ \))" by(auto simp: hb(2) intro!: qbs_s_finite.qbs_integrable_iff_integrable[OF hb(1),simplified comp_def] simp: comp_def lr_adjunction_correspondence[symmetric]) also have "... = ?rhs" by(auto simp: hs(1) comp_def lr_adjunction_correspondence[symmetric] intro!: qs.qbs_integrable_iff_integrable[symmetric]) finally show ?thesis . qed lemma qbs_nn_integral_morphism[qbs]: "qbs_nn_integral \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q qbs_borel) \\<^sub>Q qbs_borel" proof(rule curry_preserves_morphisms[OF pair_qbs_morphismI]) fix \ \ assume h:"\ \ qbs_Mx (monadM_qbs X)" "\ \ qbs_Mx (X \\<^sub>Q (qbs_borel :: ennreal quasi_borel))" from rep_qbs_Mx_monadM[OF h(1)] obtain a k where ak: "\ = (\r. \X, a, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "a \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X a (k r)" by metis have 1:"borel_measurable ((borel :: real measure) \\<^sub>M (borel :: real measure)) = qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q (qbs_borel :: ennreal quasi_borel)" by (metis borel_prod qbs_borel_prod standard_borel.standard_borel_r_full_faithful standard_borel_ne_borel standard_borel_ne_def) show "(\r. qbs_nn_integral (fst (\ r, \ r)) (snd (\ r, \ r))) \ qbs_Mx qbs_borel" unfolding qbs_Mx_qbs_borel by(rule measurable_cong[where f="\r. \\<^sup>+ x. \ r (a x) \k r",THEN iffD1], insert h ak(2)) (auto simp: qbs_s_finite.qbs_nn_integral_def[OF ak(4)] qbs_Mx_is_morphisms ak(1) 1 intro!: s_finite_kernel.nn_integral_measurable_f[OF ak(3)]) qed lemma qbs_nn_integral_return: assumes "f \ X \\<^sub>Q qbs_borel" and "x \ qbs_space X" shows "qbs_nn_integral (return_qbs X x) f = f x" using assms by(auto intro!: nn_integral_return simp: qbs_nn_integral_def2_l qbs_l_return_qbs space_L lr_adjunction_correspondence) lemma qbs_nn_integral_bind: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q monadM_qbs Y" "g \ Y \\<^sub>Q qbs_borel" shows "qbs_nn_integral (s \ f) g = qbs_nn_integral s (\y. (qbs_nn_integral (f y) g))" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qs: qbs_s_finite X \ \ by simp from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(2) qs.in_Mx]] obtain \ k where hk: "f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" "\r. qbs_s_finite Y \ (k r)" by metis note sf = qs.bind_qbs[OF hs(1) assms(2) hk(2,3,1)] qs.bind_qbs_s_finite[OF hs(1) assms(2) hk(2,3,1)] have "?lhs = (\\<^sup>+ x. g (\ x) \(\ \\<^sub>k k))" by(simp add: sf(1) qbs_s_finite.qbs_nn_integral_def[OF sf(2)]) also have "... = (\\<^sup>+ r. (\\<^sup>+ y. g (\ y) \(k r)) \\)" using assms(3) hk(2) by(auto intro!: s_finite_kernel.nn_integral_bind_kernel[OF hk(3)] qs.mu_sets simp: s_finite_kernel_cong_sets[OF qs.mu_sets] lr_adjunction_correspondence) also have "... = ?rhs" using fun_cong[OF hk(1)] by(auto simp: hs(1) qs.qbs_nn_integral_def qbs_s_finite.qbs_nn_integral_def[OF hk(4),symmetric] intro!: nn_integral_cong) finally show ?thesis . qed lemma qbs_nn_integral_bind_return: assumes [qbs]:"s \ qbs_space (monadM_qbs Y)" "f \ Z \\<^sub>Q qbs_borel" "g \ Y \\<^sub>Q Z" shows "qbs_nn_integral (s \ (\y. return_qbs Z (g y))) f = qbs_nn_integral s (f \ g)" by(auto simp: qbs_nn_integral_bind[OF assms(1) _ assms(2)] qbs_nn_integral_return intro!: qbs_nn_integral_cong[OF assms(1)]) lemma qbs_integral_morphism[qbs]: "qbs_integral \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q qbs_borel) \\<^sub>Q (qbs_borel :: ('b :: {second_countable_topology,banach}) quasi_borel)" proof(rule curry_preserves_morphisms[OF pair_qbs_morphismI]) fix \ and \ :: "_ \ _ \ 'b" assume h:"\ \ qbs_Mx (monadM_qbs X)" "\ \ qbs_Mx (X \\<^sub>Q qbs_borel)" from rep_qbs_Mx_monadM[OF this(1)] obtain \ k where hk: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis have 1:"borel_measurable ((borel :: real measure) \\<^sub>M (borel :: real measure)) = qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q (qbs_borel :: (_ :: {second_countable_topology,banach}) quasi_borel)" by (metis borel_prod qbs_borel_prod standard_borel.standard_borel_r_full_faithful standard_borel_ne_borel standard_borel_ne_def) show "(\r. qbs_integral (fst (\ r,\ r)) (snd (\ r,\ r))) \ qbs_Mx borel\<^sub>Q" unfolding qbs_Mx_R by(rule measurable_cong[where f="\r. \ x. \ r (\ x) \k r",THEN iffD1], insert h hk(2)) (auto simp: qbs_s_finite.qbs_integral_def[OF hk(4)] qbs_Mx_is_morphisms hk(1) 1 intro!: s_finite_kernel.integral_measurable_f[OF hk(3)]) qed lemma qbs_integral_return: assumes [qbs]:"f \ X \\<^sub>Q qbs_borel" "x \ qbs_space X" shows "qbs_integral (return_qbs X x) f = f x" by(auto simp: qbs_integral_def2_l qbs_l_return_qbs lr_adjunction_correspondence[symmetric] space_L integral_return) lemma assumes [qbs]: "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q monadM_qbs Y" "g \ Y \\<^sub>Q qbs_borel" and "qbs_integrable s (\x. \\<^sub>Q y. norm (g y) \f x)" "AE\<^sub>Q x in s. qbs_integrable (f x) g" shows qbs_integrable_bind: "qbs_integrable (s \ f) g" (is ?goal1) and qbs_integral_bind:"(\\<^sub>Q y. g y \(s \ f)) = (\\<^sub>Q x. \\<^sub>Q y. g y \f x \s)" (is "?lhs = ?rhs") proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qs: qbs_s_finite X \ \ by simp from rep_qbs_Mx_monadM[OF qbs_morphism_Mx[OF assms(2) qs.in_Mx]] obtain \ k where hk: "f \ \ = (\r. \Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx Y" "s_finite_kernel borel borel k" "\r. qbs_s_finite Y \ (k r)" by metis note sf = qs.bind_qbs[OF hs(1) assms(2) hk(2,3,1)] have g[measurable]: "\h M. h \ M \\<^sub>M qbs_to_measure Y \ (\x. g (h x)) \ M \\<^sub>M borel" using assms(3) by(auto simp: lr_adjunction_correspondence) interpret qs2: qbs_s_finite Y \ "\ \\<^sub>k k" by(rule qs.bind_qbs_s_finite[OF hs(1) assms(2) hk(2,3,1)]) show ?goal1 by(auto simp: sf qs2.qbs_integrable_def intro!: s_finite_kernel.integrable_bind_kernel[OF hk(3) qs.mu_sets]) (insert qs.AEq_AE[OF assms(5)[simplified hs(1)],simplified fun_cong[OF hk(1),simplified] qbs_s_finite.qbs_integrable_def[OF hk(4)]] assms(4)[simplified hs(1) qs.qbs_integrable_def fun_cong[OF hk(1),simplified]],auto simp: hs(1) qs.qbs_integrable_def qbs_s_finite.qbs_integral_def[OF hk(4)]) have "?lhs = (\r. g (\ r) \(\ \\<^sub>k k))" by(simp add: sf qs2.qbs_integral_def) also have "... = (\r. (\l. g (\ l)\k r) \\)" using qs.AEq_AE[OF assms(5)[simplified hs(1)],simplified fun_cong[OF hk(1),simplified] qbs_s_finite.qbs_integrable_def[OF hk(4)]] assms(4)[simplified hs(1) qs.qbs_integrable_def fun_cong[OF hk(1),simplified]] by(auto intro!: s_finite_kernel.integral_bind_kernel[OF hk(3) qs.mu_sets] simp: qbs_s_finite.qbs_integral_def[OF hk(4)]) also have "... = (\r. (\\<^sub>Q y. g y\\Y, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \\)" by(auto intro!: Bochner_Integration.integral_cong simp: qbs_s_finite.qbs_integral_def[OF hk(4)]) also have "... = ?rhs" by(auto simp: hs(1) qs.qbs_integral_def fun_cong[OF hk(1),simplified comp_def]) finally show "?lhs = ?rhs" . qed lemma qbs_integral_bind_return: assumes [qbs]:"s \ qbs_space (monadM_qbs Y)" "f \ Z \\<^sub>Q qbs_borel" "g \ Y \\<^sub>Q Z" shows "qbs_integral (s \ (\y. return_qbs Z (g y))) f = qbs_integral s (f \ g)" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" by metis then interpret qs: qbs_s_finite Y \ \ by simp have hb: "qbs_s_finite Z (g \ \) \" "s \ (\y. return_qbs Z (g y)) = \Z, g \ \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" using qs.bind_qbs_s_finite[OF hs(1) _ qbs_morphism_Mx[OF assms(3) qs.in_Mx] prob_kernel.s_finite_kernel_prob_kernel return_qbs_comp[OF qbs_morphism_Mx[OF assms(3) qs.in_Mx],simplified comp_assoc[symmetric] comp_def[of _ g]],simplified prob_kernel_def'] by(auto simp: qs.bind_qbs[OF hs(1) _ qbs_morphism_Mx[OF assms(3) qs.in_Mx] prob_kernel.s_finite_kernel_prob_kernel return_qbs_comp[OF qbs_morphism_Mx[OF assms(3) qs.in_Mx],simplified comp_assoc[symmetric] comp_def[of _ g]],simplified prob_kernel_def'] bind_kernel_return''[OF qs.mu_sets]) show ?thesis by(simp add: hb(2) qbs_s_finite.qbs_integral_def[OF hb(1)] qs.qbs_integral_def[simplified hs(1)[symmetric]]) qed subsubsection \ Binary Product Measures\ definition qbs_pair_measure :: "['a qbs_measure, 'b qbs_measure] \ ('a \ 'b) qbs_measure" (infix "\\<^sub>Q\<^sub>m\<^sub>e\<^sub>s" 80) where qbs_pair_measure_def':"qbs_pair_measure p q \ (p \ (\x. q \ (\y. return_qbs (qbs_space_of p \\<^sub>Q qbs_space_of q) (x, y))))" context pair_qbs_s_finites begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) lemma shows qbs_pair_measure: "\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = \X \\<^sub>Q Y, map_prod \ \ \ rr.from_real, distr (\ \\<^sub>M \) borel rr.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and qbs_pair_measure_s_finite: "qbs_s_finite (X \\<^sub>Q Y) (map_prod \ \ \ rr.from_real) (distr (\ \\<^sub>M \) borel rr.to_real)" by(simp_all add: qbs_pair_measure_def' pq1.qbs_l pq2.qbs_l qbs_bind_bind_return_pq qbs_s_finite_axioms) lemma qbs_l_qbs_pair_measure: "qbs_l (\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) = distr (\ \\<^sub>M \) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \)" by(simp add: qbs_pair_measure qbs_s_finite.qbs_l[OF qbs_pair_measure_s_finite] distr_distr comp_assoc) lemma qbs_nn_integral_pair_measure: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q z. f z \(\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)) = (\\<^sup>+ z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" using assms by(simp add: qbs_nn_integral_def2 qbs_pair_measure distr_distr comp_assoc nn_integral_distr lr_adjunction_correspondence) lemma qbs_integral_pair_measure: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" shows "(\\<^sub>Q z. f z \(\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)) = (\ z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" using assms by(simp add: qbs_integral_def2 qbs_pair_measure distr_distr comp_assoc integral_distr lr_adjunction_correspondence) lemma qbs_pair_measure_integrable_eq: "qbs_integrable (\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) f \ f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel \ integrable (\ \\<^sub>M \) (f \ (map_prod \ \))" (is "?h \ ?h1 \ ?h2") proof safe assume h: ?h show ?h1 by(auto intro!: qbs_integrable_morphism_dest[OF _ h] simp: qbs_pair_measure_def') have 1:"integrable (distr (\ \\<^sub>M \) borel (to_real_on (borel \\<^sub>M borel))) (f \ (map_prod \ \ \ from_real_into (borel \\<^sub>M borel)))" using h[simplified qbs_pair_measure] by(simp add: qbs_integrable_def[of f] comp_def[of f]) have "integrable (\ \\<^sub>M \) (\x. (f \ (map_prod \ \ \ from_real_into (borel \\<^sub>M borel))) (to_real_on (borel \\<^sub>M borel) x))" by(intro integrable_distr[OF _ 1]) simp thus ?h2 by(simp add: comp_def) next assume h: ?h1 ?h2 then show ?h by(simp add: qbs_pair_measure qbs_integrable_def) (simp add: lr_adjunction_correspondence integrable_distr_eq[of rr.to_real "\ \\<^sub>M \" borel "\x. f (map_prod \ \ (rr.from_real x))"] comp_def) qed end lemmas(in pair_qbs_probs) qbs_pair_measure_prob = qbs_prob_axioms context fixes X Y p q assumes p[qbs]:"p \ qbs_space (monadM_qbs X)" and q[qbs]:"q \ qbs_space (monadM_qbs Y)" begin lemma qbs_pair_measure_def: "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q = p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))" by(simp add: qbs_space_of_in[OF p] qbs_space_of_in[OF q] qbs_pair_measure_def') lemma qbs_pair_measure_def2: "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q = q \ (\y. p \ (\x. return_qbs (X \\<^sub>Q Y) (x,y)))" by(simp add: bind_qbs_return_rotate qbs_pair_measure_def) lemma assumes "f \ X \\<^sub>Q Y \\<^sub>Q monadM_qbs Z" shows qbs_pair_bind_bind_return1':"q \ (\y. p \ (\x. f (x,y))) = p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q \ f" and qbs_pair_bind_bind_return2':"p \ (\x. q \ (\y. f (x,y))) = p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q \ f" by(simp_all add: qbs_bind_bind_return1[OF assms] qbs_bind_bind_return2[OF assms] bind_qbs_return_rotate qbs_pair_measure_def) lemma assumes [qbs]:"f \ X \\<^sub>Q exp_qbs Y (monadM_qbs Z)" shows qbs_pair_bind_bind_return1'': "q \ (\y. p \ (\x. f x y)) = p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q \ (\x. f (fst x) (snd x))" and qbs_pair_bind_bind_return2'': "p \ (\x. q \ (\y. f x y)) = p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q \ (\x. f (fst x) (snd x))" by(auto intro!: qbs_pair_bind_bind_return1'[where f="\x. f (fst x) (snd x)",simplified] qbs_pair_bind_bind_return2'[where f="\x. f (fst x) (snd x)",simplified] uncurry_preserves_morphisms) qbs lemma qbs_nn_integral_Fubini_fst: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. f (x,y) \q \p) = (\\<^sup>+\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. qbs_nn_integral (return_qbs (X \\<^sub>Q Y) (x, y)) f \q \p)" by(auto intro!: qbs_nn_integral_cong p q simp: qbs_nn_integral_return) also have "... = ?rhs" by(auto intro!: qbs_nn_integral_cong[OF p] simp:qbs_nn_integral_bind[OF q _ assms] qbs_nn_integral_bind[OF p _ assms] qbs_pair_measure_def) finally show ?thesis . qed lemma qbs_nn_integral_Fubini_snd: assumes [qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q y. \\<^sup>+\<^sub>Q x. f (x,y) \p \q) = (\\<^sup>+\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+\<^sub>Q y. \\<^sup>+\<^sub>Q x. qbs_nn_integral (return_qbs (X \\<^sub>Q Y) (x, y)) f \p \q)" by(auto intro!: qbs_nn_integral_cong p q simp: qbs_nn_integral_return) also have "... = ?rhs" by(auto intro!: qbs_nn_integral_cong[OF q] simp:qbs_nn_integral_bind[OF q _ assms] qbs_nn_integral_bind[OF p _ assms] qbs_pair_measure_def2) finally show ?thesis . qed lemma qbs_ennintegral_indep_mult: assumes [qbs]: "f \ X \\<^sub>Q qbs_borel" "g \ Y \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q z. f (fst z) * g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q x. f x \p) * (\\<^sup>+\<^sub>Q y. g y \q)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y .f x * g y \q \p)" using qbs_nn_integral_Fubini_fst[where f="\z. f (fst z) * g (snd z)"] by simp also have "... = (\\<^sup>+\<^sub>Q x. f x * \\<^sup>+\<^sub>Q y . g y \q \p)" by(simp add: qbs_nn_integral_cmult[OF q]) also have "... = ?rhs" by(simp add: qbs_nn_integral_cmult[OF p] ab_semigroup_mult_class.mult.commute[where b="qbs_nn_integral q g"]) finally show ?thesis . qed end lemma qbs_l_qbs_pair_measure: assumes "standard_borel M" "standard_borel N" defines "X \ measure_to_qbs M" and "Y \ measure_to_qbs N" assumes [qbs]: "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" shows "qbs_l (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) = qbs_l p \\<^sub>M qbs_l q" proof - obtain \ \ \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" and hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" using rep_qbs_space_monadM assms(5,6) by meson have 1:"sets (qbs_to_measure (X \\<^sub>Q Y)) = sets (M \\<^sub>M N)" by(auto simp: r_preserves_product[symmetric] X_def Y_def intro!: standard_borel.lr_sets_ident pair_standard_borel assms) have "qbs_l (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) = qbs_l p \\<^sub>k qbs_l \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x,y)))" by(auto simp: qbs_pair_measure_def[of p X q Y] intro!: qbs_l_bind_qbs[of _ X _ "X \\<^sub>Q Y"]) also have "... = qbs_l p \\<^sub>k (\x. qbs_l (q \ (\y. return_qbs (X \\<^sub>Q Y) (x, y))))" by(simp add: comp_def) also have "... = qbs_l p \\<^sub>k (\x. qbs_l q \\<^sub>k qbs_l \ (\y. return_qbs (X \\<^sub>Q Y) (x, y)))" by(auto intro!: bind_kernel_cong_All qbs_l_bind_qbs[of _ "Y" _ "X \\<^sub>Q Y"] simp: space_qbs_l_in[OF assms(5)]) also have "... = qbs_l p \\<^sub>k (\x. qbs_l q \\<^sub>k (\y. return (qbs_to_measure (X \\<^sub>Q Y)) (x, y)))" by(auto simp: comp_def space_qbs_l_in[OF assms(6)] space_qbs_l_in[OF assms(5)] qbs_l_return_qbs intro!: bind_kernel_cong_All) also have "... = qbs_l p \\<^sub>k (\x. qbs_l q \\<^sub>k (\y. return (M \\<^sub>M N) (x, y)))" by(simp add: return_cong[OF 1]) also have "... = qbs_l p \\<^sub>k (\x. qbs_l q \\<^sub>k (\y. return (qbs_l p \\<^sub>M qbs_l q) (x, y)))" by(auto cong: return_cong sets_pair_measure_cong simp: sets_qbs_l[OF assms(5)] standard_borel.lr_sets_ident[OF assms(1)] sets_qbs_l[OF assms(6)] standard_borel.lr_sets_ident[OF assms(2)] X_def Y_def) also have "... = qbs_l p \\<^sub>M qbs_l q" by(auto intro!: pair_measure_eq_bind_s_finite[symmetric] qbs_l_s_finite.s_finite_measure_axioms) finally show ?thesis . qed lemma qbs_pair_measure_morphism[qbs]: "qbs_pair_measure \ monadM_qbs X \\<^sub>Q monadM_qbs Y \\<^sub>Q monadM_qbs (X \\<^sub>Q Y)" by(rule curry_preserves_morphisms,rule qbs_morphism_cong'[where f="(\(p,q). (p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x, y)))))"]) (auto simp: pair_qbs_space qbs_pair_measure_def) lemma qbs_pair_measure_morphismP: "qbs_pair_measure \ monadP_qbs X \\<^sub>Q monadP_qbs Y \\<^sub>Q monadP_qbs (X \\<^sub>Q Y)" proof - note [qbs] = return_qbs_morphismP bind_qbs_morphismP show ?thesis by(rule curry_preserves_morphisms,rule qbs_morphism_cong'[where f="(\(p,q). (p \ (\x. q \ (\y. return_qbs (X \\<^sub>Q Y) (x, y)))))"]) (auto simp: pair_qbs_space qbs_pair_measure_def[OF qbs_space_monadPM qbs_space_monadPM]) qed lemma qbs_nn_integral_indep1: assumes [qbs]:"p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadP_qbs X)" "f \ X \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q z. f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q x. f x \p)" proof - obtain Y \ \ where hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_prob Y \ \" using rep_qbs_space_monadP[OF assms(2)] by blast then interpret qbs_prob Y \ \ by simp show ?thesis by(simp add: qbs_nn_integral_const_prob[OF in_space_monadP] qbs_nn_integral_Fubini_snd[OF assms(1) in_space_monadM,symmetric] hq(1)) qed lemma qbs_nn_integral_indep2: assumes [qbs]:"q \ qbs_space (monadM_qbs Y)" "p \ qbs_space (monadP_qbs X)" "f \ Y \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q z. f (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q y. f y \q)" proof - obtain X \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_prob X \ \" using rep_qbs_space_monadP[OF assms(2)] by metis then interpret qbs_prob X \ \ by simp show ?thesis by(simp add: qbs_nn_integral_const_prob[OF in_space_monadP] qbs_nn_integral_Fubini_snd[OF in_space_monadM assms(1),symmetric] hp(1)) qed - context begin interpretation rr : standard_borel_ne "borel \\<^sub>M borel :: (real \ real) measure" by(auto intro!: pair_standard_borel_ne) lemma qbs_integrable_pair_swap: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "qbs_integrable (q \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s p) (\(x,y). f (y,x))" proof - obtain X \ \ Y \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" and hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" using rep_qbs_s_finite_measure by meson interpret p1: pair_qbs_s_finites X \ \ Y \ \ by(simp add: pair_qbs_s_finites_def hq hp) interpret p2: pair_qbs_s_finites Y \ \ X \ \ by(simp add: pair_qbs_s_finites_def hq hp) show ?thesis using assms by(auto simp: hp(1) hq(1) p1.qbs_pair_measure p2.qbs_pair_measure p1.qbs_integrable_def p2.qbs_integrable_def) (auto simp add: integrable_distr_eq lr_adjunction_correspondence qbs_Mx_is_morphisms map_prod_def split_beta' intro!:integrable_product_swap_iff_s_finite[OF p1.pq2.s_finite_measure_axioms p1.pq1.s_finite_measure_axioms,THEN iffD1]) qed lemma qbs_integrable_pair1': assumes [qbs]:"p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" "f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" "qbs_integrable p (\x. \\<^sub>Q y. norm (f (x,y)) \q)" and "AE\<^sub>Q x in p. qbs_integrable q (\y. f (x,y))" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" proof - obtain \ \ \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" and hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" using rep_qbs_space_monadM assms(1,2) by meson then interpret pqs: pair_qbs_s_finites X \ \ Y \ \ by(simp add: pair_qbs_s_finites_def) have [measurable]: "f \ borel_measurable (qbs_to_measure (X \\<^sub>Q Y))" by(simp add: lr_adjunction_correspondence[symmetric]) show ?thesis using assms(4) pqs.pq1.AEq_AE[OF assms(5)[simplified hp(1)]] by(auto simp add: pqs.qbs_integrable_def pqs.qbs_pair_measure hp(1) hq(1) integrable_distr_eq pqs.pq2.qbs_integrable_def pqs.pq1.qbs_integrable_def pqs.pq2.qbs_integral_def intro!: s_finite_measure.Fubini_integrable' pqs.pq2.s_finite_measure_axioms) qed lemma assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows qbs_integrable_pair1D1': "qbs_integrable p (\x. \\<^sub>Q y. f (x,y) \q)" (is ?g1) and qbs_integrable_pair1D1_norm': "qbs_integrable p (\x. \\<^sub>Q y. norm (f (x,y)) \q)" (is ?g2) and qbs_integrable_pair1D2': "AE\<^sub>Q x in p. qbs_integrable q (\y. f (x,y))" (is ?g3) and qbs_integrable_pair2D1': "qbs_integrable q (\y. \\<^sub>Q x. f (x,y) \p)" (is ?g4) and qbs_integrable_pair2D1_norm': "qbs_integrable q (\y. \\<^sub>Q x. norm (f (x,y)) \p)" (is ?g5) and qbs_integrable_pair2D2': "AE\<^sub>Q y in q. qbs_integrable p (\x. f (x,y))" (is ?g6) and qbs_integral_Fubini_fst': "(\\<^sub>Q x. \\<^sub>Q y. f (x,y) \q \p) = (\\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is ?g7) and qbs_integral_Fubini_snd': "(\\<^sub>Q y. \\<^sub>Q x. f (x,y) \p \q) = (\\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is ?g8) proof - obtain X \ \ Y \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" and hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" by (meson rep_qbs_space_of) then interpret pqs: pair_qbs_s_finites X \ \ Y \ \ by(simp add: pair_qbs_s_finites_def) have [qbs]:"p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" by(simp_all add: hp hq) note qbs_pair_measure_morphism[qbs] have f[qbs]:"f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" by(auto intro!: qbs_integrable_morphism_dest[OF _ assms]) have [measurable]: "f \ borel_measurable (qbs_to_measure (X \\<^sub>Q Y))" by(simp add: lr_adjunction_correspondence[symmetric]) show ?g1 ?g2 ?g4 ?g5 using assms by(auto simp: hp(1) hq(1) pqs.qbs_integrable_def pqs.qbs_pair_measure integrable_distr_eq pqs.pq1.qbs_integrable_def pqs.pq2.qbs_integrable_def pqs.pq2.qbs_integral_def pqs.pq1.qbs_integral_def intro!: Bochner_Integration.integrable_cong[where g="\r. \\<^sub>Q y. f (\ r, y) \\Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and f="\x. \ y. f (\ x, \ y) \\" and N0=\,THEN iffD1] Bochner_Integration.integrable_cong[where g="\r. \\<^sub>Q x. f (x, \ r) \\X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and f="\y. \ x. f (\ x, \ y) \\" and N0=\,THEN iffD1]) (auto intro!: pqs.pq2.integrable_fst''[of \] integrable_snd_s_finite[OF pqs.pq1.s_finite_measure_axioms pqs.pq2.s_finite_measure_axioms] simp: map_prod_def split_beta') show ?g3 ?g6 using assms by(auto simp: hp(1) pqs.pq1.AEq_AE_iff hq(1) pqs.pq2.AEq_AE_iff pqs.qbs_integrable_def pqs.qbs_pair_measure integrable_distr_eq) (auto simp: pqs.pq1.qbs_integrable_def pqs.pq2.qbs_integrable_def map_prod_def split_beta' intro!: pqs.pq2.AE_integrable_fst'' AE_integrable_snd_s_finite[OF pqs.pq1.s_finite_measure_axioms pqs.pq2.s_finite_measure_axioms]) show ?g7 ?g8 using assms by(auto simp: hp(1) hq(1) pqs.qbs_integrable_def pqs.qbs_pair_measure pqs.qbs_integral_def pqs.pq1.qbs_integral_def pqs.pq2.qbs_integral_def integral_distr integrable_distr_eq) (auto simp: map_prod_def split_beta' intro!: pqs.pq2.integral_fst'''[of \ "\x. f (\ (fst x),\ (snd x))",simplified] integral_snd_s_finite[OF pqs.pq1.s_finite_measure_axioms pqs.pq2.s_finite_measure_axioms,of "\x y. f (\ x, \ y)",simplified split_beta']) qed end lemma assumes h:"qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (case_prod f)" shows qbs_integrable_pair1D1: "qbs_integrable p (\x. \\<^sub>Q y. f x y \q)" and qbs_integrable_pair1D1_norm: "qbs_integrable p (\x. \\<^sub>Q y. norm (f x y) \q)" and qbs_integrable_pair1D2: "AE\<^sub>Q x in p. qbs_integrable q (\y. f x y)" and qbs_integrable_pair2D1: "qbs_integrable q (\y. \\<^sub>Q x. f x y \p)" and qbs_integrable_pair2D1_norm: "qbs_integrable q (\y. \\<^sub>Q x. norm (f x y) \p)" and qbs_integrable_pair2D2: "AE\<^sub>Q y in q. qbs_integrable p (\x. f x y)" and qbs_integral_Fubini_fst: "(\\<^sub>Q x. \\<^sub>Q y. f x y \q \p) = (\\<^sub>Q (x,y). f x y \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is ?g7) and qbs_integral_Fubini_snd: "(\\<^sub>Q y. \\<^sub>Q x. f x y \p \q) = (\\<^sub>Q (x,y). f x y \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is ?g8) using qbs_integrable_pair1D1'[OF h] qbs_integrable_pair1D1_norm'[OF h] qbs_integrable_pair1D2'[OF h] qbs_integral_Fubini_fst'[OF h] qbs_integrable_pair2D1'[OF h] qbs_integrable_pair2D1_norm'[OF h] qbs_integrable_pair2D2'[OF h] qbs_integral_Fubini_snd'[OF h] by simp_all lemma qbs_integrable_pair2': assumes "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" "f \ X \\<^sub>Q Y \\<^sub>Q qbs_borel" "qbs_integrable q (\y. \\<^sub>Q x. norm (f (x,y)) \p)" and "AE\<^sub>Q y in q. qbs_integrable p (\x. f (x,y))" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" using qbs_integrable_pair_swap[OF qbs_integrable_pair1'[OF assms(2,1) qbs_morphism_pair_swap[OF assms(3)],simplified],OF assms(4,5)] by simp lemma qbs_integrable_indep_mult: fixes f :: "_ \ _::{real_normed_div_algebra,second_countable_topology}" assumes "qbs_integrable p f" "qbs_integrable q g" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. f (fst x) * g (snd x))" proof - obtain X \ \ Y \ \ where hp: "p = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" and hq: "q = \Y, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite Y \ \" by (meson rep_qbs_space_of) then interpret pqs: pair_qbs_s_finites X \ \ Y \ \ by(simp add: pair_qbs_s_finites_def) have [qbs]:"f \ X \\<^sub>Q qbs_borel" "g \ Y \\<^sub>Q qbs_borel" "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" by(auto intro!: qbs_integrable_morphism_dest assms simp:hp hq) show ?thesis by(auto intro!: qbs_integrable_pair1'[of _ X _ Y] qbs_integrable_mult_left qbs_integrable_norm assms(1) AEq_I2[of _ X] simp: norm_mult qbs_integrable_mult_right[OF assms(2)]) qed lemma qbs_integrable_indep1: fixes f :: "_ \ _::{real_normed_div_algebra,second_countable_topology}" assumes "qbs_integrable p f" "q \ qbs_space (monadP_qbs Y)" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. f (fst x))" using qbs_integrable_indep_mult[OF assms(1) qbs_integrable_const[OF assms(2),of 1]] by simp lemma qbs_integral_indep1: fixes f :: "_ \ _::{real_normed_div_algebra,second_countable_topology}" assumes "qbs_integrable p f" "q \ qbs_space (monadP_qbs Y)" shows "(\\<^sub>Q z. f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q x. f x \p)" using qbs_integral_Fubini_snd'[OF qbs_integrable_indep1[OF assms]] by(simp add: qbs_integral_const_prob[OF assms(2)]) lemma qbs_integrable_indep2: fixes g :: "_ \ _::{real_normed_div_algebra,second_countable_topology}" assumes "qbs_integrable q g" "p \ qbs_space (monadP_qbs X)" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. g (snd x))" using qbs_integrable_pair_swap[OF qbs_integrable_indep1[OF assms]] by(simp add: split_beta') lemma qbs_integral_indep2: fixes g :: "_ \ _::{real_normed_div_algebra,second_countable_topology}" assumes "qbs_integrable q g" "p \ qbs_space (monadP_qbs X)" shows "(\\<^sub>Q z. g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q y. g y \q)" using qbs_integral_Fubini_fst'[OF qbs_integrable_indep2[OF assms]] by(simp add: qbs_integral_const_prob[OF assms(2)]) lemma qbs_integral_indep_mult1: fixes f and g:: "_ \ _::{real_normed_field,second_countable_topology}" assumes "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" and "qbs_integrable p f" "qbs_integrable q g" shows "(\\<^sub>Q z. f (fst z) * g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q x. f x \p) * (\\<^sub>Q y. g y \q)" using qbs_integral_Fubini_fst'[OF qbs_integrable_indep_mult[OF assms(3,4)]] by simp lemma qbs_integral_indep_mult2: fixes f and g:: "_ \ _::{real_normed_field,second_countable_topology}" assumes "p \ qbs_space (monadP_qbs X)" "q \ qbs_space (monadP_qbs Y)" and "qbs_integrable p f" "qbs_integrable q g" shows "(\\<^sub>Q z. g (snd z) * f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q y. g y \q) * (\\<^sub>Q x. f x \p)" using qbs_integral_indep_mult1[OF assms] by(simp add: mult.commute) subsubsection \ The Inverse Function of $l$\ definition qbs_l_inverse :: "'a measure \ 'a qbs_measure" where "qbs_l_inverse M \ \measure_to_qbs M, from_real_into M, distr M borel (to_real_on M)\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" context standard_borel_ne begin lemma qbs_l_inverse_def2: assumes [measurable_cong]: "sets \ = sets M" and "s_finite_measure \" shows "qbs_l_inverse \ = \measure_to_qbs M, from_real, distr \ borel to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - interpret s: standard_borel_ne \ using assms standard_borel_ne_axioms standard_borel_ne_sets by blast have [measurable]: "s.from_real \ borel \\<^sub>M M" using assms(1) measurable_cong_sets s.from_real_measurable by blast show ?thesis by(auto simp: distr_distr qbs_l_inverse_def qbs_s_finite_eq_def qbs_s_finite_def in_Mx_def qbs_Mx_R qbs_s_finite_axioms_def intro!: qbs_s_finite_measure_eq s_finite_measure.s_finite_measure_distr assms cong: measure_to_qbs_cong_sets[OF assms(1)]) (auto intro!: distr_cong simp: sets_eq_imp_space_eq[OF assms(1)]) qed lemma assumes [measurable_cong]:"sets \ = sets M" shows qbs_l_inverse_s_finite: "s_finite_measure \ \ qbs_s_finite (measure_to_qbs M) from_real (distr \ borel to_real)" and qbs_l_inverse_qbs_prob: "prob_space \ \ qbs_prob (measure_to_qbs M) from_real (distr \ borel to_real)" by(auto simp: qbs_s_finite_def qbs_prob_def in_Mx_def qbs_s_finite_axioms_def real_distribution_def real_distribution_axioms_def qbs_Mx_R intro!: s_finite_measure.s_finite_measure_distr prob_space.prob_space_distr) corollary assumes [measurable_cong]:"sets \ = sets M" shows qbs_l_inverse_in_space_monadM: "s_finite_measure \ \ qbs_l_inverse \ \ qbs_space (monadM_qbs M)" and qbs_l_inverse_in_space_monadP: "prob_space \ \ qbs_l_inverse \ \ qbs_space (monadP_qbs M)" by(auto simp: qbs_l_inverse_def2[OF assms(1)] qbs_l_inverse_def2[OF assms(1) prob_space.s_finite_measure_prob] assms intro!: qbs_s_finite.in_space_monadM[OF qbs_l_inverse_s_finite] qbs_prob.in_space_monadP[OF qbs_l_inverse_qbs_prob]) lemma qbs_l_qbs_l_inverse: assumes [measurable_cong]: "sets \ = sets M" "s_finite_measure \" shows "qbs_l (qbs_l_inverse \) = \" proof - interpret qbs_s_finite "measure_to_qbs M" from_real "distr \ borel to_real" by(auto intro!: qbs_l_inverse_s_finite assms) show ?thesis using distr_id'[OF assms(1),simplified sets_eq_imp_space_eq[OF assms(1)]] by(auto simp: qbs_l qbs_l_inverse_def2[OF assms] distr_distr cong: distr_cong) qed corollary qbs_l_qbs_l_inverse_prob: "sets \ = sets M \ prob_space \ \ qbs_l (qbs_l_inverse \) = \" by(auto intro!: qbs_l_qbs_l_inverse prob_space.s_finite_measure_prob) lemma qbs_l_inverse_qbs_l: assumes "s \ qbs_space (monadM_qbs (measure_to_qbs M))" shows "qbs_l_inverse (qbs_l s) = s" proof - from rep_qbs_space_monadM[OF assms] obtain \ \ where h: "s = \measure_to_qbs M, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite (measure_to_qbs M) \ \" by metis then interpret qs:qbs_s_finite "measure_to_qbs M" \ \ by simp have [simp]: "distr \ (qbs_to_measure (measure_to_qbs M)) \ = distr \ M \" by(simp cong: distr_cong) interpret s: standard_borel_ne "distr \ M \" by(rule standard_borel_ne_sets[of M]) (auto simp: standard_borel_ne_axioms) have [measurable]: "s.from_real \ borel \\<^sub>M M" "\ \ \ \\<^sub>M M" using qs.\_measurable[simplified measurable_cong_sets[OF refl lr_sets_ident]] by(auto simp: s.from_real_measurable[simplified measurable_cong_sets[OF refl sets_distr]]) interpret pqs:pair_qbs_s_finite "measure_to_qbs M" s.from_real "distr \ borel (s.to_real \ \)" \ \ by(auto simp: pair_qbs_s_finite_def h) (auto simp: qbs_s_finite_def in_Mx_def qs.s_finite_measure_axioms qbs_s_finite_axioms_def qbs_Mx_R intro!: s_finite_measure.s_finite_measure_distr) show ?thesis by(auto simp add: h(1) qs.qbs_l qbs_l_inverse_def distr_distr cong: measure_to_qbs_cong_sets intro!: pqs.qbs_s_finite_measure_eq) (insert qbs_Mx_to_X[of _ "measure_to_qbs M"], auto simp: comp_def qbs_space_R) qed corollary qbs_l_inverse_qbs_l_prob: assumes "s \ qbs_space (monadP_qbs (measure_to_qbs M))" shows "qbs_l_inverse (qbs_l s) = s" by(auto intro!: qbs_l_inverse_qbs_l qbs_space_monadPM assms) lemma s_finite_kernel_qbs_morphism: assumes "s_finite_kernel N M k" shows "(\x. qbs_l_inverse (k x)) \ measure_to_qbs N \\<^sub>Q monadM_qbs (measure_to_qbs M)" proof - interpret sfin: s_finite_kernel N M k by fact have "\measure_to_qbs M, from_real, distr (k x) borel to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n = qbs_l_inverse (k x)" if x:"x \ space N" for x proof - note sfin.kernel_sets[OF x,simp, measurable_cong] then interpret skx: standard_borel_ne "k x" using standard_borel_ne_axioms standard_borel_ne_sets by blast interpret pqs: pair_qbs_s_finite "measure_to_qbs M" from_real "distr (k x) borel to_real" skx.from_real "distr (k x) borel skx.to_real" using skx.from_real_measurable[simplified measurable_cong_sets[OF refl sfin.kernel_sets[OF x]]] by(auto simp: pair_qbs_s_finite_def qbs_s_finite_def in_Mx_def qbs_Mx_R qbs_s_finite_axioms_def sfin.image_s_finite_measure[OF x] intro!: s_finite_measure.s_finite_measure_distr) show ?thesis by(auto simp: qbs_l_inverse_def distr_distr cong: measure_to_qbs_cong_sets intro!: pqs.qbs_s_finite_measure_eq) (auto intro!: distr_cong simp: sfin.kernel_space[OF x]) qed moreover have "(\x. \measure_to_qbs M, from_real, distr (k x) borel to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ measure_to_qbs N \\<^sub>Q monadM_qbs (measure_to_qbs M)" proof(rule qbs_morphismI) fix \ :: "real \ _" assume "\ \ qbs_Mx (measure_to_qbs N)" then have [measurable]: "\ \ borel \\<^sub>M N" by(simp add: qbs_Mx_R) show "(\x. \measure_to_qbs M, from_real, distr (k x) borel to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n) \ \ \ qbs_Mx (monadM_qbs (measure_to_qbs M))" by(auto simp: monadM_qbs_Mx qbs_Mx_R intro!: exI[where x=from_real] exI[where x="\x. distr (k (\ x)) borel to_real"] s_finite_kernel.comp_measurable[OF sfin.distr_s_finite_kernel]) qed ultimately show ?thesis by(rule qbs_morphism_cong'[of "measure_to_qbs N",simplified qbs_space_R]) qed lemma prob_kernel_qbs_morphism: assumes [measurable]:"k \ N \\<^sub>M prob_algebra M" shows "(\x. qbs_l_inverse (k x)) \ measure_to_qbs N \\<^sub>Q monadP_qbs (measure_to_qbs M)" proof(safe intro!: qbs_morphism_monadPI' s_finite_kernel_qbs_morphism prob_kernel.s_finite_kernel_prob_kernel) fix x assume "x \ qbs_space (measure_to_qbs N)" then have "x \ space N" by(simp add: qbs_space_R) from measurable_space[OF assms this] have [measurable_cong, simp]: "sets (k x) = sets M" and p:"prob_space (k x)" by(auto simp: space_prob_algebra) then interpret s: standard_borel_ne "k x" using standard_borel_ne_axioms standard_borel_ne_sets by blast show "qbs_l_inverse (k x) \ qbs_space (monadP_qbs (measure_to_qbs M))" using s.qbs_l_inverse_in_space_monadP[OF refl p] by(simp cong: measure_to_qbs_cong_sets) qed(simp add: prob_kernel_def') lemma qbs_l_inverse_return: assumes "x \ space M" shows "qbs_l_inverse (return M x) = return_qbs (measure_to_qbs M) x" proof - interpret s: standard_borel_ne "return M x" by(rule standard_borel_ne_sets[of M]) (auto simp: standard_borel_ne_axioms) show ?thesis using s.qbs_l_inverse_in_space_monadP[OF refl prob_space_return[OF assms]] by(auto intro!: inj_onD[OF qbs_l_inj_P[of "measure_to_qbs M"]] return_cong qbs_l_inverse_in_space_monadP qbs_morphism_space[OF return_qbs_morphismP[of "measure_to_qbs M"]] assms simp: s.qbs_l_qbs_l_inverse_prob[OF refl prob_space_return[OF assms]] qbs_l_return_qbs[of _ M,simplified qbs_space_R,OF assms] qbs_space_R cong: measure_to_qbs_cong_sets) qed lemma qbs_l_inverse_bind_kernel: assumes "standard_borel_ne N" "s_finite_measure M" "s_finite_kernel M N k" shows "qbs_l_inverse (M \\<^sub>k k) = qbs_l_inverse M \ (\x. qbs_l_inverse (k x))" (is "?lhs = ?rhs") proof - interpret sfin: s_finite_kernel M N k by fact interpret s: standard_borel_ne N by fact have sets[simp,measurable_cong]:"sets (M \\<^sub>k k) = sets N" by(auto intro!: sets_bind_kernel[OF _ space_ne] simp: sfin.kernel_sets) then interpret s2: standard_borel_ne "M \\<^sub>k k" using s.standard_borel_ne_axioms standard_borel_ne_sets by blast have [measurable]: "s2.from_real \ borel \\<^sub>M N" using measurable_cong_sets s2.from_real_measurable sets by blast have comp1:"(\x. qbs_l_inverse (k x)) \ from_real = (\r. \measure_to_qbs N, s.from_real, distr (k (from_real r)) borel s.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" proof fix r have setskfr[measurable_cong, simp]: "sets (k (from_real r)) = sets N" by(auto intro!: sfin.kernel_sets measurable_space[OF from_real_measurable]) then interpret s3: standard_borel_ne "k (from_real r)" using s.standard_borel_ne_axioms standard_borel_ne_sets by blast have [measurable]: "s3.from_real \ borel \\<^sub>M N" using measurable_cong_sets s3.from_real_measurable setskfr by blast show "((\x. qbs_l_inverse (k x)) \ from_real) r = \measure_to_qbs N, s.from_real, distr (k (from_real r)) borel s.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n " by(auto simp: qbs_l_inverse_def qbs_s_finite_eq_def qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def qbs_Mx_R distr_distr measurable_space[OF from_real_measurable] cong: measure_to_qbs_cong_sets intro!: sfin.image_s_finite_measure s_finite_measure.s_finite_measure_distr qbs_s_finite_measure_eq) (auto intro!: distr_cong simp: sets_eq_imp_space_eq[OF setskfr]) qed have "?lhs = \measure_to_qbs (M \\<^sub>k k), s2.from_real, distr (M \\<^sub>k k) borel s2.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(simp add: qbs_l_inverse_def) also have "... = \measure_to_qbs N, s.from_real, distr (M \\<^sub>k k) borel s.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(auto cong: measure_to_qbs_cong_sets intro!: qbs_s_finite_measure_eq distr_cong s_finite_measure.s_finite_measure_distr sfin.comp_s_finite_measure assms(2) simp: qbs_s_finite_eq_def qbs_s_finite_def qbs_s_finite_axioms_def in_Mx_def qbs_Mx_R distr_distr sets_eq_imp_space_eq[OF sets]) also have "... = \measure_to_qbs N, s.from_real, M \\<^sub>k (\x. distr (k x) borel s.to_real)\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" by(simp add: sfin.distr_bind_kernel[OF space_ne refl]) also have "... = \measure_to_qbs N, s.from_real, distr M borel to_real \\<^sub>k (\r. distr (k (from_real r)) borel s.to_real)\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" proof - have "M \\<^sub>k (\x. distr (k x) borel s.to_real) = M \\<^sub>k (\x. distr (k (from_real (to_real x))) borel s.to_real)" by(auto intro!: bind_kernel_cong_All) also have "... = distr M borel to_real \\<^sub>k (\r. distr (k (from_real r)) borel s.to_real)" by(auto intro!: measure_kernel.bind_kernel_distr[symmetric,where Y=borel] space_ne measure_kernel.distr_measure_kernel[where Y=N] sfin.measure_kernel_comp) finally show ?thesis by simp qed also have "... = ?rhs" by(auto intro!: qbs_s_finite.bind_qbs[OF qbs_l_inverse_s_finite[OF refl assms(2)] _ s.s_finite_kernel_qbs_morphism[OF assms(3)] _ _ comp1,symmetric] s_finite_kernel.distr_s_finite_kernel[OF sfin.comp_measurable] simp: qbs_Mx_R) (simp add: qbs_l_inverse_def) finally show ?thesis . qed lemma qbs_l_inverse_bind: assumes "standard_borel_ne N" "s_finite_measure M" "k \ M \\<^sub>M prob_algebra N" shows "qbs_l_inverse (M \ k) = qbs_l_inverse M \ (\x. qbs_l_inverse (k x))" by(auto simp: bind_kernel_bind[OF measurable_prob_algebraD[OF assms(3)],symmetric] prob_kernel_def' intro!: qbs_l_inverse_bind_kernel assms prob_kernel.s_finite_kernel_prob_kernel) end subsubsection \ PMF and SPMF \ definition "qbs_pmf \ (\p. qbs_l_inverse (measure_pmf p))" definition "qbs_spmf \ (\p. qbs_l_inverse (measure_spmf p))" declare [[coercion qbs_pmf]] lemma qbs_pmf_qbsP: fixes p :: "(_ :: countable) pmf" shows "qbs_pmf p \ qbs_space (monadP_qbs (count_space\<^sub>Q UNIV))" by(auto simp: qbs_pmf_def measure_to_qbs_cong_sets[of "count_space UNIV" "measure_pmf p",simplified] intro!: standard_borel_ne.qbs_l_inverse_in_space_monadP measure_pmf.prob_space_axioms) lemma qbs_pmf_qbs[qbs]: fixes p :: "(_ :: countable) pmf" shows "qbs_pmf p \ qbs_space (monadM_qbs (count_space\<^sub>Q UNIV))" by (simp add: qbs_pmf_qbsP qbs_space_monadPM) lemma qbs_spmf_qbs[qbs]: fixes q :: "(_ :: countable) spmf" shows "qbs_spmf q \ qbs_space (monadM_qbs (count_space\<^sub>Q UNIV))" by(auto simp: qbs_spmf_def measure_to_qbs_cong_sets[of "count_space UNIV" "measure_spmf q",simplified] intro!: standard_borel_ne.qbs_l_inverse_in_space_monadM subprob_space.s_finite_measure_subprob) lemma [simp]: fixes p :: "(_ :: countable) pmf" and q :: "(_ :: countable) spmf" shows qbs_l_qbs_pmf: "qbs_l (qbs_pmf p) = measure_pmf p" and qbs_l_qbs_spmf: "qbs_l (qbs_spmf q) = measure_spmf q" by(auto simp: qbs_pmf_def qbs_spmf_def intro!: standard_borel_ne.qbs_l_qbs_l_inverse subprob_space.s_finite_measure_subprob measure_pmf.subprob_space_axioms) lemma qbs_pmf_return_pmf: fixes x :: "_ :: countable" shows "qbs_pmf (return_pmf x) = return_qbs (count_space\<^sub>Q UNIV) x" proof - note return_qbs_morphismP[qbs] show ?thesis by(auto intro!: inj_onD[OF qbs_l_inj_P[where X="count_space\<^sub>Q UNIV"]] return_cong qbs_pmf_qbsP simp: qbs_l_return_qbs return_pmf.rep_eq) qed lemma qbs_pmf_bind_pmf: fixes p :: "('a :: countable) pmf" and f :: "'a \ ('b :: countable) pmf" shows "qbs_pmf (p \ f) = qbs_pmf p \ (\x. qbs_pmf (f x))" by(auto simp: measure_pmf_bind qbs_pmf_def space_prob_algebra measure_pmf.prob_space_axioms intro!: standard_borel_ne.qbs_l_inverse_bind[where N="count_space UNIV"] prob_space.s_finite_measure_prob) lemma qbs_pair_pmf: fixes p :: "('a :: countable) pmf" and q :: "('b :: countable) pmf" shows "qbs_pmf p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_pmf q = qbs_pmf (pair_pmf p q)" proof(rule inj_onD[OF qbs_l_inj_P[of "count_space\<^sub>Q UNIV \\<^sub>Q count_space\<^sub>Q UNIV"]]) show "qbs_l (qbs_pmf p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_pmf q) = qbs_l (qbs_pmf (pair_pmf p q))" by(simp add: measure_pair_pmf qbs_l_qbs_pair_measure[OF standard_borel_ne.standard_borel standard_borel_ne.standard_borel,of "count_space UNIV" "count_space UNIV"]) next note [qbs] = qbs_pmf_qbsP qbs_pair_measure_morphismP show "qbs_pmf p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_pmf q \ qbs_space (monadP_qbs (count_space\<^sub>Q UNIV \\<^sub>Q count_space\<^sub>Q UNIV))" "qbs_pmf (pair_pmf p q) \ qbs_space (monadP_qbs (count_space\<^sub>Q UNIV \\<^sub>Q count_space\<^sub>Q UNIV))" by auto (simp add: qbs_count_space_prod) qed subsubsection \ Density \ lift_definition density_qbs :: "['a qbs_measure, 'a \ ennreal] \ 'a qbs_measure" is "\(X,\,\) f. if f \ X \\<^sub>Q qbs_borel then (X, \, density \ (f \ \)) else (X, SOME a. a \ qbs_Mx X, null_measure borel)" proof safe fix X Y :: "'a quasi_borel" fix \ \ \ \ and f :: "_ \ ennreal" assume 1:"qbs_s_finite_eq (X, \, \) (Y, \, \)" then interpret qs: pair_qbs_s_finite X \ \ \ \ using qbs_s_finite_eq_dest[OF 1] by(simp add: pair_qbs_s_finite_def) have [simp]:"(SOME a. a \ qbs_Mx X) \ qbs_Mx X" "(SOME a. a \ qbs_Mx Y) \ qbs_Mx X" using qs.pq1.in_Mx by(simp_all only: some_in_eq qbs_s_finite_eq_dest[OF 1]) blast+ { assume "f \ X \\<^sub>Q qbs_borel" then have "qbs_s_finite_eq (X, \, density \ (f \ \)) (Y, \, density \ (f \ \))" by(auto simp: qbs_s_finite_eq_def lr_adjunction_correspondence density_distr[symmetric] comp_def qbs_s_finite_eq_dest[OF 1] qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def qs.pq1.mu_sets qs.pq2.mu_sets AE_distr_iff intro!: qs.pq1.s_finite_measure_density qs.pq2.s_finite_measure_density) } moreover have "f \ X \\<^sub>Q qbs_borel \ f \ Y \\<^sub>Q qbs_borel \ qbs_s_finite_eq (X, \, density \ (f \ \)) (Y, (SOME a. a \ qbs_Mx Y), null_measure borel)" "f \ X \\<^sub>Q qbs_borel \ f \ Y \\<^sub>Q qbs_borel \ qbs_s_finite_eq (X, (SOME a. a \ qbs_Mx X), null_measure borel) (Y, \, density \ (f \ \))" "f \ X \\<^sub>Q qbs_borel \ f \ Y \\<^sub>Q qbs_borel \ qbs_s_finite_eq (X, (SOME a. a \ qbs_Mx X), null_measure borel) (Y, (SOME a. a \ qbs_Mx Y), null_measure borel)" by(auto simp: qbs_s_finite_eq_dest[OF 1] qbs_s_finite_eq_def qbs_s_finite_def in_Mx_def qbs_s_finite_axioms_def distr_return null_measure_distr intro!: subprob_space.s_finite_measure_subprob subprob_spaceI) ultimately show "qbs_s_finite_eq (if f \ X \\<^sub>Q borel\<^sub>Q then (X, \, density \ (f \ \)) else (X, SOME aa. aa \ qbs_Mx X, null_measure borel)) (if f \ Y \\<^sub>Q borel\<^sub>Q then (Y, \, density \ (f \ \)) else (Y, SOME a. a \ qbs_Mx Y, null_measure borel))" by auto qed lemma(in qbs_s_finite) assumes "f \ X \\<^sub>Q qbs_borel" shows density_qbs:"density_qbs \X,\, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n f = \X, \, density \ (f \ \)\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" and density_qbs_s_finite: "qbs_s_finite X \ (density \ (f \ \))" using assms by(auto simp: density_qbs.abs_eq qbs_s_finite_def in_Mx_def lr_adjunction_correspondence qbs_s_finite_axioms_def mu_sets AE_distr_iff intro!: s_finite_measure_density) lemma density_qbs_density_qbs_eq: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" shows "density_qbs (density_qbs s f) g = density_qbs s (\x. f x * g x)" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qbs_s_finite X \ \ by simp show ?thesis using assms(2,3) by(simp add: hs(1) density_qbs[OF assms(2)] qbs_s_finite.density_qbs[OF density_qbs_s_finite[OF assms(2)] assms(3)] density_qbs lr_adjunction_correspondence density_density_eq) (simp add: comp_def) qed lemma qbs_l_density_qbs: assumes "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" shows "qbs_l (density_qbs s f) = density (qbs_l s) f" proof - from rep_qbs_space_monadM[OF assms(1)] obtain \ \ where s: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" by metis then interpret qbs_s_finite X \ \ by simp show ?thesis using assms(2) by(simp add: s(1) qbs_l qbs_s_finite.density_qbs[OF s(2) assms(2)] qbs_s_finite.qbs_l[OF qbs_s_finite.density_qbs_s_finite[OF s(2) assms(2)]] density_distr lr_adjunction_correspondence) (simp add: comp_def) qed corollary qbs_l_density_qbs_indicator: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "qbs_pred X P" shows "qbs_l (density_qbs s (indicator {x\qbs_space X. P x})) (qbs_space X) = qbs_l s {x\qbs_space X. P x} " proof - have 1[measurable]: "{x \ qbs_space X. P x} \ sets (qbs_to_measure X)" by (metis qbs_pred_iff_sets space_L assms(2)) have 2[qbs]: "indicator {x \ qbs_space X. P x} \ X \\<^sub>Q qbs_borel" by(rule indicator_qbs_morphism'') qbs show ?thesis using assms(2) by(auto simp: qbs_l_density_qbs[of _ X] emeasure_density[of "indicator {x\space (qbs_to_measure X). P x}" "qbs_l s",OF _ sets.top,simplified measurable_qbs_l'[OF assms(1)],OF borel_measurable_indicator[OF predE],simplified space_L space_qbs_l_in[OF assms(1)]] qbs_pred_iff_measurable_pred nn_set_integral_space[of "qbs_l s",simplified space_qbs_l_in[OF assms(1)]] nn_integral_indicator[of _ "qbs_l s",simplified sets_qbs_l[OF assms(1)]]) qed lemma qbs_nn_integral_density_qbs: assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q x. g x \(density_qbs s f)) = (\\<^sup>+\<^sub>Q x. f x * g x \s)" by(auto simp: qbs_nn_integral_def2_l qbs_l_density_qbs[of _ X] measurable_qbs_l'[OF assms(1)] lr_adjunction_correspondence[symmetric] intro!:nn_integral_density) lemma qbs_integral_density_qbs: fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" assumes [qbs]:"s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" and "AE\<^sub>Q x in s. f x \ 0" shows "(\\<^sub>Q x. g x \(density_qbs s f)) = (\\<^sub>Q x. f x *\<^sub>R g x \s)" using assms(4) by(auto simp: qbs_integral_def2_l qbs_l_density_qbs[of _ X] measurable_qbs_l'[OF assms(1)] lr_adjunction_correspondence[symmetric] AEq_qbs_l intro!: integral_density) lemma density_qbs_morphism[qbs]: "density_qbs \ monadM_qbs X \\<^sub>Q (X \\<^sub>Q qbs_borel) \\<^sub>Q monadM_qbs X" proof(rule curry_preserves_morphisms[OF pair_qbs_morphismI]) fix \ and \ :: "_ \ _ \ ennreal" assume h:"\ \ qbs_Mx (monadM_qbs X)" "\ \ qbs_Mx (X \\<^sub>Q qbs_borel)" hence [qbs]: "\ \ qbs_borel \\<^sub>Q monadM_qbs X" "\ \ qbs_borel \\<^sub>Q X \\<^sub>Q qbs_borel" by(simp_all add: qbs_Mx_is_morphisms) from rep_qbs_Mx_monadM[OF h(1)] obtain \ k where hk: "\ = (\r. \X, \, k r\\<^sub>s\<^sub>f\<^sub>i\<^sub>n)" "\ \ qbs_Mx X" "s_finite_kernel borel borel k" "\r. qbs_s_finite X \ (k r)" by metis then interpret a: in_Mx X \ by(simp add: in_Mx_def) have [measurable]: "(\(x, y). \ x (\ y)) \ borel_measurable (borel \\<^sub>M borel)" proof - have "(\(x, y). \ x (\ y)) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q qbs_borel" by simp thus ?thesis by(simp add: lr_adjunction_correspondence qbs_borel_prod borel_prod) qed have [simp]:"density_qbs (\ r) (\ r) = \X, \, density (k r) (\ r \ \)\\<^sub>s\<^sub>f\<^sub>i\<^sub>n " for r using hk(4) by(auto simp add: hk(1) density_qbs.abs_eq[OF qbs_s_finite.qbs_s_finite_eq_refl[OF hk(4)]]) show "(\r. density_qbs (fst (\ r,\ r)) (snd (\ r,\ r))) \ qbs_Mx (monadM_qbs X)" by(auto simp: monadM_qbs_Mx comp_def intro!: exI[where x=\] exI[where x="\r. density (k r) (\ r \ \)"] s_finite_kernel.density_s_finite_kernel[OF hk(3)]) qed lemma density_qbs_cong_AE: assumes [qbs]: "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" and "AE\<^sub>Q x in s. f x = g x" shows "density_qbs s f = density_qbs s g" proof(rule inj_onD[OF qbs_l_inj[of X]]) show "qbs_l (density_qbs s f) = qbs_l (density_qbs s g)" using assms(4) by(auto simp: qbs_l_density_qbs[of _ X] measurable_qbs_l'[OF assms(1)] AEq_qbs_l lr_adjunction_correspondence[symmetric] intro!: density_cong) qed simp_all corollary density_qbs_cong: assumes [qbs]: "s \ qbs_space (monadM_qbs X)" "f \ X \\<^sub>Q qbs_borel" "g \ X \\<^sub>Q qbs_borel" and "\x. x \ qbs_space X \ f x = g x" shows "density_qbs s f = density_qbs s g" by(auto intro!: density_qbs_cong_AE[of _ X] AEq_I2[of _ X] assms(4)) lemma density_qbs_1[simp]: "density_qbs s (\x. 1) = s" proof - obtain X where s[qbs]: "s \ qbs_space (monadM_qbs X)" using in_qbs_space_of by blast show ?thesis by(auto intro!: inj_onD[OF qbs_l_inj _ _ s] simp: qbs_l_density_qbs[of _ X] density_1) qed lemma pair_density_qbs: assumes [qbs]: "p \ qbs_space (monadM_qbs X)" "q \ qbs_space (monadM_qbs Y)" and [qbs]: "f \ X \\<^sub>Q qbs_borel" "g \ Y \\<^sub>Q qbs_borel" shows "density_qbs p f \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s density_qbs q g = density_qbs (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\(x,y). f x * g y)" proof(safe intro!: qbs_measure_eqI[of _ "X \\<^sub>Q Y"]) fix h :: "_ \ ennreal" assume h[qbs]:"h \ X \\<^sub>Q Y \\<^sub>Q borel\<^sub>Q" show "(\\<^sup>+\<^sub>Q z. h z \(density_qbs p f \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s density_qbs q g)) = (\\<^sup>+\<^sub>Q z. h z \(density_qbs (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\(x, y). f x * g y)))" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. h (x, y) \density_qbs q g \density_qbs p f)" by(simp add: qbs_nn_integral_Fubini_fst[of _ X _ Y]) also have "... = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. g y * h (x, y) \q \density_qbs p f)" by(auto intro!: qbs_nn_integral_cong[of _ X] simp: qbs_nn_integral_density_qbs[of _ Y]) also have "... = ?rhs" by(auto simp add: qbs_nn_integral_density_qbs[of _ X] qbs_nn_integral_density_qbs[of _ "X \\<^sub>Q Y"] split_beta' qbs_nn_integral_Fubini_fst[of _ X _ Y,symmetric] qbs_nn_integral_cmult[of _ Y] mult.assoc intro!: qbs_nn_integral_cong[of _ X]) finally show ?thesis . qed qed simp_all subsubsection \ Normalization \ definition normalize_qbs :: "'a qbs_measure \ 'a qbs_measure" where "normalize_qbs s \ (let X = qbs_space_of s; r = qbs_l s (qbs_space X) in if r \ 0 \ r \ \ then density_qbs s (\x. 1 / r) else qbs_null_measure X)" lemma assumes "s \ qbs_space (monadM_qbs X)" shows normalize_qbs: "qbs_l s (qbs_space X) \ 0 \ qbs_l s (qbs_space X) \ \ \ normalize_qbs s = density_qbs s (\x. 1 / emeasure (qbs_l s) (qbs_space X))" and normalize_qbs0: "qbs_l s (qbs_space X) = 0 \ normalize_qbs s = qbs_null_measure X" and normalize_qbsinfty: "qbs_l s (qbs_space X) = \ \ normalize_qbs s = qbs_null_measure X" by(auto simp: qbs_space_of_in[OF assms(1)] normalize_qbs_def) lemma normalize_qbs_prob: assumes "s \ qbs_space (monadM_qbs X)" "qbs_l s (qbs_space X) \ 0" "qbs_l s (qbs_space X) \ \" shows "normalize_qbs s \ qbs_space (monadP_qbs X)" unfolding normalize_qbs[OF assms] proof - obtain \ \ where hs: "s = \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "qbs_s_finite X \ \" using rep_qbs_space_monadM assms(1) by meson interpret qs: qbs_s_finite X \ \ by fact have "density_qbs s (\x. 1 / emeasure (qbs_l s) (qbs_space X)) = density_qbs \X, \, \\\<^sub>s\<^sub>f\<^sub>i\<^sub>n (\x. 1 / emeasure (qbs_l s) (qbs_space X))" by(simp add: hs) also have "... \ qbs_space (monadP_qbs X)" by(auto simp add: qs.density_qbs monadP_qbs_space qbs_s_finite.qbs_l[OF qs.density_qbs_s_finite,of "\x. 1 / emeasure (qbs_l s) (qbs_space X)",simplified] qbs_s_finite.qbs_space_of[OF qs.density_qbs_s_finite,of "\x. 1 / emeasure (qbs_l s) (qbs_space X)",simplified] intro!: prob_space.prob_space_distr, auto intro!: prob_spaceI simp: emeasure_density) (insert assms(2,3),auto simp: hs qs.qbs_l emeasure_distr emeasure_distr[of _ _ "qbs_to_measure X",OF _ sets.top,simplified space_L] divide_eq_1_ennreal ennreal_divide_times) finally show "density_qbs s (\x. 1 / emeasure (qbs_l s) (qbs_space X)) \ qbs_space (monadP_qbs X)" . qed lemma normalize_qbs_morphism[qbs]: "normalize_qbs \ monadM_qbs X \\<^sub>Q monadM_qbs X" proof - have "(if emeasure (qbs_l s) (qbs_space X) \ 0 \ emeasure (qbs_l s) (qbs_space X) \ \ then density_qbs s (\x. 1 / emeasure (qbs_l s) (qbs_space X)) else qbs_null_measure X) = normalize_qbs s" (is "?f s = _") if s:"s \ qbs_space (monadM_qbs X)" for s by(simp add: qbs_space_of_in[OF s] normalize_qbs_def) moreover have "(\s. ?f s) \ monadM_qbs X \\<^sub>Q monadM_qbs X" proof(cases "qbs_space X = {}") case True then show ?thesis by(simp add: qbs_morphism_from_empty monadM_qbs_empty_iff[of X]) next case X:False have [qbs]:"(\s. emeasure (qbs_l s) (qbs_space X)) \ monadM_qbs X \\<^sub>Q qbs_borel" by(rule qbs_l_morphism[OF sets.top[of "qbs_to_measure X",simplified space_L]]) have [qbs]: "qbs_null_measure X \ qbs_space (monadM_qbs X)" by(auto intro!: qbs_null_measure_in_Mx X) have [qbs]: "(\s x. 1 / emeasure (qbs_l s) (qbs_space X)) \ monadM_qbs X \\<^sub>Q X \\<^sub>Q qbs_borel" by(rule arg_swap_morphism) simp show ?thesis by qbs qed ultimately show ?thesis by(simp cong: qbs_morphism_cong) qed lemma normalize_qbs_morphismP: assumes [qbs]:"s \ X \\<^sub>Q monadM_qbs Y" and "\x. x \ qbs_space X \ qbs_l (s x) (qbs_space Y) \ 0" "\x. x \ qbs_space X \ qbs_l (s x) (qbs_space Y) \ \" shows "(\x. normalize_qbs (s x)) \ X \\<^sub>Q monadP_qbs Y" by(rule qbs_morphism_monadPI'[OF normalize_qbs_prob]) (use assms(2,3) qbs_morphism_space[OF assms(1)] in auto) lemma normalize_qbs_monadP_ident: assumes "s \ qbs_space (monadP_qbs X)" shows "normalize_qbs s = s" using normalize_qbs[OF qbs_space_monadPM[OF assms]] prob_space.emeasure_space_1[OF qbs_l_prob_space[OF assms]] by(auto simp: space_qbs_l_in[OF qbs_space_monadPM[OF assms]] intro!: inj_onD[OF qbs_l_inj_P _ _ assms]) corollary normalize_qbs_idenpotent: "normalize_qbs (normalize_qbs s) = normalize_qbs s" proof - obtain X where s[qbs]: "s \ qbs_space (monadM_qbs X)" using in_qbs_space_of by blast then have X: "qbs_space X \ {}" by (metis qbs_s_space_of_not_empty qbs_space_of_in) then obtain x where x:"x \ qbs_space X" by auto consider "qbs_l s (qbs_space X) = 0" | "qbs_l s (qbs_space X) = \" | "qbs_l s (qbs_space X) \ 0" "qbs_l s (qbs_space X) \ \" by auto then show ?thesis proof cases case 1 then show ?thesis using normalize_qbs0[OF qbs_null_measure_in_Mx[OF X]] by(simp add: normalize_qbs0[OF s] qbs_null_measure_null_measure[OF X]) next case 2 then show ?thesis using normalize_qbs0[OF qbs_null_measure_in_Mx[OF X]] by(simp add: normalize_qbsinfty[OF s] qbs_null_measure_null_measure[OF X]) next case 3 have "normalize_qbs s \ qbs_space (monadP_qbs X)" by(rule qbs_morphism_space[OF normalize_qbs_morphismP[of "\x. s"],of X X x]) (auto simp: 3 x) then show ?thesis by(simp add: normalize_qbs_monadP_ident) qed qed subsubsection \ Product Measures \ definition PiQ_measure :: "['a set, 'a \ 'b qbs_measure] \ ('a \ 'b) qbs_measure" where "PiQ_measure \ (\I si. if (\i\I. \Mi. standard_borel_ne Mi \ si i \ qbs_space (monadM_qbs (measure_to_qbs Mi))) then if countable I \ (\i\I. prob_space (qbs_l (si i))) then qbs_l_inverse (\\<^sub>M i\I. qbs_l (si i)) else if finite I \ (\i\I. sigma_finite_measure (qbs_l (si i))) then qbs_l_inverse (\\<^sub>M i\I. qbs_l (si i)) else qbs_null_measure (\\<^sub>Q i\I. qbs_space_of (si i)) else qbs_null_measure (\\<^sub>Q i\I. qbs_space_of (si i)))" syntax "_PiQ_measure" :: "pttrn \ 'i set \ 'a qbs_measure \ ('i => 'a) qbs_measure" ("(3\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s _\_./ _)" 10) translations "\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s x\I. X" == "CONST PiQ_measure I (\x. X)" context fixes I and Mi assumes standard_borel_ne:"\i. i \ I \ standard_borel_ne (Mi i)" begin context assumes countableI:"countable I" begin interpretation sb:standard_borel_ne "\\<^sub>M i\I. (borel :: real measure)" by (simp add: countableI product_standard_borel_ne) interpretation sbM: standard_borel_ne "\\<^sub>M i\I. Mi i" by (simp add: countableI standard_borel_ne product_standard_borel_ne) lemma assumes "\i. i \ I \ si i \ qbs_space (monadP_qbs (measure_to_qbs (Mi i)))" and "\i. i \ I \ si i = \measure_to_qbs (Mi i), \ i, \ i\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" "\i. i \ I \ qbs_prob (measure_to_qbs (Mi i)) (\ i) (\ i)" shows PiQ_measure_prob_eq: "(\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = \measure_to_qbs (\\<^sub>M i\I. Mi i), sbM.from_real, distr (\\<^sub>M i\I. qbs_l (si i)) borel sbM.to_real\\<^sub>s\<^sub>f\<^sub>i\<^sub>n" (is "_ = ?rhs") and PiQ_measure_qbs_prob: "qbs_prob (measure_to_qbs (\\<^sub>M i\I. Mi i)) sbM.from_real (distr (\\<^sub>M i\I. qbs_l (si i)) borel sbM.to_real)" (is "?qbsprob") proof - have [measurable_cong,simp]: "prob_space (\\<^sub>M i\I. qbs_l (si i))" "sets (\\<^sub>M i\I. qbs_l (si i)) = sets (\\<^sub>M i\I. Mi i)" using sets_qbs_l[OF assms(1)[THEN qbs_space_monadPM]] standard_borel.lr_sets_ident[OF standard_borel_ne.standard_borel[OF standard_borel_ne]] by(auto cong: sets_PiM_cong intro!: prob_space_PiM qbs_l_prob_space assms(1)) show ?qbsprob by(auto simp: pair_qbs_s_finite_def intro!: qbs_prob.qbs_s_finite sbM.qbs_l_inverse_qbs_prob) have "(\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = qbs_l_inverse (\\<^sub>M i\I. qbs_l (si i))" using countableI assms(1)[THEN qbs_space_monadPM] qbs_l_prob_space[OF assms(1)] standard_borel_ne by(auto simp: PiQ_measure_def) also have "... = ?rhs" by(auto intro!: sbM.qbs_l_inverse_def2 prob_space.s_finite_measure_prob cong: sets_PiM_cong[OF refl]) finally show "(\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = ?rhs" . qed lemma qbs_l_PiQ_measure_prob: assumes "\i. i \ I \ si i \ qbs_space (monadP_qbs (measure_to_qbs (Mi i)))" shows "qbs_l (\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = (\\<^sub>M i\I. qbs_l (si i))" proof - have "qbs_l (\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = qbs_l (qbs_l_inverse (\\<^sub>M i\I. qbs_l (si i)))" using countableI assms(1)[THEN qbs_space_monadPM] qbs_l_prob_space[OF assms(1)] standard_borel_ne by(auto simp: PiQ_measure_def) also have "... = (\\<^sub>M i\I. qbs_l (si i))" using sets_qbs_l[OF assms(1)[THEN qbs_space_monadPM]] standard_borel.lr_sets_ident[OF standard_borel_ne.standard_borel[OF standard_borel_ne]] by(auto intro!: sbM.qbs_l_qbs_l_inverse_prob prob_space_PiM qbs_l_prob_space[OF assms(1)] cong: sets_PiM_cong) finally show ?thesis . qed end context assumes finI: "finite I" begin interpretation sb:standard_borel_ne "\\<^sub>M i\I. (borel :: real measure)" by (simp add: finI product_standard_borel_ne countable_finite) interpretation sbM: standard_borel_ne "\\<^sub>M i\I. Mi i" by (simp add: countable_finite finI standard_borel_ne product_standard_borel_ne) lemma qbs_l_PiQ_measure: assumes "\i. i \ I \ si i \ qbs_space (monadM_qbs (measure_to_qbs (Mi i)))" and "\i. i \ I \ sigma_finite_measure (qbs_l (si i))" shows "qbs_l (\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = (\\<^sub>M i\I. qbs_l (si i))" proof - have [simp]: "s_finite_measure (\\<^sub>M i\I. qbs_l (si i))" proof - have "(\\<^sub>M i\I. qbs_l (si i)) = (\\<^sub>M i\I. if i \ I then qbs_l (si i) else null_measure (count_space UNIV))" by(simp cong: PiM_cong) also have "s_finite_measure ..." by(auto intro!: sigma_finite_measure.s_finite_measure product_sigma_finite.sigma_finite finI simp: product_sigma_finite_def assms(2)) (auto intro!: finite_measure.sigma_finite_measure finite_measureI) finally show ?thesis . qed have "qbs_l (\\<^sub>Q\<^sub>m\<^sub>e\<^sub>a\<^sub>s i\I. si i) = qbs_l (qbs_l_inverse (\\<^sub>M i\I. qbs_l (si i)))" using finI assms(1) assms(2) standard_borel_ne by(fastforce simp: PiQ_measure_def) also have "... = (\\<^sub>M i\I. qbs_l (si i))" using sets_qbs_l[OF assms(1)] standard_borel.lr_sets_ident[OF standard_borel_ne.standard_borel[OF standard_borel_ne]] by(auto intro!: sbM.qbs_l_qbs_l_inverse prob_space_PiM cong: sets_PiM_cong) finally show ?thesis . qed end end subsection \Measures\ subsubsection \ The Lebesgue Measure \ definition lborel_qbs ("lborel\<^sub>Q") where "lborel_qbs \ qbs_l_inverse lborel" lemma lborel_qbs_qbs[qbs]: "lborel_qbs \ qbs_space (monadM_qbs qbs_borel)" by(auto simp: lborel_qbs_def measure_to_qbs_cong_sets[OF sets_lborel,symmetric] intro!: standard_borel_ne.qbs_l_inverse_in_space_monadM lborel.s_finite_measure_axioms) lemma qbs_l_lborel_qbs[simp]: "qbs_l lborel\<^sub>Q = lborel" by(auto intro!: standard_borel_ne.qbs_l_qbs_l_inverse lborel.s_finite_measure_axioms simp: lborel_qbs_def) corollary shows qbs_integral_lborel: "(\\<^sub>Q x. f x \lborel_qbs) = (\x. f x \lborel)" and qbs_nn_integral_lborel: "(\\<^sup>+\<^sub>Q x. f x \lborel_qbs) = (\\<^sup>+x. f x \lborel)" by(simp_all add: qbs_integral_def2_l qbs_nn_integral_def2_l) lemma(in standard_borel_ne) measure_with_args_morphism: assumes "s_finite_kernel X M k" shows "qbs_l_inverse \ k \ measure_to_qbs X \\<^sub>Q monadM_qbs (measure_to_qbs M)" proof(safe intro!: qbs_morphismI) fix \ :: "real \ _" assume "\ \ qbs_Mx (measure_to_qbs X)" then have h[measurable]:"\ \ borel \\<^sub>M X" by(simp add: qbs_Mx_R) interpret s:s_finite_kernel X M k by fact have 1: "\r. sets (k (\ r)) = sets M" "\r. s_finite_measure (k (\ r))" using measurable_space[OF h] s.kernel_sets by(auto intro!: s.image_s_finite_measure) show "qbs_l_inverse \ k \ \ \ qbs_Mx (monadM_qbs (measure_to_qbs M))" by(auto intro!: exI[where x=from_real] exI[where x="(\r. distr (k (\ r)) borel to_real)"] s_finite_kernel.comp_measurable[OF s_finite_kernel.distr_s_finite_kernel[OF assms]] simp: monadM_qbs_Mx qbs_Mx_R qbs_l_inverse_def2[OF 1] comp_def) qed lemma(in standard_borel_ne) measure_with_args_morphismP: assumes [measurable]:"\ \ X \\<^sub>M prob_algebra M" shows "qbs_l_inverse \ \ \ measure_to_qbs X \\<^sub>Q monadP_qbs (measure_to_qbs M)" by(rule qbs_morphism_monadPI'[OF _ measure_with_args_morphism]) (insert measurable_space[OF assms], auto simp: qbs_space_R space_prob_algebra prob_kernel_def' intro!: qbs_l_inverse_in_space_monadP prob_kernel.s_finite_kernel_prob_kernel) subsubsection \ Counting Measure \ abbreviation "counting_measure_qbs A \ qbs_l_inverse (count_space A)" lemma qbs_nn_integral_count_space_nat: fixes f :: "nat \ ennreal" shows "(\\<^sup>+\<^sub>Q i. f i \counting_measure_qbs UNIV) = (\i. f i)" by(simp add: standard_borel_ne.qbs_l_qbs_l_inverse[OF _ refl sigma_finite_measure.s_finite_measure[OF sigma_finite_measure_count_space]] qbs_nn_integral_def2_l nn_integral_count_space_nat) subsubsection \ Normal Distribution \ lemma qbs_normal_distribution_qbs: "(\\ \. density_qbs lborel\<^sub>Q (normal_density \ \)) \ qbs_borel \\<^sub>Q qbs_borel \\<^sub>Q monadM_qbs qbs_borel" by simp lemma qbs_l_qbs_normal_distribution[simp]: "qbs_l (density_qbs lborel\<^sub>Q (normal_density \ \)) = density lborel (normal_density \ \)" by(auto simp: qbs_l_density_qbs[of _ qbs_borel]) lemma qbs_normal_distribution_P: "\ > 0 \ density_qbs lborel\<^sub>Q (normal_density \ \) \ qbs_space (monadP_qbs qbs_borel)" by(auto simp: monadP_qbs_def sub_qbs_space prob_space_normal_density) lemma qbs_normal_distribution_integral: "(\\<^sub>Q x. f x \ (density_qbs lborel\<^sub>Q (normal_density \ \))) = (\ x. f x \ (density lborel (\x. ennreal (normal_density \ \ x))))" by(auto simp: qbs_integral_def2_l) lemma qbs_normal_distribution_expectation: assumes [measurable]:"f \ borel_measurable borel" and [arith]: "\ > 0" shows "(\\<^sub>Q x. f x \ (density_qbs lborel\<^sub>Q (normal_density \ \))) = (\ x. normal_density \ \ x * f x \ lborel)" by(simp add: qbs_normal_distribution_integral integral_real_density integral_density) lemma qbs_normal_posterior: assumes [arith]: "\ > 0" "\' > 0" shows "normalize_qbs (density_qbs (density_qbs lborel\<^sub>Q (normal_density \ \)) (normal_density \' \')) = density_qbs lborel\<^sub>Q (normal_density ((\ * \'\<^sup>2 + \' * \\<^sup>2) / (\\<^sup>2 + \'\<^sup>2)) (\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2)))" (is "?lhs = ?rhs") proof - have 0: "\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2) > 0" "sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)) > 0" by (simp_all add: power2_eq_square sum_squares_gt_zero_iff) have 1:"qbs_l (density_qbs lborel\<^sub>Q (\x. ennreal (1 / sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)) * exp (- ((\ - \')\<^sup>2 / (2 * \\<^sup>2 + 2 * \'\<^sup>2))) * normal_density ((\ * \'\<^sup>2 + \' * \\<^sup>2) / (\\<^sup>2 + \'\<^sup>2)) (\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2)) x))) UNIV = ennreal (exp (- ((\ - \')\<^sup>2 / (2 * \\<^sup>2 + 2 * \'\<^sup>2))) / sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)))" using prob_space.emeasure_space_1[OF prob_space_normal_density[OF 0(1)]] by(auto simp add: qbs_l_density_qbs[of _ qbs_borel] emeasure_density ennreal_mult' nn_integral_cmult simp del: times_divide_eq_left) (simp add: ennreal_mult'[symmetric]) have "?lhs = normalize_qbs (density_qbs lborel\<^sub>Q (\x. ennreal (1 / sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)) * exp (- ((\ - \')\<^sup>2 / (2 * \\<^sup>2 + 2 * \'\<^sup>2))) * normal_density ((\ * \'\<^sup>2 + \' * \\<^sup>2) / (\\<^sup>2 + \'\<^sup>2)) (\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2)) x)))" by(simp add: density_qbs_density_qbs_eq[of _ qbs_borel] ennreal_mult'[symmetric] normal_density_times del: times_divide_eq_left) also have "... = density_qbs (density_qbs lborel\<^sub>Q (\x. ennreal (1 / sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)) * exp (- ((\ - \')\<^sup>2 / (2 * \\<^sup>2 + 2 * \'\<^sup>2))) * normal_density ((\ * \'\<^sup>2 + \' * \\<^sup>2) / (\\<^sup>2 + \'\<^sup>2)) (\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2)) x))) (\x. 1 / emeasure (qbs_l (density_qbs lborel\<^sub>Q (\x. ennreal (1 / sqrt (2 * pi * (\\<^sup>2 + \'\<^sup>2)) * exp (- ((\ - \')\<^sup>2 / (2 * \\<^sup>2 + 2 * \'\<^sup>2))) * normal_density ((\ * \'\<^sup>2 + \' * \\<^sup>2) / (\\<^sup>2 + \'\<^sup>2)) (\ * \' / sqrt (\\<^sup>2 + \'\<^sup>2)) x)))) (qbs_space borel\<^sub>Q))" by(rule normalize_qbs) (simp_all add: 1 del: times_divide_eq_left) also have "... = ?rhs" by(simp add: 1 density_qbs_density_qbs_eq[of _ qbs_borel] del: times_divide_eq_left, auto intro!: density_qbs_cong[of _ qbs_borel]) (insert 0, auto simp: ennreal_1[symmetric] ennreal_mult'[symmetric] divide_ennreal normal_density_def simp del: ennreal_1) finally show ?thesis . qed subsubsection \ Uniform Distribution \ definition uniform_qbs :: "'a qbs_measure \ 'a set \ 'a qbs_measure" where "uniform_qbs \ (\s A. qbs_l_inverse (uniform_measure (qbs_l s) A))" lemma(in standard_borel_ne) qbs_l_uniform_qbs': assumes "sets \ = sets M" "s_finite_measure \" "\ A \ 0" shows "qbs_l (uniform_qbs (qbs_l_inverse \) A) = uniform_measure \ A" (is "?lhs = ?rhs") proof - have "?lhs = qbs_l (qbs_l_inverse (uniform_measure \ A))" by(simp add: qbs_l_qbs_l_inverse[OF assms(1,2)] uniform_qbs_def) also have "... = ?rhs" proof(rule qbs_l_qbs_l_inverse) consider "\ A = \" | "\ A \ \" by auto then show "s_finite_measure (uniform_measure \ A)" proof cases case 1 have A[measurable]: "A \ sets \" using assms(3) emeasure_notin_sets by blast have "uniform_measure \ A = density \ (\x. 0)" by(auto simp: uniform_measure_def 1 intro!: density_cong) also have "... = null_measure \" by(simp add: null_measure_eq_density) finally show ?thesis by(auto intro!: finite_measure.s_finite_measure_finite_measure finite_measureI) next case 2 show ?thesis by(rule prob_space.s_finite_measure_prob[OF prob_space_uniform_measure[OF assms(3) 2]]) qed qed(simp add: assms) finally show ?thesis . qed corollary(in standard_borel_ne) qbs_l_uniform_qbs: assumes "s \ qbs_space (monadM_qbs (measure_to_qbs M))" "qbs_l s A \ 0" shows "qbs_l (uniform_qbs s A) = uniform_measure (qbs_l s) A" by(simp add: qbs_l_uniform_qbs'[OF sets_qbs_l[OF assms(1),simplified lr_sets_ident] qbs_l_s_finite.s_finite_measure_axioms assms(2),symmetric] qbs_l_inverse_qbs_l[OF assms(1)]) lemma interval_uniform_qbs: "(\a b. uniform_qbs lborel\<^sub>Q {a<.. borel\<^sub>Q \\<^sub>Q borel\<^sub>Q \\<^sub>Q monadM_qbs borel\<^sub>Q" proof(rule curry_preserves_morphisms) have "(\xy. uniform_qbs lborel\<^sub>Q {fst xy<.. (\xy. uniform_measure lborel {fst xy<.. measure_to_qbs (borel \\<^sub>M borel) \\<^sub>Q monadM_qbs borel\<^sub>Q" proof(safe intro!: standard_borel_ne.measure_with_args_morphism measure_kernel.s_finite_kernel_finite_bounded) show "measure_kernel (borel \\<^sub>M borel) borel (\xy. uniform_measure lborel {fst xy<.. sets borel" have [simp]:"emeasure lborel ({fst x<.. B) / emeasure lborel {fst x<.. snd x then emeasure lborel ({fst x<.. B) / ennreal (snd x - fst x) else 0)" for x by auto show "(\x. emeasure (uniform_measure lborel {fst x<.. borel_measurable (borel \\<^sub>M borel)" by (simp, measurable) auto qed auto next show "(a, b) \ space (borel \\<^sub>M borel) \ emeasure (uniform_measure lborel {fst (a, b)<.." for a b :: real by(cases "a \ b") (use ennreal_divide_eq_top_iff top.not_eq_extremum in auto) qed simp finally show "(\xy. uniform_qbs lborel\<^sub>Q {fst xy<.. borel\<^sub>Q \\<^sub>Q borel\<^sub>Q \\<^sub>Q monadM_qbs borel\<^sub>Q" by (simp add: borel_prod qbs_borel_prod) qed context fixes a b :: real assumes [arith]:"a < b" begin lemma qbs_uniform_distribution_expectation: assumes "f \ qbs_borel \\<^sub>Q qbs_borel" shows "(\\<^sup>+\<^sub>Q x. f x \uniform_qbs lborel\<^sub>Q {a<..\<^sup>+x \ {a<..lborel) / (b - a)" proof - have [measurable]: "f \ borel_measurable borel" using assms qbs_Mx_is_morphisms qbs_Mx_qbs_borel by blast show ?thesis by(auto simp: qbs_nn_integral_def2_l standard_borel_ne.qbs_l_uniform_qbs[of borel lborel_qbs] nn_integral_uniform_measure) qed end subsubsection \ Bernoulli Distribution \ abbreviation qbs_bernoulli :: "real \ bool qbs_measure" where "qbs_bernoulli \ (\x. qbs_pmf (bernoulli_pmf x))" lemma bernoulli_measurable: "(\x. measure_pmf (bernoulli_pmf x)) \ borel \\<^sub>M prob_algebra (count_space UNIV)" proof(rule measurable_prob_algebra_generated[where \=UNIV and G=UNIV]) fix A :: "bool set" have "A \ {True,False}" by auto then consider "A = {}" | "A = {True}" | "A = {False}" | "A = {False,True}" by auto thus "(\a. emeasure (measure_pmf (bernoulli_pmf a)) A) \ borel_measurable borel" by(cases,simp_all add: emeasure_measure_pmf_finite bernoulli_pmf.rep_eq UNIV_bool[symmetric]) qed (auto simp add: sets_borel_eq_count_space Int_stable_def measure_pmf.prob_space_axioms) lemma qbs_bernoulli_morphism: "qbs_bernoulli \ qbs_borel \\<^sub>Q monadP_qbs (qbs_count_space UNIV)" using standard_borel_ne.measure_with_args_morphismP[OF _ bernoulli_measurable] by (simp add: qbs_pmf_def comp_def) lemma qbs_bernoulli_expectation: assumes [simp]: "0 \ p" "p \ 1" shows "(\\<^sub>Q x. f x \qbs_bernoulli p) = f True * p + f False * (1 - p)" by(simp add: qbs_integral_def2_l) end \ No newline at end of file diff --git a/thys/S_Finite_Measure_Monad/QBS_Morphism.thy b/thys/S_Finite_Measure_Monad/QBS_Morphism.thy --- a/thys/S_Finite_Measure_Monad/QBS_Morphism.thy +++ b/thys/S_Finite_Measure_Monad/QBS_Morphism.thy @@ -1,427 +1,427 @@ (* Title: QBS_Morphism.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) subsection \ Morphisms of Quasi-Borel Spaces \ theory QBS_Morphism imports "QuasiBorel" begin abbreviation qbs_morphism :: "['a quasi_borel, 'b quasi_borel] \ ('a \ 'b) set" (infixr "\\<^sub>Q" 60) where "X \\<^sub>Q Y \ qbs_space (X \\<^sub>Q Y)" lemma qbs_morphismI: "(\\. \ \ qbs_Mx X \ f \ \ \ qbs_Mx Y) \ f \ X \\<^sub>Q Y" by(auto simp: exp_qbs_space) lemma qbs_morphism_def: "X \\<^sub>Q Y = {f\qbs_space X \ qbs_space Y. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}" unfolding exp_qbs_space proof safe fix f x assume h:"x \ qbs_space X " "\\\qbs_Mx X. f \ \ \ qbs_Mx Y" then have "(\r. x) \ qbs_Mx X" by simp hence "f \ (\r. x) \ qbs_Mx Y" using h by blast with qbs_Mx_to_X show "f x \ qbs_space Y" by auto qed auto lemma qbs_morphism_Mx: assumes "f \ X \\<^sub>Q Y" "\ \ qbs_Mx X" shows "f \ \ \ qbs_Mx Y" using assms by(auto simp: qbs_morphism_def) lemma qbs_morphism_space: assumes "f \ X \\<^sub>Q Y" "x \ qbs_space X" shows "f x \ qbs_space Y" using assms by(auto simp: qbs_morphism_def) lemma qbs_morphism_ident[simp]: "id \ X \\<^sub>Q X" by(auto intro: qbs_morphismI) lemma qbs_morphism_ident'[simp]: "(\x. x) \ X \\<^sub>Q X" using qbs_morphism_ident by(simp add: id_def) lemma qbs_morphism_comp: assumes "f \ X \\<^sub>Q Y" "g \ Y \\<^sub>Q Z" shows "g \ f \ X \\<^sub>Q Z" using assms by (simp add: comp_assoc Pi_def qbs_morphism_def) lemma qbs_morphism_compose_rev: assumes "f \ Y \\<^sub>Q Z" and "g \ X \\<^sub>Q Y" shows "(\x. f (g x)) \ X \\<^sub>Q Z" using qbs_morphism_comp[OF assms(2,1)] by(simp add: comp_def) lemma qbs_morphism_compose: assumes "g \ X \\<^sub>Q Y" and "f \ Y \\<^sub>Q Z" shows "(\x. f (g x)) \ X \\<^sub>Q Z" using qbs_morphism_compose_rev[OF assms(2,1)] . lemma qbs_morphism_cong': assumes "\x. x \ qbs_space X \ f x = g x" and "f \ X \\<^sub>Q Y" shows "g \ X \\<^sub>Q Y" proof(rule qbs_morphismI) fix \ assume 1:"\ \ qbs_Mx X" have "g \ \ = f \ \" proof fix x have "\ x \ qbs_space X" using 1 qbs_decomp[of X] qbs_Mx_to_X by auto thus "(g \ \) x = (f \ \) x" using assms(1) by simp qed thus "g \ \ \ qbs_Mx Y" using 1 assms(2) by(simp add: qbs_morphism_def) qed lemma qbs_morphism_cong: assumes "\x. x \ qbs_space X \ f x = g x" shows "f \ X \\<^sub>Q Y \ g \ X \\<^sub>Q Y" using assms by(auto simp: qbs_morphism_cong'[of _ f g] qbs_morphism_cong'[of _ g f]) lemma qbs_morphism_const: assumes "y \ qbs_space Y" shows "(\x. y) \ X \\<^sub>Q Y" using assms by (auto intro: qbs_morphismI) lemma qbs_morphism_from_empty: "qbs_space X = {} \ f \ X \\<^sub>Q Y" by(auto intro!: qbs_morphismI simp: qbs_empty_equiv) lemma unit_quasi_borel_terminal: "\! f. f \ X \\<^sub>Q unit_quasi_borel" by(fastforce simp: qbs_morphism_def) definition to_unit_quasi_borel :: "'a \ unit" ("!\<^sub>Q") where "to_unit_quasi_borel \ (\r.())" lemma to_unit_quasi_borel_morphism: "!\<^sub>Q \ X \\<^sub>Q unit_quasi_borel" by(auto simp add: to_unit_quasi_borel_def qbs_morphism_def) lemma qbs_morphism_subD: assumes "f \ X \\<^sub>Q sub_qbs Y A" shows "f \ X \\<^sub>Q Y" using qbs_morphism_Mx[OF assms] by(auto intro!: qbs_morphismI simp: sub_qbs_Mx) lemma qbs_morphism_subI1: assumes "f \ X \\<^sub>Q Y" "\x. x \ qbs_space X \ f x \ A" shows "f \ X \\<^sub>Q sub_qbs Y A" using qbs_morphism_space[OF assms(1)] qbs_morphism_Mx[OF assms(1)] assms(2) qbs_Mx_to_X[of _ X] by(auto intro!: qbs_morphismI simp: sub_qbs_Mx) lemma qbs_morphism_subI2: assumes "f \ X \\<^sub>Q Y" shows "f \ sub_qbs X A \\<^sub>Q Y" using qbs_morphism_Mx[OF assms] by(auto intro!: qbs_morphismI simp: sub_qbs_Mx) corollary qbs_morphism_subsubI: assumes "f \ X \\<^sub>Q Y" "\x. x \ A \ f x \ B" shows "f \ sub_qbs X A \\<^sub>Q sub_qbs Y B" by(rule qbs_morphism_subI1) (auto intro!: qbs_morphism_subI2 assms simp: sub_qbs_space) lemma map_qbs_morphism_f: "f \ X \\<^sub>Q map_qbs f X" by(auto intro!: qbs_morphismI simp: map_qbs_Mx) lemma map_qbs_morphism_inverse_f: assumes "\x. x \ qbs_space X \ g (f x) = x" shows "g \ map_qbs f X \\<^sub>Q X" proof - { fix \ assume h:"\ \ qbs_Mx X" from qbs_Mx_to_X[OF this] assms have "g \ (f \ \) = \" by auto with h have "g \ (f \ \) \ qbs_Mx X" by simp } thus ?thesis by(auto intro!: qbs_morphismI simp: map_qbs_Mx) qed lemma pair_qbs_morphismI: assumes "\\ \. \ \ qbs_Mx X \ \ \ qbs_Mx Y \ (\r. f (\ r, \ r)) \ qbs_Mx Z" shows "f \ (X \\<^sub>Q Y) \\<^sub>Q Z" using assms by(fastforce intro!: qbs_morphismI simp: pair_qbs_Mx comp_def) lemma pair_qbs_MxD: assumes "\ \ qbs_Mx (X \\<^sub>Q Y)" obtains \ \ where "\ \ qbs_Mx X" "\ \ qbs_Mx Y" "\ = (\x. (\ x, \ x))" using assms by(auto simp: pair_qbs_Mx) lemma pair_qbs_MxI: assumes "(\x. fst (\ x)) \ qbs_Mx X" and "(\x. snd (\ x)) \ qbs_Mx Y" shows "\ \ qbs_Mx (X \\<^sub>Q Y)" using assms by(auto simp: pair_qbs_Mx comp_def) lemma shows fst_qbs_morphism: "fst \ X \\<^sub>Q Y \\<^sub>Q X" and snd_qbs_morphism: "snd \ X \\<^sub>Q Y \\<^sub>Q Y" by(auto intro!: pair_qbs_morphismI simp: comp_def) lemma qbs_morphism_pair_iff: "f \ X \\<^sub>Q Y \\<^sub>Q Z \ fst \ f \ X \\<^sub>Q Y \ snd \ f \ X \\<^sub>Q Z" by(auto intro!: qbs_morphism_comp fst_qbs_morphism snd_qbs_morphism) (auto dest: qbs_morphism_Mx intro!: qbs_morphismI simp: pair_qbs_Mx comp_assoc[symmetric]) lemma qbs_morphism_Pair: assumes "f \ Z \\<^sub>Q X" and "g \ Z \\<^sub>Q Y" shows "(\z. (f z, g z)) \ Z \\<^sub>Q X \\<^sub>Q Y" unfolding qbs_morphism_pair_iff using assms by (auto simp: comp_def) lemma qbs_morphism_curry: "curry \ exp_qbs (X \\<^sub>Q Y) Z \\<^sub>Q exp_qbs X (exp_qbs Y Z)" by(auto intro!: qbs_morphismI simp: pair_qbs_Mx exp_qbs_Mx comp_def) corollary curry_preserves_morphisms: assumes "(\xy. f (fst xy) (snd xy)) \ X \\<^sub>Q Y \\<^sub>Q Z" shows "f \ X \\<^sub>Q Y \\<^sub>Q Z" using qbs_morphism_space[OF qbs_morphism_curry assms] by (auto simp: curry_def) lemma qbs_morphism_eval: "(\fx. (fst fx) (snd fx)) \ (X \\<^sub>Q Y) \\<^sub>Q X \\<^sub>Q Y" by(auto intro!: qbs_morphismI simp: pair_qbs_Mx exp_qbs_Mx comp_def) corollary qbs_morphism_app: assumes "f \ X \\<^sub>Q (Y \\<^sub>Q Z)" "g \ X \\<^sub>Q Y" shows "(\x. (f x) (g x)) \ X \\<^sub>Q Z" by(rule qbs_morphism_cong'[where f="(\fx. (fst fx) (snd fx)) \ (\x. (f x, g x))",OF _ qbs_morphism_comp[OF qbs_morphism_Pair[OF assms] qbs_morphism_eval]]) auto ML_file \qbs.ML\ attribute_setup qbs = \ Attrib.add_del Qbs.qbs_add Qbs.qbs_del\ "declaration of qbs rule" method_setup qbs = \ Scan.lift (Scan.succeed (METHOD o Qbs.qbs_tac))\ simproc_setup qbs ("x \ qbs_space X") = \K Qbs.simproc\ declare fst_qbs_morphism[qbs] snd_qbs_morphism[qbs] qbs_morphism_const[qbs] qbs_morphism_ident[qbs] qbs_morphism_ident'[qbs] qbs_morphism_curry[qbs] lemma [qbs]: shows qbs_morphism_Pair1: "Pair \ X \\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q Y)" by(auto intro!: qbs_morphismI simp: exp_qbs_Mx pair_qbs_Mx comp_def) lemma qbs_morphism_case_prod[qbs]: "case_prod \ exp_qbs X (exp_qbs Y Z) \\<^sub>Q exp_qbs (X \\<^sub>Q Y) Z" by(fastforce intro!: qbs_morphismI simp: exp_qbs_Mx pair_qbs_Mx comp_def split_beta') lemma uncurry_preserves_morphisms: assumes [qbs]:"(\x y. f (x,y)) \ X \\<^sub>Q Y \\<^sub>Q Z" shows "f \ X \\<^sub>Q Y \\<^sub>Q Z" by(rule qbs_morphism_cong'[where f="case_prod (\x y. f (x,y))"],simp) qbs lemma qbs_morphism_comp'[qbs]:"comp \ Y \\<^sub>Q Z \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q X \\<^sub>Q Z" by(auto intro!: qbs_morphismI simp: exp_qbs_Mx) lemma arg_swap_morphism: assumes "f \ X \\<^sub>Q exp_qbs Y Z" shows "(\y x. f x y) \ Y \\<^sub>Q exp_qbs X Z" using assms by simp lemma exp_qbs_comp_morphism: assumes "f \ W \\<^sub>Q exp_qbs X Y" and "g \ W \\<^sub>Q exp_qbs Y Z" shows "(\w. g w \ f w) \ W \\<^sub>Q exp_qbs X Z" using assms by qbs lemma arg_swap_morphism_map_qbs1: assumes "g \ exp_qbs W (exp_qbs X Y) \\<^sub>Q Z" shows "(\k. g (k \ f)) \ exp_qbs (map_qbs f W) (exp_qbs X Y) \\<^sub>Q Z" using assms map_qbs_morphism_f by qbs lemma qbs_morphism_map_prod[qbs]: "map_prod \ X \\<^sub>Q Y \\<^sub>Q (W \\<^sub>Q Z) \\<^sub>Q (X \\<^sub>Q W) \\<^sub>Q (Y \\<^sub>Q Z)" by(auto intro!: qbs_morphismI simp: exp_qbs_Mx pair_qbs_Mx map_prod_def comp_def case_prod_beta') lemma qbs_morphism_pair_swap: assumes "f \ X \\<^sub>Q Y \\<^sub>Q Z" shows "(\(x,y). f (y,x)) \ Y \\<^sub>Q X \\<^sub>Q Z" using assms by simp lemma shows qbs_morphism_pair_assoc1: "(\((x,y),z). (x,(y,z))) \ (X \\<^sub>Q Y) \\<^sub>Q Z \\<^sub>Q X \\<^sub>Q (Y \\<^sub>Q Z)" and qbs_morphism_pair_assoc2: "(\(x,(y,z)). ((x,y),z)) \ X \\<^sub>Q (Y \\<^sub>Q Z) \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q Z" by simp_all lemma Inl_qbs_morphism[qbs]: "Inl \ X \\<^sub>Q X \\<^sub>Q Y" by(auto intro!: qbs_morphismI bexI[where x="{}"] simp: copair_qbs_Mx copair_qbs_Mx_def comp_def) lemma Inr_qbs_morphism[qbs]: "Inr \ Y \\<^sub>Q X \\<^sub>Q Y" by(auto intro!: qbs_morphismI bexI[where x="UNIV"] simp: copair_qbs_Mx copair_qbs_Mx_def comp_def) lemma case_sum_qbs_morphism[qbs]: "case_sum \ X \\<^sub>Q Z \\<^sub>Q (Y \\<^sub>Q Z) \\<^sub>Q (X \\<^sub>Q Y \\<^sub>Q Z)" by(auto intro!: qbs_morphismI qbs_Mx_indicat simp: copair_qbs_Mx copair_qbs_Mx_def exp_qbs_Mx case_sum_if) lemma map_sum_qbs_morphism[qbs]: "map_sum \ X \\<^sub>Q Y \\<^sub>Q (X' \\<^sub>Q Y') \\<^sub>Q (X \\<^sub>Q X' \\<^sub>Q Y \\<^sub>Q Y')" proof(rule qbs_morphismI) fix \ assume "\ \ qbs_Mx (X \\<^sub>Q Y)" then have ha[measurable]: "\(k :: real \ real)\borel_measurable borel. \a\qbs_Mx X. (\r. \ (k r) (a r)) \ qbs_Mx Y" by (auto simp: exp_qbs_Mx) show "map_sum \ \ \ qbs_Mx ((X' \\<^sub>Q Y') \\<^sub>Q X \\<^sub>Q X' \\<^sub>Q Y \\<^sub>Q Y')" unfolding exp_qbs_Mx proof safe fix \ b and f g :: "real \ real" assume h[measurable]: "\(k :: real \ real)\borel_measurable borel. \b\qbs_Mx X'. (\r. \ (k r) (b r)) \ qbs_Mx Y'" "f \ borel_measurable borel" "g \ borel_measurable borel" and b: "b \ qbs_Mx (X \\<^sub>Q X')" show "(\r. (map_sum \ \) (f (g r)) (\ (g r)) (b r)) \ qbs_Mx (Y \\<^sub>Q Y')" proof(rule copair_qbs_MxD[OF b]) fix a assume "a \ qbs_Mx X" "b = (\r. Inl (a r))" with ha show "(\r. (map_sum \ \) (f (g r)) (\ (g r)) (b r)) \ qbs_Mx (Y \\<^sub>Q Y')" by(auto simp: copair_qbs_Mx copair_qbs_Mx_def intro!: bexI[where x="{}"]) next fix a assume "a \ qbs_Mx X'" "b = (\r. Inr (a r))" with h(1) show "(\r. (map_sum \ \) (f (g r)) (\ (g r)) (b r)) \ qbs_Mx (Y \\<^sub>Q Y')" by(auto simp: copair_qbs_Mx copair_qbs_Mx_def intro!: bexI[where x="UNIV"]) next fix S a a' assume "S \ sets borel" "S \ {}" "S \ UNIV" "a \ qbs_Mx X" "a' \ qbs_Mx X'" "b = (\r. if r \ S then Inl (a r) else Inr (a' r))" with h ha show "(\r. (map_sum \ \) (f (g r)) (\ (g r)) (b r)) \ qbs_Mx (Y \\<^sub>Q Y')" by simp (fastforce simp: copair_qbs_Mx copair_qbs_Mx_def intro!: bexI[where x=S]) qed qed qed lemma qbs_morphism_component_singleton[qbs]: assumes "i \ I" shows "(\x. x i) \ (\\<^sub>Q i\I. (M i)) \\<^sub>Q M i" by(auto intro!: qbs_morphismI simp: comp_def assms PiQ_Mx) lemma qbs_morphism_component_singleton': assumes "f \ Y \\<^sub>Q (\\<^sub>Q i\I. X i)" "g \ Z \\<^sub>Q Y" "i \ I" shows "(\x. f (g x) i) \ Z \\<^sub>Q X i" by(auto intro!: qbs_morphism_compose[OF assms(2)] qbs_morphism_compose[OF assms(1)] qbs_morphism_component_singleton assms) lemma product_qbs_canonical1: assumes "\i. i \ I \ f i \ Y \\<^sub>Q X i" and "\i. i \ I \ f i = (\y. undefined)" shows "(\y i. f i y) \ Y \\<^sub>Q (\\<^sub>Q i\I. X i)" using assms qbs_morphism_Mx[OF assms(1)] by(auto intro!: qbs_morphismI simp: PiQ_Mx comp_def) lemma product_qbs_canonical2: assumes "\i. i \ I \ f i \ Y \\<^sub>Q X i" "\i. i \ I \ f i = (\y. undefined)" "g \ Y \\<^sub>Q (\\<^sub>Q i\I. X i)" "\i. i \ I \ f i = (\x. x i) \ g" and "y \ qbs_space Y" shows "g y = (\i. f i y)" proof(intro ext) fix i show "g y i = f i y" proof(cases "i \ I") case True then show ?thesis using assms(4)[of i] by simp next case False with qbs_morphism_space[OF assms(3)] assms(2,3,5) show ?thesis by(auto simp: PiQ_Mx PiQ_space) qed qed lemma merge_qbs_morphism: "merge I J \ (\\<^sub>Q i\I. (M i)) \\<^sub>Q (\\<^sub>Q j\J. (M j)) \\<^sub>Q (\\<^sub>Q i\I\J. (M i))" proof(rule qbs_morphismI) fix \ assume h:"\ \ qbs_Mx ((\\<^sub>Q i\I. (M i)) \\<^sub>Q (\\<^sub>Q j\J. (M j)))" show "merge I J \ \ \ qbs_Mx (\\<^sub>Q i\I\J. (M i))" proof - { fix i assume "i \ I \ J" then consider "i \ I" | "i \ I \ i \ J" | "i \ I \ i \ J" by auto hence "(\r. (merge I J \ \) r i) \ qbs_Mx (M i)" by cases (insert h, auto simp: merge_def split_beta' pair_qbs_Mx PiQ_Mx) } thus ?thesis by(auto simp: PiQ_Mx) (auto simp: merge_def split_beta') qed qed lemma ini_morphism[qbs]: assumes "j \ I" shows "(\x. (j,x)) \ X j \\<^sub>Q (\\<^sub>Q i\I. X i)" - by(fastforce intro!: qbs_morphismI exI[where x="\r. j"] simp: coprod_qbs_Mx_def comp_def assms coprod_qbs_Mx) + by(fastforce intro!: qbs_morphismI exI[where x="\r. j"] simp: coPiQ_Mx_def comp_def assms coPiQ_Mx) -lemma coprod_qbs_canonical1: +lemma coPiQ_canonical1: assumes "countable I" and "\i. i \ I \ f i \ X i \\<^sub>Q Y" shows "(\(i,x). f i x) \ (\\<^sub>Q i \I. X i) \\<^sub>Q Y" proof(rule qbs_morphismI) fix \ - assume "\ \ qbs_Mx (coprod_qbs I X)" + assume "\ \ qbs_Mx (coPiQ I X)" then obtain \ g where ha: "\i. i \ range g \ \ i \ qbs_Mx (X i)" "\ = (\r. (g r, \ (g r) r))" and hg[measurable]:"g \ borel \\<^sub>M count_space I" - by(fastforce simp: coprod_qbs_Mx_def coprod_qbs_Mx) + by(fastforce simp: coPiQ_Mx_def coPiQ_Mx) define f' where "f' \ (\i r. f i (\ i r))" have "range g \ I" using measurable_space[OF hg] by auto hence 1:"(\i. i \ range g \ f' i \ qbs_Mx Y)" using qbs_morphism_Mx[OF assms(2) ha(1),simplified comp_def] by(auto simp: f'_def) have "(\(i, x). f i x) \ \ = (\r. f' (g r) r)" by(auto simp: ha(2) f'_def) also have "... \ qbs_Mx Y" by(auto intro!: qbs_closed3_dest2'[OF assms(1) hg,of f',OF 1]) finally show "(\(i, x). f i x) \ \ \ qbs_Mx Y " . qed -lemma coprod_qbs_canonical1': +lemma coPiQ_canonical1': assumes "countable I" and "\i. i \ I \ (\x. f (i,x)) \ X i \\<^sub>Q Y" shows "f \ (\\<^sub>Q i \I. X i) \\<^sub>Q Y" - using coprod_qbs_canonical1[where f="curry f"] assms by(auto simp: curry_def) + using coPiQ_canonical1[where f="curry f"] assms by(auto simp: curry_def) lemma None_qbs[qbs]: "None \ qbs_space (option_qbs X)" by(simp add: option_qbs_space) lemma Some_qbs[qbs]: "Some \ X \\<^sub>Q option_qbs X" proof - have 1: "Some = (\x. case x of Inl y \ Some y | Inr y \ None) \ Inl" by standard auto show ?thesis unfolding option_qbs_def by(rule qbs_morphism_cong'[OF _ qbs_morphism_comp[OF Inl_qbs_morphism map_qbs_morphism_f]]) (simp add: 1) qed lemma case_option_qbs_morphism[qbs]: "case_option \ qbs_space (Y \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q option_qbs X \\<^sub>Q Y)" proof(rule curry_preserves_morphisms[OF arg_swap_morphism]) have "(\x y. case x of None \ fst y | Some z \ snd y z) = (\x y. case x of Inr _ \ fst y | Inl z \ snd y z) \ (\z. case z of Some x \ Inl x | None \ Inr ())" by standard+ (simp add: option.case_eq_if) also have "... \ option_qbs X \\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q Y" unfolding option_qbs_def by(rule qbs_morphism_comp[OF map_qbs_morphism_inverse_f]) (auto simp: copair_qbs_space) finally show " (\x y. case x of None \ fst y | Some x \ snd y x) \ option_qbs X \\<^sub>Q Y \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q Y" . qed lemma rec_option_qbs_morphism[qbs]: "rec_option \ qbs_space (Y \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q option_qbs X \\<^sub>Q Y)" proof - have [simp]: "rec_option = case_option" by standard+ (metis option.case_eq_if option.exhaust_sel option.simps(6) option.simps(7)) show ?thesis by simp qed lemma bind_option_qbs_morphism[qbs]: "(\) \ qbs_space (option_qbs X \\<^sub>Q (X \\<^sub>Q option_qbs Y) \\<^sub>Q option_qbs Y)" by(simp add: Option.bind_def) lemma Let_qbs_morphism[qbs]: "Let \ X \\<^sub>Q (X \\<^sub>Q Y) \\<^sub>Q Y" proof - have [simp]:"Let = (\x f. f x)" by standard+ auto show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/S_Finite_Measure_Monad/QuasiBorel.thy b/thys/S_Finite_Measure_Monad/QuasiBorel.thy --- a/thys/S_Finite_Measure_Monad/QuasiBorel.thy +++ b/thys/S_Finite_Measure_Monad/QuasiBorel.thy @@ -1,1465 +1,1492 @@ (* Title: QuasiBorel.thy Author: Michikazu Hirata, Yasuhiko Minamide Tokyo Institute of Technology *) section \Quasi-Borel Spaces\ theory QuasiBorel imports "HOL-Probability.Probability" begin subsection \ Definitions \ subsubsection \ Quasi-Borel Spaces\ definition qbs_closed1 :: "(real \ 'a) set \ bool" where "qbs_closed1 Mx \ (\a \ Mx. \f \ (borel :: real measure) \\<^sub>M (borel :: real measure). a \ f \ Mx)" definition qbs_closed2 :: "['a set, (real \ 'a) set] \ bool" where "qbs_closed2 X Mx \ (\x \ X. (\r. x) \ Mx)" definition qbs_closed3 :: "(real \ 'a) set \ bool" where "qbs_closed3 Mx \ (\P::real \ nat. \Fi::nat \ real \ 'a. (P \ borel \\<^sub>M count_space UNIV) \ (\i. Fi i \ Mx) \ (\r. Fi (P r) r) \ Mx)" lemma separate_measurable: fixes P :: "real \ nat" assumes "\i. P -` {i} \ sets borel" shows "P \ borel \\<^sub>M count_space UNIV" by (auto simp add: assms measurable_count_space_eq_countable) lemma measurable_separate: fixes P :: "real \ nat" assumes "P \ borel \\<^sub>M count_space UNIV" shows "P -` {i} \ sets borel" by (metis assms borel_singleton measurable_sets_borel sets.empty_sets sets_borel_eq_count_space) definition "is_quasi_borel X Mx \ Mx \ UNIV \ X \ qbs_closed1 Mx \ qbs_closed2 X Mx \ qbs_closed3 Mx" lemma is_quasi_borel_intro[simp]: assumes "Mx \ UNIV \ X" and "qbs_closed1 Mx" "qbs_closed2 X Mx" "qbs_closed3 Mx" shows "is_quasi_borel X Mx" using assms by(simp add: is_quasi_borel_def) typedef 'a quasi_borel = "{(X::'a set, Mx). is_quasi_borel X Mx}" proof show "(UNIV, UNIV) \ {(X::'a set, Mx). is_quasi_borel X Mx}" by (simp add: is_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def) qed definition qbs_space :: "'a quasi_borel \ 'a set" where "qbs_space X \ fst (Rep_quasi_borel X)" definition qbs_Mx :: "'a quasi_borel \ (real \ 'a) set" where "qbs_Mx X \ snd (Rep_quasi_borel X)" declare [[coercion qbs_space]] lemma qbs_decomp : "(qbs_space X,qbs_Mx X) \ {(X::'a set, Mx). is_quasi_borel X Mx}" by (simp add: qbs_space_def qbs_Mx_def Rep_quasi_borel[simplified]) lemma qbs_Mx_to_X: assumes "\ \ qbs_Mx X" shows "\ r \ qbs_space X" using qbs_decomp assms by(auto simp: is_quasi_borel_def) lemma qbs_closed1I: assumes "\\ f. \ \ Mx \ f \ borel \\<^sub>M borel \ \ \ f \ Mx" shows "qbs_closed1 Mx" using assms by(simp add: qbs_closed1_def) lemma qbs_closed1_dest[simp]: assumes "\ \ qbs_Mx X" and "f \ borel \\<^sub>M borel" shows "\ \ f \ qbs_Mx X" using assms qbs_decomp by (auto simp add: is_quasi_borel_def qbs_closed1_def) lemma qbs_closed1_dest'[simp]: assumes "\ \ qbs_Mx X" and "f \ borel \\<^sub>M borel" shows "(\r. \ (f r)) \ qbs_Mx X" using qbs_closed1_dest[OF assms] by (simp add: comp_def) lemma qbs_closed2I: assumes "\x. x \ X \ (\r. x) \ Mx" shows "qbs_closed2 X Mx" using assms by(simp add: qbs_closed2_def) lemma qbs_closed2_dest[simp]: assumes "x \ qbs_space X" shows "(\r. x) \ qbs_Mx X" using assms qbs_decomp[of X] by (auto simp add: is_quasi_borel_def qbs_closed2_def) lemma qbs_closed3I: assumes "\(P :: real \ nat) Fi. P \ borel \\<^sub>M count_space UNIV \ (\i. Fi i \ Mx) \ (\r. Fi (P r) r) \ Mx" shows "qbs_closed3 Mx" using assms by(auto simp: qbs_closed3_def) lemma qbs_closed3I': assumes "\(P :: real \ nat) Fi. (\i. P -` {i} \ sets borel) \ (\i. Fi i \ Mx) \ (\r. Fi (P r) r) \ Mx" shows "qbs_closed3 Mx" using assms by(auto intro!: qbs_closed3I dest: measurable_separate) lemma qbs_closed3_dest[simp]: fixes P::"real \ nat" and Fi :: "nat \ real \ _" assumes "P \ borel \\<^sub>M count_space UNIV" and "\i. Fi i \ qbs_Mx X" shows "(\r. Fi (P r) r) \ qbs_Mx X" using assms qbs_decomp[of X] by (auto simp add: is_quasi_borel_def qbs_closed3_def) lemma qbs_closed3_dest': fixes P::"real \ nat" and Fi :: "nat \ real \ _" assumes "\i. P -` {i} \ sets borel" and "\i. Fi i \ qbs_Mx X" shows "(\r. Fi (P r) r) \ qbs_Mx X" using qbs_closed3_dest[OF separate_measurable[OF assms(1)] assms(2)] . lemma qbs_closed3_dest2: assumes "countable I" and [measurable]: "P \ borel \\<^sub>M count_space I" and "\i. i \ I \ Fi i \ qbs_Mx X" shows "(\r. Fi (P r) r) \ qbs_Mx X" proof - have 0:"I \ {}" using measurable_empty_iff[of "count_space I" P borel] assms(2) by fastforce define P' where "P' \ to_nat_on I \ P" define Fi' where "Fi' \ Fi \ (from_nat_into I)" have 1:"P' \ borel \\<^sub>M count_space UNIV" by(simp add: P'_def) have 2:"\i. Fi' i \ qbs_Mx X" using assms(3) from_nat_into[OF 0] by(simp add: Fi'_def) have "(\r. Fi' (P' r) r) \ qbs_Mx X" using 1 2 measurable_separate by auto thus ?thesis using from_nat_into_to_nat_on[OF assms(1)] measurable_space[OF assms(2)] by(auto simp: Fi'_def P'_def) qed lemma qbs_closed3_dest2': assumes "countable I" and [measurable]: "P \ borel \\<^sub>M count_space I" and "\i. i \ range P \ Fi i \ qbs_Mx X" shows "(\r. Fi (P r) r) \ qbs_Mx X" proof - have 0:"range P \ I = range P" using measurable_space[OF assms(2)] by auto have 1:"P \ borel \\<^sub>M count_space (range P)" using restrict_count_space[of I "range P"] measurable_restrict_space2[OF _ assms(2),of "range P"] by(simp add: 0) have 2:"countable (range P)" using countable_Int2[OF assms(1),of "range P"] by(simp add: 0) show ?thesis by(auto intro!: qbs_closed3_dest2[OF 2 1 assms(3)]) qed lemma qbs_Mx_indicat: assumes "S \ sets borel" "\ \ qbs_Mx X" "\ \ qbs_Mx X" shows "(\r. if r \ S then \ r else \ r) \ qbs_Mx X" proof - have "(\r::real. if r \ S then \ r else \ r) = (\r. (\b. if b then \ else \) (r \ S) r)" by(auto simp: indicator_def) also have "... \ qbs_Mx X" by(rule qbs_closed3_dest2[where I=UNIV and Fi="\b. if b then \ else \"]) (use assms in auto) finally show ?thesis . qed lemma qbs_space_Mx: "qbs_space X = {\ x |x \. \ \ qbs_Mx X}" proof safe fix x assume 1:"x \ qbs_space X" show "\xa \. x = \ xa \ \ \ qbs_Mx X" by(auto intro!: exI[where x=0] exI[where x="(\r. x)"] simp: 1) qed(simp add: qbs_Mx_to_X) lemma qbs_space_eq_Mx: assumes "qbs_Mx X = qbs_Mx Y" shows "qbs_space X = qbs_space Y" by(simp add: qbs_space_Mx assms) lemma qbs_eqI: assumes "qbs_Mx X = qbs_Mx Y" shows "X = Y" by (metis Rep_quasi_borel_inverse prod.exhaust_sel qbs_Mx_def qbs_space_def assms qbs_space_eq_Mx[OF assms]) subsubsection \ Empty Space \ definition empty_quasi_borel :: "'a quasi_borel" where "empty_quasi_borel \ Abs_quasi_borel ({},{})" lemma shows eqb_space[simp]: "qbs_space empty_quasi_borel = ({} :: 'a set)" and eqb_Mx[simp]: "qbs_Mx empty_quasi_borel = ({} :: (real \ 'a) set)" proof - have "Rep_quasi_borel empty_quasi_borel = ({} :: 'a set, {})" using Abs_quasi_borel_inverse by(auto simp add: Abs_quasi_borel_inverse empty_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def is_quasi_borel_def) thus "qbs_space empty_quasi_borel = ({} :: 'a set)" "qbs_Mx empty_quasi_borel = ({} :: (real \ 'a) set)" by(auto simp add: qbs_space_def qbs_Mx_def) qed lemma qbs_empty_equiv :"qbs_space X = {} \ qbs_Mx X = {}" proof safe fix x assume "qbs_Mx X = {}" and h:"x \ qbs_space X" have "(\r. x) \ qbs_Mx X" using h by simp thus "x \ {}" using \qbs_Mx X = {}\ by simp qed(use qbs_Mx_to_X in blast) lemma empty_quasi_borel_iff: "qbs_space X = {} \ X = empty_quasi_borel" by(auto intro!: qbs_eqI simp: qbs_empty_equiv) subsubsection \ Unit Space \ definition unit_quasi_borel :: "unit quasi_borel" ("1\<^sub>Q") where "unit_quasi_borel \ Abs_quasi_borel (UNIV,UNIV)" lemma shows unit_qbs_space[simp]: "qbs_space unit_quasi_borel = {()}" and unit_qbs_Mx[simp]: "qbs_Mx unit_quasi_borel = {\r. ()}" proof - have "Rep_quasi_borel unit_quasi_borel = (UNIV,UNIV)" using Abs_quasi_borel_inverse by(auto simp add: unit_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def is_quasi_borel_def) thus "qbs_space unit_quasi_borel = {()}" "qbs_Mx unit_quasi_borel = {\r. ()}" by(auto simp add: qbs_space_def qbs_Mx_def UNIV_unit) qed subsubsection \ Sub-Spaces \ definition sub_qbs :: "['a quasi_borel, 'a set] \ 'a quasi_borel" where "sub_qbs X U \ Abs_quasi_borel (qbs_space X \ U,{\. \ \ qbs_Mx X \ (\r. \ r \ U)})" lemma shows sub_qbs_space: "qbs_space (sub_qbs X U) = qbs_space X \ U" and sub_qbs_Mx: "qbs_Mx (sub_qbs X U) = {\. \ \ qbs_Mx X \ (\r. \ r \ U)}" proof - have "qbs_closed1 {\. \ \ qbs_Mx X \ (\r. \ r \ U)}" "qbs_closed2 (qbs_space X \ U) {\. \ \ qbs_Mx X \ (\r. \ r \ U)}" "qbs_closed3 {\. \ \ qbs_Mx X \ (\r. \ r \ U)}" unfolding qbs_closed1_def qbs_closed2_def qbs_closed3_def by auto hence "Rep_quasi_borel (sub_qbs X U) = (qbs_space X \ U,{\. \ \ qbs_Mx X \ (\r. \ r \ U)})" by(auto simp: sub_qbs_def is_quasi_borel_def qbs_Mx_to_X intro!: Abs_quasi_borel_inverse) thus "qbs_space (sub_qbs X U) = qbs_space X \ U" "qbs_Mx (sub_qbs X U) = {\. \ \ qbs_Mx X \ (\r. \ r \ U)}" by(simp_all add: qbs_Mx_def qbs_space_def) qed lemma sub_qbs: assumes "U \ qbs_space X" shows "(qbs_space (sub_qbs X U), qbs_Mx (sub_qbs X U)) = (U, {f \ UNIV \ U. f \ qbs_Mx X})" using assms by (auto simp: sub_qbs_space sub_qbs_Mx) lemma sub_qbs_ident: "sub_qbs X (qbs_space X) = X" by(auto intro!: qbs_eqI simp: sub_qbs_Mx qbs_Mx_to_X) lemma sub_qbs_sub_qbs: "sub_qbs (sub_qbs X A) B = sub_qbs X (A \ B)" by(auto intro!: qbs_eqI simp: sub_qbs_Mx sub_qbs_space) subsubsection \ Image Spaces \ definition map_qbs :: "['a \ 'b] \ 'a quasi_borel \ 'b quasi_borel" where "map_qbs f X = Abs_quasi_borel (f ` (qbs_space X),{f \ \ |\. \\ qbs_Mx X})" lemma shows map_qbs_space: "qbs_space (map_qbs f X) = f ` (qbs_space X)" and map_qbs_Mx: "qbs_Mx (map_qbs f X) = {f \ \ |\. \\ qbs_Mx X}" proof - have "{f \ \ |\. \\ qbs_Mx X} \ UNIV \ f ` (qbs_space X)" using qbs_Mx_to_X by fastforce moreover have "qbs_closed1 {f \ \ |\. \\ qbs_Mx X}" unfolding qbs_closed1_def using qbs_closed1_dest by(fastforce simp: comp_def) moreover have "qbs_closed2 (f ` (qbs_space X)) {f \ \ |\. \\ qbs_Mx X}" unfolding qbs_closed2_def by fastforce moreover have "qbs_closed3 {f \ \ |\. \\ qbs_Mx X}" proof(rule qbs_closed3I') fix P :: "real \ nat" and Fi assume h:"\i::nat. P -` {i} \ sets borel" "\i::nat. Fi i \ {f \ \ |\. \\ qbs_Mx X}" then obtain \i where ha: "\i::nat. \i i \ qbs_Mx X" "\i. Fi i = f \ (\i i)" by auto metis hence 1:"(\r. \i (P r) r) \ qbs_Mx X" using h(1) qbs_closed3_dest' by blast show "(\r. Fi (P r) r) \ {f \ \ |\. \\ qbs_Mx X}" by(auto intro!: bexI[where x="(\r. \i (P r) r)"] simp add: 1 ha comp_def) qed ultimately have "Rep_quasi_borel (map_qbs f X) = (f ` (qbs_space X),{f \ \ |\. \\ qbs_Mx X})" unfolding map_qbs_def by(auto intro!: Abs_quasi_borel_inverse) thus "qbs_space (map_qbs f X) = f ` (qbs_space X)" "qbs_Mx (map_qbs f X) = {f \ \ |\. \\ qbs_Mx X}" by(simp_all add: qbs_space_def qbs_Mx_def) qed subsubsection \ Binary Product Spaces \ definition pair_qbs :: "['a quasi_borel, 'b quasi_borel] \ ('a \ 'b) quasi_borel" (infixr "\\<^sub>Q" 80) where "pair_qbs X Y = Abs_quasi_borel (qbs_space X \ qbs_space Y, {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y})" lemma shows pair_qbs_space: "qbs_space (X \\<^sub>Q Y) = qbs_space X \ qbs_space Y" and pair_qbs_Mx: "qbs_Mx (X \\<^sub>Q Y) = {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" proof - have "{f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y} \ UNIV \ qbs_space X \ qbs_space Y" by (auto simp: mem_Times_iff[of _ "qbs_space X" "qbs_space Y"]; use qbs_Mx_to_X in fastforce) moreover have "qbs_closed1 {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" unfolding qbs_closed1_def by (metis (no_types, lifting) comp_assoc mem_Collect_eq qbs_closed1_dest) moreover have "qbs_closed2 (qbs_space X \ qbs_space Y) {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" unfolding qbs_closed2_def by auto moreover have "qbs_closed3 {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" proof(safe intro!: qbs_closed3I) fix P :: "real \ nat" fix Fi :: "nat \ real \ 'a \ 'b" define Fj :: "nat \ real \ 'a" where "Fj \ \j.(fst \ Fi j)" assume "\i. Fi i \ {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" then have "\i. Fj i \ qbs_Mx X" by (simp add: Fj_def) moreover assume "P \ borel \\<^sub>M count_space UNIV" ultimately have "(\r. Fj (P r) r) \ qbs_Mx X" by auto moreover have "fst \ (\r. Fi (P r) r) = (\r. Fj (P r) r)" by (auto simp add: Fj_def) ultimately show "fst \ (\r. Fi (P r) r) \ qbs_Mx X" by simp next fix P :: "real \ nat" fix Fi :: "nat \ real \ 'a \ 'b" define Fj :: "nat \ real \ 'b" where "Fj \ \j.(snd \ Fi j)" assume "\i. Fi i \ {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" then have "\i. Fj i \ qbs_Mx Y" by (simp add: Fj_def) moreover assume "P \ borel \\<^sub>M count_space UNIV" ultimately have "(\r. Fj (P r) r) \ qbs_Mx Y" by auto moreover have "snd \ (\r. Fi (P r) r) = (\r. Fj (P r) r)" by (auto simp add: Fj_def) ultimately show "snd \ (\r. Fi (P r) r) \ qbs_Mx Y" by simp qed ultimately have "Rep_quasi_borel (X \\<^sub>Q Y) = (qbs_space X \ qbs_space Y, {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y})" unfolding pair_qbs_def by(auto intro!: Abs_quasi_borel_inverse is_quasi_borel_intro) thus "qbs_space (X \\<^sub>Q Y) = qbs_space X \ qbs_space Y" "qbs_Mx (X \\<^sub>Q Y) = {f. fst \ f \ qbs_Mx X \ snd \ f \ qbs_Mx Y}" by(simp_all add: qbs_space_def qbs_Mx_def) qed lemma pair_qbs_fst: assumes "qbs_space Y \ {}" shows "map_qbs fst (X \\<^sub>Q Y) = X" proof(rule qbs_eqI) obtain \y where hy:"\y \ qbs_Mx Y" using qbs_empty_equiv[of Y] assms by auto show "qbs_Mx (map_qbs fst (X \\<^sub>Q Y)) = qbs_Mx X" by(auto simp: map_qbs_Mx pair_qbs_Mx hy comp_def intro!: exI[where x="\r. (_ r, \y r)"]) qed lemma pair_qbs_snd: assumes "qbs_space X \ {}" shows "map_qbs snd (X \\<^sub>Q Y) = Y" proof(rule qbs_eqI) obtain \x where hx:"\x \ qbs_Mx X" using qbs_empty_equiv[of X] assms by auto show "qbs_Mx (map_qbs snd (X \\<^sub>Q Y)) = qbs_Mx Y" by(auto simp: map_qbs_Mx pair_qbs_Mx hx comp_def intro!: exI[where x="\r. (\x r, _ r)"]) qed subsubsection \ Binary Coproduct Spaces \ definition copair_qbs_Mx :: "['a quasi_borel, 'b quasi_borel] \ (real => 'a + 'b) set" where "copair_qbs_Mx X Y \ {g. \ S \ sets borel. (S = {} \ (\ \1\ qbs_Mx X. g = (\r. Inl (\1 r)))) \ (S = UNIV \ (\ \2\ qbs_Mx Y. g = (\r. Inr (\2 r)))) \ ((S \ {} \ S \ UNIV) \ (\ \1\ qbs_Mx X. \ \2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))))}" - definition copair_qbs :: "['a quasi_borel, 'b quasi_borel] \ ('a + 'b) quasi_borel" (infixr "\\<^sub>Q" 65) where "copair_qbs X Y \ Abs_quasi_borel (qbs_space X <+> qbs_space Y, copair_qbs_Mx X Y)" - text \ The following is an equivalent definition of @{term copair_qbs_Mx}. \ definition copair_qbs_Mx2 :: "['a quasi_borel, 'b quasi_borel] \ (real => 'a + 'b) set" where "copair_qbs_Mx2 X Y \ {g. (if qbs_space X = {} \ qbs_space Y = {} then False else if qbs_space X \ {} \ qbs_space Y = {} then (\\1\ qbs_Mx X. g = (\r. Inl (\1 r))) else if qbs_space X = {} \ qbs_space Y \ {} then (\\2\ qbs_Mx Y. g = (\r. Inr (\2 r))) else (\S \ sets borel. \\1\ qbs_Mx X. \\2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r))))) }" lemma copair_qbs_Mx_equiv :"copair_qbs_Mx (X :: 'a quasi_borel) (Y :: 'b quasi_borel) = copair_qbs_Mx2 X Y" proof safe (* \ *) fix g :: "real \ 'a + 'b" assume "g \ copair_qbs_Mx X Y" then obtain S where hs:"S\ sets borel \ (S = {} \ (\ \1\ qbs_Mx X. g = (\r. Inl (\1 r)))) \ (S = UNIV \ (\ \2\ qbs_Mx Y. g = (\r. Inr (\2 r)))) \ ((S \ {} \ S \ UNIV) \ (\ \1\ qbs_Mx X. \ \2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))))" by (auto simp add: copair_qbs_Mx_def) consider "S = {}" | "S = UNIV" | "S \ {} \ S \ UNIV" by auto then show "g \ copair_qbs_Mx2 X Y" proof cases assume "S = {}" from hs this have "\ \1\ qbs_Mx X. g = (\r. Inl (\1 r))" by simp then obtain \1 where h1:"\1\ qbs_Mx X \ g = (\r. Inl (\1 r))" by auto have "qbs_space X \ {}" using qbs_empty_equiv h1 by auto then have "(qbs_space X \ {} \ qbs_space Y = {}) \ (qbs_space X \ {} \ qbs_space Y \ {})" by simp then show "g \ copair_qbs_Mx2 X Y" proof assume "qbs_space X \ {} \ qbs_space Y = {}" then show "g \ copair_qbs_Mx2 X Y" by(simp add: copair_qbs_Mx2_def \\ \1\ qbs_Mx X. g = (\r. Inl (\1 r))\) next assume "qbs_space X \ {} \ qbs_space Y \ {}" then obtain \2 where "\2 \ qbs_Mx Y" using qbs_empty_equiv by force define S' :: "real set" where "S' \ UNIV" define g' :: "real \ 'a + 'b" where "g' \ (\r::real. (if (r \ S') then Inl (\1 r) else Inr (\2 r)))" from \qbs_space X \ {} \ qbs_space Y \ {}\ h1 \\2 \ qbs_Mx Y\ have "g' \ copair_qbs_Mx2 X Y" by(force simp add: S'_def g'_def copair_qbs_Mx2_def) moreover have "g = g'" using h1 by(simp add: g'_def S'_def) ultimately show ?thesis by simp qed next assume "S = UNIV" from hs this have "\ \2\ qbs_Mx Y. g = (\r. Inr (\2 r))" by simp then obtain \2 where h2:"\2\ qbs_Mx Y \ g = (\r. Inr (\2 r))" by auto have "qbs_space Y \ {}" using qbs_empty_equiv h2 by auto then have "(qbs_space X = {} \ qbs_space Y \ {}) \ (qbs_space X \ {} \ qbs_space Y \ {})" by simp then show "g \ copair_qbs_Mx2 X Y" proof assume "qbs_space X = {} \ qbs_space Y \ {}" then show ?thesis by(simp add: copair_qbs_Mx2_def \\ \2\ qbs_Mx Y. g = (\r. Inr (\2 r))\) next assume "qbs_space X \ {} \ qbs_space Y \ {}" then obtain \1 where "\1 \ qbs_Mx X" using qbs_empty_equiv by force define S' :: "real set" where "S' \ {}" define g' :: "real \ 'a + 'b" where "g' \ (\r::real. (if (r \ S') then Inl (\1 r) else Inr (\2 r)))" from \qbs_space X \ {} \ qbs_space Y \ {}\ h2 \\1 \ qbs_Mx X\ have "g' \ copair_qbs_Mx2 X Y" by(force simp add: S'_def g'_def copair_qbs_Mx2_def) moreover have "g = g'" using h2 by(simp add: g'_def S'_def) ultimately show ?thesis by simp qed next assume "S \ {} \ S \ UNIV" then have h: "\ \1\ qbs_Mx X. \ \2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))" using hs by simp then have "qbs_space X \ {} \ qbs_space Y \ {}" by (metis empty_iff qbs_empty_equiv) thus ?thesis using hs h by(auto simp add: copair_qbs_Mx2_def) qed (* \ *) next fix g :: "real \ 'a + 'b" assume "g \ copair_qbs_Mx2 X Y" then have h: "if qbs_space X = {} \ qbs_space Y = {} then False else if qbs_space X \ {} \ qbs_space Y = {} then (\\1\ qbs_Mx X. g = (\r. Inl (\1 r))) else if qbs_space X = {} \ qbs_space Y \ {} then (\\2\ qbs_Mx Y. g = (\r. Inr (\2 r))) else (\S \ sets borel. \\1\ qbs_Mx X. \\2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r))))" by(simp add: copair_qbs_Mx2_def) consider "(qbs_space X = {} \ qbs_space Y = {})" | "(qbs_space X \ {} \ qbs_space Y = {})" | "(qbs_space X = {} \ qbs_space Y \ {})" | "(qbs_space X \ {} \ qbs_space Y \ {})" by auto then show "g \ copair_qbs_Mx X Y" proof cases assume "qbs_space X = {} \ qbs_space Y = {}" then show ?thesis using \g \ copair_qbs_Mx2 X Y\ by(simp add: copair_qbs_Mx2_def) next assume "qbs_space X \ {} \ qbs_space Y = {}" from h this have "\\1\ qbs_Mx X. g = (\r. Inl (\1 r))" by simp thus ?thesis by(auto simp add: copair_qbs_Mx_def) next assume "qbs_space X = {} \ qbs_space Y \ {}" from h this have "\\2\ qbs_Mx Y. g = (\r. Inr (\2 r))" by simp thus ?thesis unfolding copair_qbs_Mx_def by(force simp add: copair_qbs_Mx_def) next assume "qbs_space X \ {} \ qbs_space Y \ {}" from h this obtain S \1 \2 where Sag: "S \ sets borel" "\1 \ qbs_Mx X" "\2 \ qbs_Mx Y" "g = (\r. if r \ S then Inl (\1 r) else Inr (\2 r))" by auto consider "S = {}" | "S = UNIV" | "S \ {}" "S \ UNIV" by auto then show "g \ copair_qbs_Mx X Y" proof cases assume "S = {}" then have [simp]: "(\r. if r \ S then Inl (\1 r) else Inr (\2 r)) = (\r. Inr (\2 r))" by simp show ?thesis using \\2 \ qbs_Mx Y\ unfolding copair_qbs_Mx_def by(auto intro! : bexI[where x=UNIV] simp: Sag) next assume "S = UNIV" then have "(\r. if r \ S then Inl (\1 r) else Inr (\2 r)) = (\r. Inl (\1 r))" by simp then show ?thesis using Sag by(auto simp add: copair_qbs_Mx_def) next assume "S \ {}" "S \ UNIV" then show ?thesis using Sag by(auto simp add: copair_qbs_Mx_def) qed qed qed lemma shows copair_qbs_space: "qbs_space (X \\<^sub>Q Y) = qbs_space X <+> qbs_space Y" (is ?goal1) and copair_qbs_Mx: "qbs_Mx (X \\<^sub>Q Y) = copair_qbs_Mx X Y" (is ?goal2) proof - have "copair_qbs_Mx X Y \ UNIV \ qbs_space X <+> qbs_space Y" proof fix g assume "g \ copair_qbs_Mx X Y" then obtain S where hs:"S\ sets borel \ (S = {} \ (\ \1\ qbs_Mx X. g = (\r. Inl (\1 r)))) \ (S = UNIV \ (\ \2\ qbs_Mx Y. g = (\r. Inr (\2 r)))) \ ((S \ {} \ S \ UNIV) \ (\ \1\ qbs_Mx X. \ \2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))))" by (auto simp add: copair_qbs_Mx_def) consider "S = {}" | "S = UNIV" | "S \ {} \ S \ UNIV" by auto then show "g \ UNIV \ qbs_space X <+> qbs_space Y" proof cases assume "S = {}" then show ?thesis using hs qbs_Mx_to_X by auto next assume "S = UNIV" then show ?thesis using hs qbs_Mx_to_X by auto next assume "S \ {} \ S \ UNIV" then have "\ \1\ qbs_Mx X. \ \2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))" using hs by simp then show ?thesis by(auto dest: qbs_Mx_to_X) qed qed moreover have "qbs_closed1 (copair_qbs_Mx X Y)" proof(rule qbs_closed1I) fix g and f :: "real \ real" assume "g \ copair_qbs_Mx X Y" and [measurable]: "f \ borel \\<^sub>M borel" then have "g \ copair_qbs_Mx2 X Y" using copair_qbs_Mx_equiv by auto consider "(qbs_space X = {} \ qbs_space Y = {})" | "(qbs_space X \ {} \ qbs_space Y = {})" | "(qbs_space X = {} \ qbs_space Y \ {})" | "(qbs_space X \ {} \ qbs_space Y \ {})" by auto then have "g \ f \ copair_qbs_Mx2 X Y" proof cases assume "qbs_space X = {} \ qbs_space Y = {}" then show ?thesis using \g \ copair_qbs_Mx2 X Y\ qbs_empty_equiv by(simp add: copair_qbs_Mx2_def) next assume "qbs_space X \ {} \ qbs_space Y = {}" then obtain \1 where h1:"\1\ qbs_Mx X \ g = (\r. Inl (\1 r))" using \g \ copair_qbs_Mx2 X Y\ by(auto simp add: copair_qbs_Mx2_def) then have "\1 \ f \ qbs_Mx X" by auto moreover have "g \ f = (\r. Inl ((\1 \ f) r))" using h1 by auto ultimately show ?thesis using \qbs_space X \ {} \ qbs_space Y = {}\ by(force simp add: copair_qbs_Mx2_def) next assume "(qbs_space X = {} \ qbs_space Y \ {})" then obtain \2 where h2:"\2\ qbs_Mx Y \ g = (\r. Inr (\2 r))" using \g \ copair_qbs_Mx2 X Y\ by(auto simp add: copair_qbs_Mx2_def) then have "\2 \ f \ qbs_Mx Y" by auto moreover have "g \ f = (\r. Inr ((\2 \ f) r))" using h2 by auto ultimately show ?thesis using \(qbs_space X = {} \ qbs_space Y \ {})\ by(force simp add: copair_qbs_Mx2_def) next assume "qbs_space X \ {} \ qbs_space Y \ {}" then have "\S \ sets borel. \\1\ qbs_Mx X. \\2\ qbs_Mx Y. g = (\r::real. (if (r \ S) then Inl (\1 r) else Inr (\2 r)))" using \g \ copair_qbs_Mx2 X Y\ by(simp add: copair_qbs_Mx2_def) then show ?thesis proof safe fix S \1 \2 assume [measurable]:"S \ sets borel" and "\1\ qbs_Mx X" "\2 \ qbs_Mx Y" "g = (\r. if r \ S then Inl (\1 r) else Inr (\2 r))" have "f -` S \ sets borel" using \S \ sets borel\ \f \ borel_measurable borel\ measurable_sets_borel by blast moreover have "\1 \ f \ qbs_Mx X" using \\1\ qbs_Mx X\ by(auto simp add: qbs_closed1_def) moreover have "\2 \ f \ qbs_Mx Y" using \\2\ qbs_Mx Y\ by(auto simp add: qbs_closed1_def) moreover have "(\r. if r \ S then Inl (\1 r) else Inr (\2 r)) \ f = (\r. if r \ f -` S then Inl ((\1 \ f) r) else Inr ((\2 \ f) r))" by auto ultimately show "(\r. if r \ S then Inl (\1 r) else Inr (\2 r)) \ f \ copair_qbs_Mx2 X Y" using \qbs_space X \ {} \ qbs_space Y \ {}\ by(force simp add: copair_qbs_Mx2_def) qed qed thus "g \ f \ copair_qbs_Mx X Y" using copair_qbs_Mx_equiv by auto qed moreover have "qbs_closed2 (qbs_space X <+> qbs_space Y) (copair_qbs_Mx X Y)" proof(rule qbs_closed2I) fix y assume "y \ qbs_space X <+> qbs_space Y" then consider "y \ Inl ` (qbs_space X)" | "y \ Inr ` (qbs_space Y)" by auto thus "(\r. y) \ copair_qbs_Mx X Y" proof cases case 1 then obtain x where x: "y = Inl x" "x \ qbs_space X" by auto define \1 :: "real \ _" where "\1 \ (\r. x)" have "\1 \ qbs_Mx X" using \x \ qbs_space X\ qbs_decomp by(force simp add: qbs_closed2_def \1_def) moreover have "(\r. Inl x) = (\l. Inl (\1 l))" by (simp add: \1_def) moreover have "{} \ sets borel" by auto ultimately show "(\r. y) \ copair_qbs_Mx X Y" by(auto simp add: copair_qbs_Mx_def x) next case 2 then obtain x where x: "y = Inr x" "x \ qbs_space Y" by auto define \2 :: "real \ _" where "\2 \ (\r. x)" have "\2 \ qbs_Mx Y" using \x \ qbs_space Y\ qbs_decomp by(force simp add: qbs_closed2_def \2_def) moreover have "(\r. Inr x) = (\l. Inr (\2 l))" by (simp add: \2_def) moreover have "UNIV \ sets borel" by auto ultimately show "(\r. y) \ copair_qbs_Mx X Y" unfolding copair_qbs_Mx_def by(auto intro!: bexI[where x=UNIV] simp: x) qed qed moreover have "qbs_closed3 (copair_qbs_Mx X Y)" proof(safe intro!: qbs_closed3I) fix P :: "real \ nat" fix Fi :: "nat \ real \_ + _" assume "P \ borel \\<^sub>M count_space UNIV" "\i. Fi i \ copair_qbs_Mx X Y" then have "\i. Fi i \ copair_qbs_Mx2 X Y" using copair_qbs_Mx_equiv by blast consider "(qbs_space X = {} \ qbs_space Y = {})" | "(qbs_space X \ {} \ qbs_space Y = {})" | "(qbs_space X = {} \ qbs_space Y \ {})" | "(qbs_space X \ {} \ qbs_space Y \ {})" by auto then have "(\r. Fi (P r) r) \ copair_qbs_Mx2 X Y" proof cases assume "qbs_space X = {} \ qbs_space Y = {}" then show ?thesis using \\i. Fi i \ copair_qbs_Mx2 X Y\ qbs_empty_equiv by(simp add: copair_qbs_Mx2_def) next assume "qbs_space X \ {} \ qbs_space Y = {}" then have "\i. \\i. \i \ qbs_Mx X \ Fi i = (\r. Inl (\i r))" using \\i. Fi i \ copair_qbs_Mx2 X Y\ by(auto simp add: copair_qbs_Mx2_def) then have "\\1. \i. \1 i \ qbs_Mx X \ Fi i = (\r. Inl (\1 i r))" by(rule choice) then obtain \1 :: "nat \ real \ _" where h1: "\i. \1 i \ qbs_Mx X \ Fi i = (\r. Inl (\1 i r))" by auto define \ :: "real \ _" where "\ \ (\r. \1 (P r) r)" from \P \ borel \\<^sub>M count_space UNIV\ h1 have "\ \ qbs_Mx X" by (simp add: \_def) moreover have "(\r. Fi (P r) r) = (\r. Inl (\ r))" using h1 by(simp add: \_def) ultimately show ?thesis using \qbs_space X \ {} \ qbs_space Y = {}\ by (auto simp add: copair_qbs_Mx2_def) next assume "qbs_space X = {} \ qbs_space Y \ {}" then have "\i. \\i. \i \ qbs_Mx Y \ Fi i = (\r. Inr (\i r))" using \\i. Fi i \ copair_qbs_Mx2 X Y\ by(auto simp add: copair_qbs_Mx2_def) then have "\\2. \i. \2 i \ qbs_Mx Y \ Fi i = (\r. Inr (\2 i r))" by(rule choice) then obtain \2 :: "nat \ real \ _" where h2: "\i. \2 i \ qbs_Mx Y \ Fi i = (\r. Inr (\2 i r))" by auto define \ :: "real \ _" where "\ \ (\r. \2 (P r) r)" from \P \ borel \\<^sub>M count_space UNIV\ h2 have "\ \ qbs_Mx Y" by(simp add: \_def) moreover have "(\r. Fi (P r) r) = (\r. Inr (\ r))" using h2 by(simp add: \_def) ultimately show ?thesis using \qbs_space X = {} \ qbs_space Y \ {}\ by (auto simp add: copair_qbs_Mx2_def) next assume "qbs_space X \ {} \ qbs_space Y \ {}" then have "\i. \Si. Si \ sets borel \ (\\1i\ qbs_Mx X. \\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ Si) then Inl (\1i r) else Inr (\2i r))))" using \\i. Fi i \ copair_qbs_Mx2 X Y\ by (auto simp add: copair_qbs_Mx2_def) then have "\S. \i. S i \ sets borel \ (\\1i\ qbs_Mx X. \\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ S i) then Inl (\1i r) else Inr (\2i r))))" by(rule choice) then obtain S :: "nat \ real set" where hs :"\i. S i \ sets borel \ (\\1i\ qbs_Mx X. \\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ S i) then Inl (\1i r) else Inr (\2i r))))" by auto then have "\i. \\1i. \1i \ qbs_Mx X \ (\\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ S i) then Inl (\1i r) else Inr (\2i r))))" by blast then have "\\1. \i. \1 i \ qbs_Mx X \ (\\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ S i) then Inl (\1 i r) else Inr (\2i r))))" by(rule choice) then obtain \1 where h1: "\i. \1 i \ qbs_Mx X \ (\\2i\ qbs_Mx Y. Fi i = (\r::real. (if (r \ S i) then Inl (\1 i r) else Inr (\2i r))))" by auto define \1 :: "real \ _" where "\1 \ (\r. \1 (P r) r)" from \P \ borel \\<^sub>M count_space UNIV\ h1 have "\1 \ qbs_Mx X" by(simp add: \1_def) from h1 have "\i. \\2i. \2i\ qbs_Mx Y \ Fi i = (\r::real. (if (r \ S i) then Inl (\1 i r) else Inr (\2i r)))" by auto then have "\\2. \i. \2 i\ qbs_Mx Y \ Fi i = (\r::real. (if (r \ S i) then Inl (\1 i r) else Inr (\2 i r)))" by(rule choice) then obtain \2 where h2: "\i. \2 i\ qbs_Mx Y \ Fi i = (\r::real. (if (r \ S i) then Inl (\1 i r) else Inr (\2 i r)))" by auto define \2 :: "real \ _" where "\2 \ (\r. \2 (P r) r)" from \P \ borel \\<^sub>M count_space UNIV\ h2 have "\2 \ qbs_Mx Y" by(simp add: \2_def) define A :: "nat \ real set" where "A \ (\i. S i \ P -` {i})" have [measurable]:"\i. A i \ sets borel" using A_def hs measurable_separate[OF \P \ borel \\<^sub>M count_space UNIV\] by blast define S' :: "real set" where "S' \ {r. r \ S (P r)}" have "S' = (\i::nat. A i)" by(auto simp add: S'_def A_def) hence "S' \ sets borel" by auto from h2 have "(\r. Fi (P r) r) = (\r. (if r \ S' then Inl (\1 r) else Inr (\2 r)))" by(auto simp add: \1_def \2_def S'_def) thus "(\r. Fi (P r) r) \ copair_qbs_Mx2 X Y" using \qbs_space X \ {} \ qbs_space Y \ {}\ \S' \ sets borel\ \\1 \ qbs_Mx X\ \\2 \ qbs_Mx Y\ by(auto simp add: copair_qbs_Mx2_def) qed thus "(\r. Fi (P r) r) \ copair_qbs_Mx X Y" using copair_qbs_Mx_equiv by auto qed ultimately have "Rep_quasi_borel (copair_qbs X Y) = (qbs_space X <+> qbs_space Y, copair_qbs_Mx X Y)" unfolding copair_qbs_def by(auto intro!: Abs_quasi_borel_inverse) thus ?goal1 ?goal2 by(simp_all add: qbs_space_def qbs_Mx_def) qed lemma copair_qbs_MxD: assumes "g \ qbs_Mx (X \\<^sub>Q Y)" and "\\. \ \ qbs_Mx X \ g = (\r. Inl (\ r)) \ P g" and "\\. \ \ qbs_Mx Y \ g = (\r. Inr (\ r)) \ P g" and "\S \ \. (S :: real set) \ sets borel \ S \ {} \ S \ UNIV \ \ \ qbs_Mx X \ \ \ qbs_Mx Y \ g = (\r. if r \ S then Inl (\ r) else Inr (\ r)) \ P g" shows "P g" using assms by(fastforce simp: copair_qbs_Mx copair_qbs_Mx_def) subsubsection \ Product Spaces \ definition PiQ :: "'a set \ ('a \ 'b quasi_borel) \ ('a \ 'b) quasi_borel" where "PiQ I X \ Abs_quasi_borel (\\<^sub>E i\I. qbs_space (X i), {\. \i. (i \ I \ (\r. \ r i) \ qbs_Mx (X i)) \ (i \ I \ (\r. \ r i) = (\r. undefined))})" syntax "_PiQ" :: "pttrn \ 'i set \ 'a quasi_borel \ ('i => 'a) quasi_borel" ("(3\\<^sub>Q _\_./ _)" 10) translations "\\<^sub>Q x\I. X" == "CONST PiQ I (\x. X)" lemma shows PiQ_space: "qbs_space (PiQ I X) = (\\<^sub>E i\I. qbs_space (X i))" (is ?goal1) and PiQ_Mx: "qbs_Mx (PiQ I X) = {\. \i. (i \ I \ (\r. \ r i) \ qbs_Mx (X i)) \ (i \ I \ (\r. \ r i) = (\r. undefined))}" (is "_ = ?Mx") proof - have "?Mx \ UNIV \ (\\<^sub>E i\I. qbs_space (X i))" using qbs_Mx_to_X[of _ "X _"] by auto metis moreover have "qbs_closed1 ?Mx" proof(safe intro!: qbs_closed1I) fix \ i and f :: "real \ real" assume h[measurable]:"\i. (i \ I \ (\r. \ r i) \ qbs_Mx (X i)) \ (i \ I \ (\r. \ r i) = (\r. undefined))" "f \ borel \\<^sub>M borel" show "(\r. (\ \ f) r i) \ qbs_Mx (X i)" if i:"i \ I" proof - have "(\r. \ r i) \ f \ qbs_Mx (X i)" using h i by auto thus "(\r. (\ \ f) r i) \ qbs_Mx (X i)" by(simp add: comp_def) qed show "i \ I \ (\r. (\ \ f) r i) = (\r. undefined)" by (metis comp_apply h(1)) qed moreover have "qbs_closed2 (\\<^sub>E i\I. qbs_space (X i)) ?Mx" by(rule qbs_closed2I) (auto simp: PiE_def extensional_def Pi_def) moreover have "qbs_closed3 ?Mx" proof(rule qbs_closed3I) fix P :: "real \ nat" and Fi assume h:"P \ borel \\<^sub>M count_space UNIV" "\i::nat. Fi i \ ?Mx" show "(\r. Fi (P r) r) \ ?Mx" proof safe fix i assume hi:"i \ I" then show "(\r. Fi (P r) r i) \ qbs_Mx (X i)" using h qbs_closed3_dest[OF h(1),of "\j r. Fi j r i"] by auto next show "\i. i \ I \ (\r. Fi (P r) r i) = (\r. undefined)" using h by auto meson qed qed ultimately have "Rep_quasi_borel (PiQ I X) = (\\<^sub>E i\I. qbs_space (X i), ?Mx)" by(auto intro!: Abs_quasi_borel_inverse is_quasi_borel_intro simp: PiQ_def) thus ?goal1 "qbs_Mx (PiQ I X) = ?Mx" by(simp_all add: qbs_space_def qbs_Mx_def) qed lemma prod_qbs_MxI: assumes "\i. i \ I \ (\r. \ r i) \ qbs_Mx (X i)" and "\i. i \ I \ (\r. \ r i) = (\r. undefined)" shows "\ \ qbs_Mx (PiQ I X)" using assms by(auto simp: PiQ_Mx) lemma prod_qbs_MxD: assumes "\ \ qbs_Mx (PiQ I X)" shows "\i. i \ I \ (\r. \ r i) \ qbs_Mx (X i)" and "\i. i \ I \ (\r. \ r i) = (\r. undefined)" and "\i r. i \ I \ \ r i = undefined" using assms by(auto simp: PiQ_Mx dest: fun_cong[where g="(\r. undefined)"]) lemma PiQ_eqI: assumes "\i. i \ I \ X i = Y i" shows "PiQ I X = PiQ I Y" by(auto intro!: qbs_eqI simp: PiQ_Mx assms) lemma PiQ_empty: "qbs_space (PiQ {} X) = {\i. undefined}" by(auto simp: PiQ_space) lemma PiQ_empty_Mx: "qbs_Mx (PiQ {} X) = {\r i. undefined}" by(auto simp: PiQ_Mx) meson subsubsection \ Coproduct Spaces \ -definition coprod_qbs_Mx :: "['a set, 'a \ 'b quasi_borel] \ (real \ 'a \ 'b) set" where -"coprod_qbs_Mx I X \ { \r. (f r, \ (f r) r) |f \. f \ borel \\<^sub>M count_space I \ (\i\range f. \ i \ qbs_Mx (X i))}" +definition coPiQ_Mx :: "['a set, 'a \ 'b quasi_borel] \ (real \ 'a \ 'b) set" where +"coPiQ_Mx I X \ { \r. (f r, \ (f r) r) |f \. f \ borel \\<^sub>M count_space I \ (\i\range f. \ i \ qbs_Mx (X i))}" -definition coprod_qbs_Mx' :: "['a set, 'a \ 'b quasi_borel] \ (real \ 'a \ 'b) set" where -"coprod_qbs_Mx' I X \ { \r. (f r, \ (f r) r) |f \. f \ borel \\<^sub>M count_space I \ (\i. (i \ range f \ qbs_space (X i) \ {}) \ \ i \ qbs_Mx (X i))}" +definition coPiQ_Mx' :: "['a set, 'a \ 'b quasi_borel] \ (real \ 'a \ 'b) set" where +"coPiQ_Mx' I X \ { \r. (f r, \ (f r) r) |f \. f \ borel \\<^sub>M count_space I \ (\i. (i \ range f \ qbs_space (X i) \ {}) \ \ i \ qbs_Mx (X i))}" -lemma coproduct_qbs_Mx_eq: - "coprod_qbs_Mx I X = coprod_qbs_Mx' I X" +lemma coPiQ_Mx_eq: + "coPiQ_Mx I X = coPiQ_Mx' I X" proof safe fix \ - assume "\ \ coprod_qbs_Mx I X" + assume "\ \ coPiQ_Mx I X" then obtain f \ where hfb: "f \ borel \\<^sub>M count_space I" "\i. i \ range f \ \ i \ qbs_Mx (X i)" "\ = (\r. (f r, \ (f r) r))" - unfolding coprod_qbs_Mx_def by blast + unfolding coPiQ_Mx_def by blast define \' where "\' \ (\i. if i \ range f then \ i else if qbs_space (X i) \ {} then (SOME \. \ \ qbs_Mx (X i)) else \ i)" have 1:"\ = (\r. (f r, \' (f r) r))" by(simp add: hfb(3) \'_def) have 2:"\i. qbs_space (X i) \ {} \ \' i \ qbs_Mx (X i)" proof - fix i assume hne:"qbs_space (X i) \ {}" then obtain x where "x \ qbs_space (X i)" by auto hence "(\r. x) \ qbs_Mx (X i)" by auto thus "\' i \ qbs_Mx (X i)" by(cases "i \ range f") (auto simp: \'_def hfb(2) hne intro!: someI2[where a="\r. x"]) qed - show "\ \ coprod_qbs_Mx' I X" - using hfb(1,2) 1 2 \'_def by(auto simp: coprod_qbs_Mx'_def intro!: exI[where x=f] exI[where x=\']) + show "\ \ coPiQ_Mx' I X" + using hfb(1,2) 1 2 \'_def by(auto simp: coPiQ_Mx'_def intro!: exI[where x=f] exI[where x=\']) next fix \ - assume "\ \ coprod_qbs_Mx' I X" + assume "\ \ coPiQ_Mx' I X" then obtain f \ where hfb: "f \ borel \\<^sub>M count_space I" "\i. qbs_space (X i) \ {} \ \ i \ qbs_Mx (X i)" "\i. i \ range f \ \ i \ qbs_Mx (X i)" "\ = (\r. (f r, \ (f r) r))" - unfolding coprod_qbs_Mx'_def by blast - show "\ \ coprod_qbs_Mx I X" - by(auto simp: hfb(4) coprod_qbs_Mx_def intro!: hfb(1) hfb(3)) + unfolding coPiQ_Mx'_def by blast + show "\ \ coPiQ_Mx I X" + by(auto simp: hfb(4) coPiQ_Mx_def intro!: hfb(1) hfb(3)) qed -definition coprod_qbs :: "['a set, 'a \ 'b quasi_borel] \ ('a \ 'b) quasi_borel" where -"coprod_qbs I X \ Abs_quasi_borel (SIGMA i:I. qbs_space (X i), coprod_qbs_Mx I X)" +definition coPiQ :: "['a set, 'a \ 'b quasi_borel] \ ('a \ 'b) quasi_borel" where +"coPiQ I X \ Abs_quasi_borel (SIGMA i:I. qbs_space (X i), coPiQ_Mx I X)" syntax - "_coprod_qbs" :: "pttrn \ 'i set \ 'a quasi_borel \ ('i \ 'a) quasi_borel" ("(3\\<^sub>Q _\_./ _)" 10) + "_coPiQ" :: "pttrn \ 'i set \ 'a quasi_borel \ ('i \ 'a) quasi_borel" ("(3\\<^sub>Q _\_./ _)" 10) translations - "\\<^sub>Q x\I. X" \ "CONST coprod_qbs I (\x. X)" + "\\<^sub>Q x\I. X" \ "CONST coPiQ I (\x. X)" lemma - shows coprod_qbs_space: "qbs_space (coprod_qbs I X) = (SIGMA i:I. qbs_space (X i))" (is ?goal1) - and coprod_qbs_Mx: "qbs_Mx (coprod_qbs I X) = coprod_qbs_Mx I X" (is ?goal2) + shows coPiQ_space: "qbs_space (coPiQ I X) = (SIGMA i:I. qbs_space (X i))" (is ?goal1) + and coPiQ_Mx: "qbs_Mx (coPiQ I X) = coPiQ_Mx I X" (is ?goal2) proof - - have "coprod_qbs_Mx I X \ UNIV \ (SIGMA i:I. qbs_space (X i))" - by(fastforce simp: coprod_qbs_Mx_def dest: measurable_space qbs_Mx_to_X) - moreover have "qbs_closed1 (coprod_qbs_Mx I X)" + have "coPiQ_Mx I X \ UNIV \ (SIGMA i:I. qbs_space (X i))" + by(fastforce simp: coPiQ_Mx_def dest: measurable_space qbs_Mx_to_X) + moreover have "qbs_closed1 (coPiQ_Mx I X)" proof(rule qbs_closed1I) fix \ and f :: "real \ real" - assume "\ \ coprod_qbs_Mx I X" + assume "\ \ coPiQ_Mx I X" and 1[measurable]: "f \ borel \\<^sub>M borel" then obtain \ g where ha: "\i. i \ range g \ \ i \ qbs_Mx (X i)" "\ = (\r. (g r, \ (g r) r))" and [measurable]:"g \ borel \\<^sub>M count_space I" - by(fastforce simp: coprod_qbs_Mx_def) + by(fastforce simp: coPiQ_Mx_def) then have "\i. i \ range g \ \ i \ f \ qbs_Mx (X i)" by simp - thus "\ \ f \ coprod_qbs_Mx I X" - unfolding coprod_qbs_Mx_def by (auto intro!: exI[where x="g \ f"] exI[where x="\i. \ i \ f"] simp: ha(2)) + thus "\ \ f \ coPiQ_Mx I X" + unfolding coPiQ_Mx_def by (auto intro!: exI[where x="g \ f"] exI[where x="\i. \ i \ f"] simp: ha(2)) qed - moreover have "qbs_closed2 (SIGMA i:I. qbs_space (X i)) (coprod_qbs_Mx I X)" + moreover have "qbs_closed2 (SIGMA i:I. qbs_space (X i)) (coPiQ_Mx I X)" proof(safe intro!: qbs_closed2I) fix i x assume "i \ I" "x \ qbs_space (X i)" - then show "(\r. (i,x)) \ coprod_qbs_Mx I X" - by(auto simp: coprod_qbs_Mx_def intro!: exI[where x="\r. i"]) + then show "(\r. (i,x)) \ coPiQ_Mx I X" + by(auto simp: coPiQ_Mx_def intro!: exI[where x="\r. i"]) qed - moreover have "qbs_closed3 (coprod_qbs_Mx I X)" + moreover have "qbs_closed3 (coPiQ_Mx I X)" proof(rule qbs_closed3I) fix P :: "real \ nat" and Fi assume h[measurable]:"P \ borel \\<^sub>M count_space UNIV" - "\i :: nat. Fi i \ coprod_qbs_Mx I X" + "\i :: nat. Fi i \ coPiQ_Mx I X" then have "\i. \fi \i. Fi i = (\r. (fi r, \i (fi r) r)) \ fi \ borel \\<^sub>M count_space I \ (\j. (j \ range fi \ qbs_space (X j) \ {}) \ \i j \ qbs_Mx (X j))" - by(auto simp: coproduct_qbs_Mx_eq coprod_qbs_Mx'_def) + by(auto simp: coPiQ_Mx_eq coPiQ_Mx'_def) then obtain fi where "\i. \\i. Fi i = (\r. (fi i r, \i (fi i r) r)) \ fi i \ borel \\<^sub>M count_space I \ (\j. (j \ range (fi i) \ qbs_space (X j) \ {}) \ \i j \ qbs_Mx (X j))" by(fastforce intro!: choice) then obtain \i where "\i. Fi i = (\r. (fi i r, \i i (fi i r) r)) \ fi i \ borel \\<^sub>M count_space I \ (\j. (j \ range (fi i) \ qbs_space (X j) \ {}) \ \i i j \ qbs_Mx (X j))" by(fastforce intro!: choice) then have hf[measurable]: "\i. Fi i = (\r. (fi i r, \i i (fi i r) r))" "\i. fi i \ borel \\<^sub>M count_space I" "\i j. j \ range (fi i) \ \i i j \ qbs_Mx (X j)" "\i j. qbs_space (X j) \ {} \ \i i j \ qbs_Mx (X j)" by auto define f' where "f' \ (\r. fi (P r) r)" define \' where "\' \ (\i r. \i (P r) i r)" have 1:"(\r. Fi (P r) r) = (\r. (f' r, \' (f' r) r))" by(simp add: \'_def f'_def hf) have "f' \ borel \\<^sub>M count_space I" by(simp add: f'_def) moreover have "\i. i \ range f' \ \' i \ qbs_Mx (X i)" proof - fix i assume hi:"i \ range f'" then obtain r where hr: "i = fi (P r) r" by(auto simp: f'_def) hence "i \ range (fi (P r))" by simp hence "\i (P r) i \ qbs_Mx (X i)" by(simp add: hf) hence "qbs_space (X i) \ {}" by(auto simp: qbs_empty_equiv) hence "\j. \i j i \ qbs_Mx (X i)" by(simp add: hf(4)) then show "\' i \ qbs_Mx (X i)" by(auto simp: \'_def h(1) intro!: qbs_closed3_dest[of P "\j. \i j i"]) qed - ultimately show "(\r. Fi (P r) r) \ coprod_qbs_Mx I X" - by(auto simp: 1 coprod_qbs_Mx_def intro!: exI[where x=f']) + ultimately show "(\r. Fi (P r) r) \ coPiQ_Mx I X" + by(auto simp: 1 coPiQ_Mx_def intro!: exI[where x=f']) qed - ultimately have "Rep_quasi_borel (coprod_qbs I X) = (SIGMA i:I. qbs_space (X i), coprod_qbs_Mx I X)" - unfolding coprod_qbs_def by(fastforce intro!: Abs_quasi_borel_inverse) + ultimately have "Rep_quasi_borel (coPiQ I X) = (SIGMA i:I. qbs_space (X i), coPiQ_Mx I X)" + unfolding coPiQ_def by(fastforce intro!: Abs_quasi_borel_inverse) thus ?goal1 ?goal2 by(simp_all add: qbs_space_def qbs_Mx_def) qed -lemma coprod_qbs_MxI: +lemma coPiQ_MxI: assumes "f \ borel \\<^sub>M count_space I" and "\i. i \ range f \ \ i \ qbs_Mx (X i)" - shows "(\r. (f r, \ (f r) r)) \ qbs_Mx (coprod_qbs I X)" - using assms unfolding coprod_qbs_Mx_def coprod_qbs_Mx by blast + shows "(\r. (f r, \ (f r) r)) \ qbs_Mx (coPiQ I X)" + using assms unfolding coPiQ_Mx_def coPiQ_Mx by blast -lemma coprod_qbs_eqI: +lemma coPiQ_eqI: assumes "\i. i \ I \ X i = Y i" - shows "coprod_qbs I X = coprod_qbs I Y" - using assms by(auto intro!: qbs_eqI simp: coprod_qbs_Mx coprod_qbs_Mx_def) (metis UNIV_I measurable_space space_borel space_count_space)+ + shows "coPiQ I X = coPiQ I Y" + using assms by(auto intro!: qbs_eqI simp: coPiQ_Mx coPiQ_Mx_def) (metis UNIV_I measurable_space space_borel space_count_space)+ subsubsection \ List Spaces \ text \ We define the quasi-Borel spaces on list using the following isomorphism. \begin{align*} List(X) \cong \coprod_{n\in \mathbb{N}} \prod_{0\leq i < n} X \end{align*}\ -definition "list_of X \ \\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. " *) definition list_nil :: "nat \ (nat \ 'a)" where "list_nil \ (0, \n. undefined)" definition list_cons :: "['a, nat \ (nat \ 'a)] \ nat \ (nat \ 'a)" where "list_cons x l \ (Suc (fst l), (\n. if n = 0 then x else (snd l) (n - 1)))" fun from_list :: "'a list \ nat \ (nat \ 'a)" where "from_list [] = list_nil" | "from_list (a#l) = list_cons a (from_list l)" fun to_list' :: "nat \ (nat \ 'a) \ 'a list" where "to_list' 0 _ = []" | "to_list' (Suc n) f = f 0 # to_list' n (\n. f (Suc n))" definition to_list :: "nat \ (nat \ 'a) \ 'a list" where "to_list \ case_prod to_list'" +lemma inj_on_to_list: "inj_on (to_list :: nat \ (nat \ 'a) \ 'a list) (SIGMA n:UNIV. PiE {.. 'a" + assume h:"to_list (n, x) = to_list (m, y)" + have 1:"\a. length (to_list (n, a)) = n" for n + by(induction n) (auto simp: to_list_def) + show n:"n = m" + using 1 arg_cong[OF h,of length] by metis + show "x \ PiE {.. y \ PiE {.. x = y" + using h unfolding n + proof(induction m arbitrary: x y A) + case ih:(Suc m) + then have "to_list (m, (\n. x (Suc n))) = to_list (m, (\n. y (Suc n)))" + by(auto simp: to_list_def) + hence 1:"(\n. x (Suc n)) = (\n. y (Suc n))" + using ih(2-4) by(intro ih(1)[of _ "\n. A (Suc n)"]) auto + show ?case + proof + fix n + show "x n = y n" + proof(cases n) + assume "n = 0" + then show "x n = y n" + using ih(4) by(auto simp: to_list_def) + qed(use fun_cong[OF 1] in auto) + qed + qed(auto simp: to_list_def) +qed + text \ Definition \ definition list_qbs :: "'a quasi_borel \ 'a list quasi_borel" where -"list_qbs X \ map_qbs to_list (list_of X)" +"list_qbs X \ map_qbs to_list (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. (nat \ 'a) \ 'a" where "list_head l = snd l 0" definition list_tail :: "nat \ (nat \ 'a) \ nat \ (nat \ 'a)" where "list_tail l = (fst l - 1, \m. (snd l) (Suc m))" lemma list_simp1: "list_nil \ list_cons x l" by (simp add: list_nil_def list_cons_def) lemma list_simp2: assumes "list_cons a al = list_cons b bl" shows "a = b" "al = bl" proof - have "a = snd (list_cons a al) 0" "b = snd (list_cons b bl) 0" by (auto simp: list_cons_def) thus "a = b" by(simp add: assms) next have "fst al = fst bl" using assms by (simp add: list_cons_def) moreover have "snd al = snd bl" proof fix n have "snd al n = snd (list_cons a al) (Suc n)" by (simp add: list_cons_def) also have "... = snd (list_cons b bl) (Suc n)" by (simp add: assms) also have "... = snd bl n" by (simp add: list_cons_def) finally show "snd al n = snd bl n" . qed ultimately show "al = bl" by (simp add: prod.expand) qed lemma shows list_simp3:"list_head (list_cons a l) = a" and list_simp4:"list_tail (list_cons a l) = l" by(simp_all add: list_head_def list_cons_def list_tail_def) lemma list_decomp1: - assumes "l \ qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. - (\a l'. a \ qbs_space X \ l' \ qbs_space (list_of X) \ l = list_cons a l')" + (\a l'. a \ qbs_space X \ l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. l = list_cons a l')" proof(cases l) case hl:(Pair n f) show ?thesis proof(cases n) case 0 then show ?thesis - using assms hl by (simp add: list_of_def list_nil_def coprod_qbs_space PiQ_space) + using assms hl by (simp add: list_nil_def coPiQ_space PiQ_space) next case hn:(Suc n') define f' where "f' \ \m. f (Suc m)" have "l = list_cons (f 0) (n',f')" unfolding hl hn list_cons_def proof safe fix m show "f = (\m. if m = 0 then f 0 else snd (n', f') (m - 1))" proof fix m show "f m = (if m = 0 then f 0 else snd (n', f') (m - 1))" using assms hl by(cases m; fastforce simp: f'_def) qed qed simp - moreover have "(n', f') \ qbs_space (list_of X)" + moreover have "(n', f') \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..x. x \ {.. f' x \ qbs_space X" - using assms hl hn by(fastforce simp: f'_def list_of_def coprod_qbs_space PiQ_space) + using assms hl hn by(fastforce simp: f'_def coPiQ_space PiQ_space) moreover { fix x assume 1:"x \ {.. qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. list_nil" shows "l = list_cons (list_head l) (list_tail l)" proof - obtain a l' where hl: - "a \ qbs_space X" "l' \ qbs_space (list_of X)" "l = list_cons a l'" + "a \ qbs_space X" "l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" - by (simp add: list_nil_def list_of_def coprod_qbs_space PiQ_space) + "list_nil \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space X" - and "l \ qbs_space (list_of X)" - shows "list_cons a l \ qbs_space (list_of X)" - using assms by(fastforce simp: PiE_def extensional_def list_cons_def list_of_def coprod_qbs_space PiQ_space) + and "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..a l'. a \ qbs_space X \ l' \ qbs_space (list_of X) \ P (list_cons a l')" + and "\a l'. a \ qbs_space X \ l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. P (list_cons a l')" shows "P l" by(rule disjE[OF list_decomp1[OF assms(1)]]) (use assms in auto) lemma list_induct_rule: - assumes "l \ qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{..a l'. a \ qbs_space X \ l' \ qbs_space (list_of X) \ P l' \ P (list_cons a l')" + and "\a l'. a \ qbs_space X \ l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. P l' \ P (list_cons a l')" shows "P l" proof(cases l) case hl:(Pair n f) then show ?thesis using assms(1) proof(induction n arbitrary: f l) case 0 then show ?case - using assms(2) by (simp add: list_of_def coprod_qbs_space PiQ_space list_nil_def) + using assms(2) by (simp add: coPiQ_space PiQ_space list_nil_def) next case ih:(Suc n) then obtain a l' where hl: - "a \ qbs_space X" "l' \ qbs_space (list_of X)" "l = list_cons a l'" + "a \ qbs_space X" "l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space X" by(rule list_induct_rule[OF assms]) (auto simp: to_list_simp1 to_list_simp2) lemma from_list_length: "fst (from_list l) = length l" by(induction l, simp_all add: list_cons_def list_nil_def) lemma from_list_in_list_of: assumes "set l \ qbs_space X" - shows "from_list l \ qbs_space (list_of X)" - using assms by(induction l) (auto simp: PiE_def extensional_def Pi_def list_of_def coprod_qbs_space PiQ_space list_nil_def list_cons_def) + shows "from_list l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of (Abs_quasi_borel (UNIV,UNIV)))" +lemma from_list_in_list_of': "from_list l \ qbs_space ((\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (Abs_quasi_borel (UNIV,UNIV))" by(simp add: qbs_space_def Abs_quasi_borel_inverse[of "(UNIV,UNIV)",simplified is_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def,simplified]) thus ?thesis using from_list_in_list_of by blast qed lemma list_cons_in_list_of: assumes "set (a#l) \ qbs_space X" - shows "list_cons a (from_list l) \ qbs_space (list_of X)" + shows "list_cons a (from_list l) \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space (list_of X)" + assume h: "l' \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. ('a \ (nat \ (nat \ 'a)) \ 'b \ 'b) \ (nat \ (nat \ 'a)) \ 'b" where "rec_list' t0 f l \ (rec_list t0 (\x l'. f x (from_list l')) (to_list l))" lemma rec_list'_simp1: "rec_list' t f list_nil = t" by(simp add: rec_list'_def to_list_simp1) lemma rec_list'_simp2: - assumes "l \ qbs_space (list_of X)" + assumes "l \ qbs_space (\\<^sub>Q n\(UNIV :: nat set).\\<^sub>Q i\{.. qbs_space X}" - using to_list_set by(auto simp: list_qbs_def map_qbs_space image_def from_list_to_list_ident from_list_in_list_of intro!: bexI[where x="from_list _"]) +lemma list_qbs_space: "qbs_space (list_qbs X) = lists (qbs_space X)" + using to_list_set by(auto simp: list_qbs_def map_qbs_space image_def from_list_to_list_ident from_list_in_list_of subset_iff intro!: bexI[where x="from_list _"]) subsubsection \ Option Spaces \ text \ The option spaces is defined using the following isomorphism. \begin{align*} Option(X) \cong X + 1 \end{align*}\ definition option_qbs :: "'a quasi_borel \ 'a option quasi_borel" where "option_qbs X = map_qbs (\x. case x of Inl y \ Some y | Inr y \ None) (X \\<^sub>Q 1\<^sub>Q)" lemma option_qbs_space: "qbs_space (option_qbs X) = {Some x|x. x \ qbs_space X} \ {None}" by(auto simp: option_qbs_def map_qbs_space copair_qbs_space) (metis InrI image_eqI insert_iff old.sum.simps(6), metis InlI image_iff sum.case(1)) subsubsection \ Function Spaces \ definition exp_qbs :: "['a quasi_borel, 'b quasi_borel] \ ('a \ 'b) quasi_borel" (infixr "\\<^sub>Q" 61) where "X \\<^sub>Q Y \ Abs_quasi_borel ({f. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}, {g. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y})" lemma shows exp_qbs_space: "qbs_space (exp_qbs X Y) = {f. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}" and exp_qbs_Mx: "qbs_Mx (exp_qbs X Y) = {g. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y}" proof - have "{g:: real \ _. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y} \ UNIV \ {f. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}" proof safe fix g :: "real \ _" and r :: real and \ assume h:"\\\borel_measurable borel. \\\qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y" "\ \ qbs_Mx X" have [simp]: "g r \ \ = (\l. g r (\ l))" by (auto simp: comp_def) thus "g r \ \ \ qbs_Mx Y" using h by auto qed moreover have "qbs_closed3 {g. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y}" by(rule qbs_closed3I, auto) (rule qbs_closed3_dest,auto) ultimately have "Rep_quasi_borel (exp_qbs X Y) = ({f. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}, {g. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y})" unfolding exp_qbs_def by(auto intro!: Abs_quasi_borel_inverse is_quasi_borel_intro qbs_closed1I qbs_closed2I simp: comp_def) thus "qbs_space (exp_qbs X Y) = {f. \\ \ qbs_Mx X. f \ \ \ qbs_Mx Y}" "qbs_Mx (exp_qbs X Y) = {g. \\\ borel_measurable borel. \\\ qbs_Mx X. (\r. g (\ r) (\ r)) \ qbs_Mx Y}" by(simp_all add: qbs_space_def qbs_Mx_def) qed subsubsection \ Ordering on Quasi-Borel Spaces \ inductive_set generating_Mx :: "'a set \ (real \ 'a) set \ (real \ 'a) set" for X :: "'a set" and Mx :: "(real \ 'a) set" where Basic: "\ \ Mx \ \ \ generating_Mx X Mx" | Const: "x \ X \ (\r. x) \ generating_Mx X Mx" | Comp : "f \ (borel :: real measure) \\<^sub>M (borel :: real measure) \ \ \ generating_Mx X Mx \ \ \ f \ generating_Mx X Mx" | Part : "(\i. Fi i \ generating_Mx X Mx) \ P \ borel \\<^sub>M count_space (UNIV :: nat set) \ (\r. Fi (P r) r) \ generating_Mx X Mx" lemma generating_Mx_to_space: assumes "Mx \ UNIV \ X" shows "generating_Mx X Mx \ UNIV \ X" proof fix \ assume "\ \ generating_Mx X Mx" then show "\ \ UNIV \ X" by(induct rule: generating_Mx.induct) (use assms in auto) qed lemma generating_Mx_closed1: "qbs_closed1 (generating_Mx X Mx)" by (simp add: generating_Mx.Comp qbs_closed1I) lemma generating_Mx_closed2: "qbs_closed2 X (generating_Mx X Mx)" by (simp add: generating_Mx.Const qbs_closed2I) lemma generating_Mx_closed3: "qbs_closed3 (generating_Mx X Mx)" by(simp add: qbs_closed3I generating_Mx.Part) lemma generating_Mx_Mx: "generating_Mx (qbs_space X) (qbs_Mx X) = qbs_Mx X" proof safe fix \ assume "\ \ generating_Mx (qbs_space X) (qbs_Mx X)" then show "\ \ qbs_Mx X" by(rule generating_Mx.induct) (auto intro!: qbs_closed1_dest[simplified comp_def] simp: qbs_closed3_dest') next fix \ assume "\ \ qbs_Mx X" then show "\ \ generating_Mx (qbs_space X) (qbs_Mx X)" .. qed instantiation quasi_borel :: (type) order_bot begin inductive less_eq_quasi_borel :: "'a quasi_borel \ 'a quasi_borel \ bool" where "qbs_space X \ qbs_space Y \ less_eq_quasi_borel X Y" | "qbs_space X = qbs_space Y \ qbs_Mx Y \ qbs_Mx X \ less_eq_quasi_borel X Y" lemma le_quasi_borel_iff: "X \ Y \ (if qbs_space X = qbs_space Y then qbs_Mx Y \ qbs_Mx X else qbs_space X \ qbs_space Y)" by(auto elim: less_eq_quasi_borel.cases intro: less_eq_quasi_borel.intros) definition less_quasi_borel :: "'a quasi_borel \ 'a quasi_borel \ bool" where "less_quasi_borel X Y \ (X \ Y \ \ Y \ X)" definition bot_quasi_borel :: "'a quasi_borel" where "bot_quasi_borel = empty_quasi_borel" instance proof show "bot \ a" for a :: "'a quasi_borel" using qbs_empty_equiv by(auto simp add: le_quasi_borel_iff bot_quasi_borel_def) qed (auto simp: le_quasi_borel_iff less_quasi_borel_def split: if_split_asm intro: qbs_eqI) end definition inf_quasi_borel :: "['a quasi_borel, 'a quasi_borel] \ 'a quasi_borel" where "inf_quasi_borel X X' = Abs_quasi_borel (qbs_space X \ qbs_space X', qbs_Mx X \ qbs_Mx X')" lemma inf_quasi_borel_correct: "Rep_quasi_borel (inf_quasi_borel X X') = (qbs_space X \ qbs_space X', qbs_Mx X \ qbs_Mx X')" by(auto intro!: Abs_quasi_borel_inverse simp: inf_quasi_borel_def is_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def dest: qbs_Mx_to_X) lemma inf_qbs_space[simp]: "qbs_space (inf_quasi_borel X X') = qbs_space X \ qbs_space X'" by (simp add: qbs_space_def inf_quasi_borel_correct) lemma inf_qbs_Mx[simp]: "qbs_Mx (inf_quasi_borel X X') = qbs_Mx X \ qbs_Mx X'" by(simp add: qbs_Mx_def inf_quasi_borel_correct) definition max_quasi_borel :: "'a set \ 'a quasi_borel" where "max_quasi_borel X = Abs_quasi_borel (X, UNIV \ X)" lemma max_quasi_borel_correct: "Rep_quasi_borel (max_quasi_borel X) = (X, UNIV \ X)" by(fastforce intro!: Abs_quasi_borel_inverse simp: max_quasi_borel_def qbs_closed1_def qbs_closed2_def qbs_closed3_def is_quasi_borel_def) lemma max_qbs_space[simp]: "qbs_space (max_quasi_borel X) = X" by(simp add: qbs_space_def max_quasi_borel_correct) lemma max_qbs_Mx[simp]: "qbs_Mx (max_quasi_borel X) = UNIV \ X" by(simp add: qbs_Mx_def max_quasi_borel_correct) instantiation quasi_borel :: (type) semilattice_sup begin definition sup_quasi_borel :: "'a quasi_borel \ 'a quasi_borel \ 'a quasi_borel" where "sup_quasi_borel X Y \ (if qbs_space X = qbs_space Y then inf_quasi_borel X Y else if qbs_space X \ qbs_space Y then Y else if qbs_space Y \ qbs_space X then X else max_quasi_borel (qbs_space X \ qbs_space Y))" instance proof fix X Y :: "'a quasi_borel" let ?X = "qbs_space X" let ?Y = "qbs_space Y" consider "?X = ?Y" | "?X \ ?Y" | "?Y \ ?X" | "?X \ ?X \ ?Y \ ?Y \ ?X \ ?Y" by auto then show "X \ X \ Y" proof(cases) case 1 show ?thesis unfolding sup_quasi_borel_def by(rule less_eq_quasi_borel.intros(2),simp_all add: 1) next case 2 then show ?thesis unfolding sup_quasi_borel_def by (simp add: less_eq_quasi_borel.intros(1)) next case 3 then show ?thesis unfolding sup_quasi_borel_def by auto next case 4 then show ?thesis unfolding sup_quasi_borel_def by(auto simp: less_eq_quasi_borel.intros(1)) qed next fix X Y :: "'a quasi_borel" let ?X = "qbs_space X" let ?Y = "qbs_space Y" consider "?X = ?Y" | "?X \ ?Y" | "?Y \ ?X" | "?X \ ?X \ ?Y \ ?Y \ ?X \ ?Y" by auto then show "Y \ X \ Y" proof(cases) case 1 show ?thesis unfolding sup_quasi_borel_def by(rule less_eq_quasi_borel.intros(2)) (simp_all add: 1) next case 2 then show ?thesis unfolding sup_quasi_borel_def by auto next case 3 then show ?thesis unfolding sup_quasi_borel_def by (auto simp add: less_eq_quasi_borel.intros(1)) next case 4 then show ?thesis unfolding sup_quasi_borel_def by(auto simp: less_eq_quasi_borel.intros(1)) qed next fix X Y Z :: "'a quasi_borel" assume h:"X \ Z" "Y \ Z" let ?X = "qbs_space X" let ?Y = "qbs_space Y" let ?Z = "qbs_space Z" consider "?X = ?Y" | "?X \ ?Y" | "?Y \ ?X" | "?X \ ?X \ ?Y \ ?Y \ ?X \ ?Y" by auto then show "sup X Y \ Z" proof cases case 1 show ?thesis unfolding sup_quasi_borel_def apply(simp add: 1,rule less_eq_quasi_borel.cases[OF h(1)]) apply(rule less_eq_quasi_borel.intros(1)) apply auto[1] apply simp apply(rule less_eq_quasi_borel.intros(2)) apply(simp add: 1) apply(rule less_eq_quasi_borel.cases[OF h(2)]) using 1 apply fastforce apply simp by (metis "1" h(2) inf_qbs_Mx le_inf_iff le_quasi_borel_iff) next case 2 then show ?thesis unfolding sup_quasi_borel_def using h(2) by auto next case 3 then show ?thesis unfolding sup_quasi_borel_def using h(1) by auto next case 4 then have [simp]:"?X \ ?Y" "~ (?X \ ?Y)" "~ (?Y \ ?X)" by auto have [simp]:"?X \ ?Z" "?Y \ ?Z" by (metis h(1) dual_order.order_iff_strict less_eq_quasi_borel.cases) (metis h(2) dual_order.order_iff_strict less_eq_quasi_borel.cases) then consider "?X \ ?Y = ?Z" | "?X \ ?Y \ ?Z" by blast then show ?thesis unfolding sup_quasi_borel_def apply cases apply simp apply(rule less_eq_quasi_borel.intros(2)) apply simp using qbs_Mx_to_X apply auto[1] by(simp add: less_eq_quasi_borel.intros(1)) qed qed end end \ No newline at end of file diff --git a/thys/Standard_Borel_Spaces/StandardBorel.thy b/thys/Standard_Borel_Spaces/StandardBorel.thy --- a/thys/Standard_Borel_Spaces/StandardBorel.thy +++ b/thys/Standard_Borel_Spaces/StandardBorel.thy @@ -1,1468 +1,1479 @@ (* Title: StandardBorel.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) section \Standard Borel Spaces\ subsection \Standard Borel Spaces\ theory StandardBorel imports Abstract_Metrizable_Topology begin locale standard_borel = fixes M :: "'a measure" assumes polish_space: "\S. polish_space S \ sets M = sets (borel_of S)" begin lemma singleton_sets: assumes "x \ space M" shows "{x} \ sets M" proof - obtain S where s:"polish_space S" "sets M = sets (borel_of S)" using polish_space by blast have "closedin S {x}" using assms by(simp add: sets_eq_imp_space_eq[OF s(2)] closedin_Hausdorff_sing_eq[OF metrizable_imp_Hausdorff_space[OF polish_space_imp_metrizable_space[OF s(1)]]] space_borel_of) thus ?thesis using borel_of_closed s by simp qed corollary countable_sets: assumes "A \ space M" "countable A" shows "A \ sets M" using sets.countable[OF singleton_sets assms(2)] assms(1) by auto lemma standard_borel_restrict_space: assumes "A \ sets M" shows "standard_borel (restrict_space M A)" proof - obtain S where s:"polish_space S" "sets M = sets (borel_of S)" using polish_space by blast obtain S' where S':"polish_space S'" "sets M = sets (borel_of S')" "openin S' A" using sets_clopen_topology[OF s(1),simplified s(2)[symmetric],OF assms] by auto show ?thesis using polish_space_openin[OF S'(1,3)] S'(2) by(auto simp: standard_borel_def borel_of_subtopology sets_restrict_space intro!: exI[where x="subtopology S' A"] ) qed end locale standard_borel_ne = standard_borel + assumes space_ne: "space M \ {}" begin lemma standard_borel_ne_restrict_space: assumes "A \ sets M" "A \ {}" shows "standard_borel_ne (restrict_space M A)" using assms by(auto simp: standard_borel_ne_def standard_borel_ne_axioms_def standard_borel_restrict_space) lemma standard_borel: "standard_borel M" by(rule standard_borel_axioms) end lemma standard_borel_sets: assumes "standard_borel M" and "sets M = sets N" shows "standard_borel N" using assms by(simp add: standard_borel_def) lemma standard_borel_ne_sets: assumes "standard_borel_ne M" and "sets M = sets N" shows "standard_borel_ne N" using assms by(simp add: standard_borel_def standard_borel_ne_def sets_eq_imp_space_eq[OF assms(2)] standard_borel_ne_axioms_def) lemma pair_standard_borel: assumes "standard_borel M" "standard_borel N" shows "standard_borel (M \\<^sub>M N)" proof - obtain S S' where hs: "polish_space S" "sets M = sets (borel_of S)" "polish_space S'" "sets N = sets (borel_of S')" using assms by(auto simp: standard_borel_def) have "sets (M \\<^sub>M N) = sets (borel_of (prod_topology S S'))" unfolding borel_of_prod[OF polish_space_imp_second_countable[OF hs(1)] polish_space_imp_second_countable[OF hs(3)],symmetric] using sets_pair_measure_cong[OF hs(2,4)] . thus ?thesis unfolding standard_borel_def by(auto intro!: exI[where x="prod_topology S S'"] simp: polish_space_prod[OF hs(1,3)]) qed lemma pair_standard_borel_ne: assumes "standard_borel_ne M" "standard_borel_ne N" shows "standard_borel_ne (M \\<^sub>M N)" using assms by(auto simp: pair_standard_borel standard_borel_ne_def standard_borel_ne_axioms_def space_pair_measure) lemma product_standard_borel: assumes "countable I" and "\i. i \ I \ standard_borel (M i)" shows "standard_borel (\\<^sub>M i\I. M i)" proof - obtain S where hs: "\i. i \ I \ polish_space (S i)" "\i. i \ I \ sets (M i) = sets (borel_of (S i))" using assms(2) by(auto simp: standard_borel_def) metis have "sets (\\<^sub>M i\I. M i) = sets (\\<^sub>M i\I. borel_of (S i))" using hs(2) by(auto intro!: sets_PiM_cong) also have "... = sets (borel_of (product_topology S I))" using assms(1) polish_space_imp_second_countable[OF hs(1)] by(auto intro!: sets_PiM_equal_borel_of) finally have 1:"sets (\\<^sub>M i\I. M i) = sets (borel_of (product_topology S I))". show ?thesis unfolding standard_borel_def using assms(1) hs(1) by(auto intro!: exI[where x="product_topology S I"] polish_space_product simp: 1) qed lemma product_standard_borel_ne: assumes "countable I" and "\i. i \ I \ standard_borel_ne (M i)" shows "standard_borel_ne (\\<^sub>M i\I. M i)" using assms by(auto simp: standard_borel_ne_def standard_borel_ne_axioms_def product_standard_borel) lemma closed_set_standard_borel[simp]: fixes U :: "'a :: topological_space set" assumes "polish_space (euclidean :: 'a topology)" "closed U" shows "standard_borel (restrict_space borel U)" by(auto simp: standard_borel_def borel_of_euclidean borel_of_subtopology assms intro!: exI[where x="subtopology euclidean U"] polish_space_closedin) lemma closed_set_standard_borel_ne[simp]: fixes U :: "'a :: topological_space set" assumes "polish_space (euclidean :: 'a topology)" "closed U" "U \ {}" shows "standard_borel_ne (restrict_space borel U)" using assms by(simp add: standard_borel_ne_def standard_borel_ne_axioms_def) lemma open_set_standard_borel[simp]: fixes U :: "'a :: topological_space set" assumes "polish_space (euclidean :: 'a topology)" "open U" shows "standard_borel (restrict_space borel U)" by(auto simp: standard_borel_def borel_of_euclidean borel_of_subtopology assms intro!: exI[where x="subtopology euclidean U"] polish_space_openin) lemma open_set_standard_borel_ne[simp]: fixes U :: "'a :: topological_space set" assumes "polish_space (euclidean :: 'a topology)" "open U" "U \ {}" shows "standard_borel_ne (restrict_space borel U)" using assms by(simp add: standard_borel_ne_def standard_borel_ne_axioms_def) lemma standard_borel_ne_borel[simp]: "standard_borel_ne (borel :: ('a :: polish_space) measure)" and standard_borel_ne_lborel[simp]: "standard_borel_ne lborel" unfolding standard_borel_def standard_borel_ne_def standard_borel_ne_axioms_def by(auto intro!: exI[where x=euclidean] simp: borel_of_euclidean) lemma count_space_standard'[simp]: assumes "countable I" shows "standard_borel (count_space I)" by(rule standard_borel_sets[OF _ sets_borel_of_discrete_topology]) (auto simp add: assms polish_space_discrete_topology standard_borel_def intro!: exI[where x="discrete_topology I"]) lemma count_space_standard_ne[simp]: "standard_borel_ne (count_space (UNIV :: (_ :: countable) set))" by (simp add: standard_borel_ne_def standard_borel_ne_axioms_def) corollary measure_pmf_standard_borel_ne[simp]: "standard_borel_ne (measure_pmf (p :: (_ :: countable) pmf))" using count_space_standard_ne sets_measure_pmf_count_space standard_borel_ne_sets by blast corollary measure_spmf_standard_borel_ne[simp]: "standard_borel_ne (measure_spmf (p :: (_ :: countable) spmf))" using count_space_standard_ne sets_measure_spmf standard_borel_ne_sets by blast corollary countable_standard_ne[simp]: "standard_borel_ne (borel :: 'a :: {countable,t2_space} measure)" by(simp add: standard_borel_sets[OF _ sets_borel_eq_count_space[symmetric]] standard_borel_ne_def standard_borel_ne_axioms_def) lemma(in standard_borel) countable_discrete_space: assumes "countable (space M)" shows "sets M = Pow (space M)" proof safe fix A assume "A \ space M" with assms have "countable A" by(simp add: countable_subset) thus "A \ sets M" using \A \ space M\ singleton_sets by(auto intro!: sets.countable[of A]) qed(use sets.sets_into_space in auto) lemma(in standard_borel) measurable_isomorphic_standard: assumes "M measurable_isomorphic N" shows "standard_borel N" proof - obtain S where S:"polish_space S" "sets M = sets (borel_of S)" using polish_space by auto from measurable_isomorphic_borels[OF S(2) assms] obtain S' where S': "S homeomorphic_space S' \ sets N = sets (borel_of S')" by auto thus ?thesis by(auto simp: standard_borel_def homeomorphic_polish_space_aux[OF S(1)] intro!:exI[where x=S']) qed lemma(in standard_borel_ne) measurable_isomorphic_standard_ne: assumes "M measurable_isomorphic N" shows "standard_borel_ne N" using measurable_ismorphic_empty2[OF _ assms] by(auto simp: measurable_isomorphic_standard[OF assms] standard_borel_ne_def standard_borel_ne_axioms_def space_ne) +lemma(in standard_borel) standard_borel_embed_measure: + assumes"inj_on f (space M)" + shows "standard_borel (embed_measure M f)" + using measurable_embed_measure2'[OF assms] + by(auto intro!: measurable_isomorphic_standard exI[where x=f] simp: measurable_isomorphic_def measurable_isomorphic_map_def assms in_sets_embed_measure measurable_def sets.sets_into_space space_embed_measure the_inv_into_into the_inv_into_vimage bij_betw_def) + +corollary(in standard_borel_ne) standard_borel_ne_embed_measure: + assumes"inj_on f (space M)" + shows "standard_borel_ne (embed_measure M f)" + by (simp add: assms space_embed_measure space_ne standard_borel_embed_measure standard_borel_ne_axioms_def standard_borel_ne_def) + lemma ereal_standard_ne: "standard_borel_ne (borel :: ereal measure)" proof - interpret s: standard_borel_ne "restrict_space borel {0..1::real}" by auto define f :: "real \ ereal" where "f \ (\r. if r = 0 then bot else if r = 1 then top else tan (pi * r - (pi / 2)))" define g :: "ereal \ real" where "g \ (\r. if r = top then 1 else if r = bot then 0 else arctan (real_of_ereal r) / pi + 1 / 2)" show ?thesis proof(rule s.measurable_isomorphic_standard_ne[OF measurable_isomorphic_byWitness[where f=f and g = g]]) show "f \ borel_measurable (restrict_space borel {0..1})" proof - have 1:"{0..1} \ {r. r \ 0} \ {x. x \ 1} = {0<..<1::real}" by auto have 2:"(\x. ereal (tan (pi * x - pi / 2))) \ borel_measurable (restrict_space borel ({0..1} \ {r. r \ 0} \ {x. x \ 1}))" unfolding 1 proof(safe intro!: borel_measurable_continuous_on_restrict continuous_on_ereal Transcendental.continuous_on_tan) show "continuous_on {0<..<1} (\x::real. pi * x - pi / 2)" by(auto intro!: continuous_at_imp_continuous_on) next fix x :: real assume h:"cos (pi * x - pi / 2) = 0" "x \ {0<..<1}" hence "- (pi / 2) < pi * x - pi / 2" "pi * x - pi / 2 < pi / 2" by simp_all from cos_gt_zero_pi[OF this] h(1) show False by simp qed have "{r:: real. r = 0 \ 0 \ r \ r \ 1} \ sets (restrict_space borel {0..1})" "{x::real. x = 1 \ 0 \ x \ x \ 1 \ x \ 0} \ sets (restrict_space borel ({0..1} \ {r. r \ 0}))" by(auto simp: sets_restrict_space) with 2 show ?thesis by(auto intro!: measurable_If_restrict_space_iff[THEN iffD2] simp: restrict_restrict_space f_def) qed next show "g \ borel \\<^sub>M restrict_space borel {0..1}" unfolding g_def measurable_restrict_space2_iff proof safe fix x :: ereal have "-1 / 2 < arctan (real_of_ereal x) / pi" "arctan (real_of_ereal x) / pi < 1 / 2" using arctan_lbound[of "real_of_ereal x"] arctan_ubound[of "real_of_ereal x"] by (simp_all add: mult_imp_less_div_pos) hence "0 \ arctan (real_of_ereal x) / pi + 1 / 2" "arctan (real_of_ereal x) / pi + 1 / 2 \ 1" by linarith+ thus "(if x = \ then 1 else if x = \ then 0 else arctan (real_of_ereal x) / pi + 1 / 2) \ {0..1}" by auto qed measurable next fix r ::real assume "r \ space (restrict_space borel {0..1})" then consider "r = 0" | "r = 1" | "0 < r" "r < 1" by auto linarith then show " g (f r) = r" proof cases case 3 then have 1:"- (pi / 2) < pi * r - pi / 2" "pi * r - pi / 2 < pi / 2" by simp_all have "arctan (tan (pi * r - pi / 2)) / pi + 1 / 2 = r" by(simp add: arctan_tan[OF 1] diff_divide_distrib) thus ?thesis by(auto simp: f_def g_def top_ereal_def bot_ereal_def) qed(auto simp: g_def f_def top_ereal_def bot_ereal_def) next fix y :: ereal consider "y = top" | "y = bot" | "y \ bot" "y \ top" by auto then show "f (g y) = y" proof cases case 3 hence [simp]: "\y\ \ \" by(auto simp: top_ereal_def bot_ereal_def) have "-1 / 2 < arctan (real_of_ereal y) / pi" "arctan (real_of_ereal y) / pi < 1 / 2" using arctan_lbound[of "real_of_ereal y"] arctan_ubound[of "real_of_ereal y"] by (simp_all add: mult_imp_less_div_pos) hence "arctan (real_of_ereal y) / pi + 1 / 2 < 1" "arctan (real_of_ereal y) / pi + 1 / 2 > 0" by linarith+ thus ?thesis using arctan_lbound[of "real_of_ereal y"] arctan_ubound[of "real_of_ereal y"] by(auto simp: f_def g_def distrib_left tan_arctan ereal_real') qed(auto simp: f_def g_def) qed qed corollary ennreal_stanndard_ne: "standard_borel_ne (borel :: ennreal measure)" by(auto intro!: standard_borel_ne.measurable_isomorphic_standard_ne[OF standard_borel_ne.standard_borel_ne_restrict_space[OF ereal_standard_ne,of "{0..}",simplified]] measurable_isomorphic_byWitness[where f=e2ennreal and g=enn2ereal] measurable_restrict_space1 measurable_restrict_space2 enn2ereal_e2ennreal) text \ Cantor space $\mathscr{C}$ \ definition Cantor_space :: "(nat \ real) measure" where "Cantor_space \ (\\<^sub>M i\ UNIV. restrict_space borel {0,1})" lemma Cantor_space_standard_ne: "standard_borel_ne Cantor_space" by(auto simp: Cantor_space_def intro!: product_standard_borel_ne) lemma Cantor_space_borel: "sets (borel_of Cantor_space_topology) = sets Cantor_space" (is "?lhs = _") proof - have "?lhs = sets (\\<^sub>M i\ UNIV. borel_of (top_of_set {0,1}))" by(auto intro!: sets_PiM_equal_borel_of[symmetric] second_countable_subtopology) thus ?thesis by(simp add: borel_of_subtopology Cantor_space_def borel_of_euclidean) qed text \ Hilbert cube $\mathscr{H}$ \ definition Hilbert_cube :: "(nat \ real) measure" where "Hilbert_cube \ (\\<^sub>M i\ UNIV. restrict_space borel {0..1})" lemma Hilbert_cube_standard_ne: "standard_borel_ne Hilbert_cube" by(auto simp: Hilbert_cube_def intro!: product_standard_borel_ne) lemma Hilbert_cube_borel: "sets (borel_of Hilbert_cube_topology) = sets Hilbert_cube" (is "?lhs = _") proof - have "?lhs = sets (\\<^sub>M i\ UNIV. borel_of (top_of_set {0..1}))" by(auto intro!: sets_PiM_equal_borel_of[symmetric] second_countable_subtopology) thus ?thesis by(simp add: borel_of_subtopology Hilbert_cube_def borel_of_euclidean) qed subsection \ Isomorphism between $\mathscr{C}$ and $\mathscr{H}$\ lemma Cantor_space_isomorphic_to_Hilbert_cube: "Cantor_space measurable_isomorphic Hilbert_cube" proof - text \ Isomorphism between $\mathscr{C}$ and $[0,1]$\ have Cantor_space_isomorphic_to_01closed: "Cantor_space measurable_isomorphic (restrict_space borel {0..1::real})" proof - have space_Cantor_space: "space Cantor_space = (\\<^sub>E i\ UNIV. {0,1})" by(simp add: Cantor_space_def space_PiM) have space_Cantor_space_01[simp]: "0 \ x n" "x n \ 1" "x n \ {0,1}" if "x \ space Cantor_space" for x n using PiE_mem[OF that[simplified space_Cantor_space],of n] by auto have Cantor_minus_abs_cantor: "(\n. \x n - y n\) \ space Cantor_space" if assms:"x \ space Cantor_space" "y \ space Cantor_space" for x y unfolding space_Cantor_space proof safe fix n assume "\x n - y n\ \ 0" then consider "x n = 0 \ y n = 1" | "x n = 1 \ y n = 0" using space_Cantor_space_01[OF assms(1),of n] space_Cantor_space_01[OF assms(2),of n] by auto thus "\x n - y n\ = 1" by cases auto qed simp define Cantor_to_01 :: "(nat \ real) \ real" where "Cantor_to_01 \ (\x. (\n. (1/3)^(Suc n)* x n))" text \ @{term Cantor_to_01} is a measurable injective embedding.\ have Cantor_to_01_summable'[simp]: "summable (\n. (1/3)^(Suc n)* x n)" if "x \ space Cantor_space" for x proof(rule summable_comparison_test'[where g="\n. (1/3)^ n" and N=0]) show "norm ((1 / 3) ^ Suc n * x n) \ (1 / 3) ^ n" for n using space_Cantor_space_01[OF that,of n] by auto qed simp have Cantor_to_01_summable[simp]: "\x. x \ space Cantor_space \ summable (\n. (1/3)^ n* x n)" using Cantor_to_01_summable' by simp have Cantor_to_01_subst_summable[simp]: "summable (\n. (1/3)^ n* (x n - y n))" if assms:"x \ space Cantor_space" "y \ space Cantor_space" for x y proof(rule summable_comparison_test'[where g="\n. (1/3)^ n" and N=0]) show " norm ((1 / 3) ^ n * (x n - y n)) \ (1 / 3) ^ n" for n using space_Cantor_space_01[OF Cantor_minus_abs_cantor[OF assms],of n] by(auto simp: idom_abs_sgn_class.abs_mult) qed simp have Cantor_to_01_image: "Cantor_to_01 \ space Cantor_space \ {0..1}" proof fix x assume h:"x \ space Cantor_space" have "Cantor_to_01 x \ (\n. (1/3)^(Suc n))" unfolding Cantor_to_01_def by(rule suminf_le) (use h Cantor_to_01_summable[OF h] in auto) also have "... = (\n. (1 / 3) ^ n) - (1::real)" using suminf_minus_initial_segment[OF complete_algebra_summable_geometric[of "1/3::real"],of 1] by auto finally have "Cantor_to_01 x \ 1" by(simp add: suminf_geometric[of "1/3"]) moreover have "0 \ Cantor_to_01 x" unfolding Cantor_to_01_def by(rule suminf_nonneg) (use Cantor_to_01_summable[OF h] h in auto) ultimately show "Cantor_to_01 x \ {0..1}" by simp qed have Cantor_to_01_measurable: "Cantor_to_01 \ Cantor_space \\<^sub>M restrict_space borel {0..1}" proof(rule measurable_restrict_space2) show "Cantor_to_01 \ borel_measurable Cantor_space" unfolding Cantor_to_01_def proof(rule borel_measurable_suminf) fix n have "(\x. x n) \ Cantor_space \\<^sub>M restrict_space borel {0, 1}" by(simp add: Cantor_space_def) hence "(\x. x n) \ borel_measurable Cantor_space" by(simp add: measurable_restrict_space2_iff) thus "(\x. (1 / 3) ^ Suc n * x n) \ borel_measurable Cantor_space" by simp qed qed(rule Cantor_to_01_image) have Cantor_to_01_inj: "inj_on Cantor_to_01 (space Cantor_space)" and Cantor_to_01_preserves_sets: "A \ sets Cantor_space \ Cantor_to_01 ` A \ sets (restrict_space borel {0..1})" for A proof - have sets_Cantor: "sets Cantor_space = sets (borel_of (product_topology (\_. subtopology euclidean {0,1}) UNIV))" (is "?lhs = _") proof - have "?lhs = sets (\\<^sub>M i\ UNIV. borel_of (subtopology euclidean {0,1}))" by (simp add: Cantor_space_def borel_of_euclidean borel_of_subtopology) thus ?thesis by(auto intro!: sets_PiM_equal_borel_of second_countable_subtopology polish_space_imp_second_countable[of "euclideanreal"]) qed have s:"space Cantor_space = topspace (product_topology (\_. subtopology euclidean {0,1}) UNIV)" by(simp add: space_Cantor_space) let ?d = "\x y::real. if (x = 0 \ x = 1) \ (y = 0 \ y = 1) then dist x y else 0" interpret d01: Metric_space "{0,1::real}" ?d by(auto simp: Metric_space_def) have d01: "d01.mtopology = top_of_set {0,1}" d01.mcomplete proof - interpret Metric_space "{0,1}" dist by (simp add: Met_TC.subspace) have "d01.mtopology = mtopology" by(auto intro!: Metric_space_eq_mtopology simp: Metric_space_def metric_space_class.dist_commute) also have "... = top_of_set {0,1}" by(auto intro!: Submetric.mtopology_submetric[of UNIV dist "{0,1::real}",simplified] simp: Submetric_def Metric_space_def Submetric_axioms_def dist_real_def) finally show "d01.mtopology = top_of_set {0,1}" . show d01.mcomplete using Metric_space_eq_mcomplete[OF d01.Metric_space_axioms,of dist] d01.compact_space_eq_Bolzano_Weierstrass d01.compact_space_imp_mcomplete finite.emptyI finite_subset by blast qed interpret pd: Product_metric "1/3" UNIV id id "\_. {0,1::real}" "\_. ?d" 1 by(auto intro!: product_metric_natI d01.Metric_space_axioms) have mpd_top: "pd.Product_metric.mtopology = Cantor_space_topology" by(auto simp: pd.Product_metric_mtopology_eq[symmetric] d01 intro!: product_topology_cong) have pd_mcomplete: pd.Product_metric.mcomplete by(auto intro!: pd.mcomplete_Mi_mcomplete_M d01) interpret m01: Submetric UNIV dist "{0..1::real}" by(simp add: Submetric_def Submetric_axioms_def Met_TC.Metric_space_axioms) have "restrict_space borel {0..1} = borel_of m01.sub.mtopology" by (simp add: borel_of_euclidean borel_of_subtopology m01.mtopology_submetric) have pd_def: "pd.product_dist x y = (\n. (1/3)^n * \x n - y n\)" if "x \ space Cantor_space" "y \ space Cantor_space" for x y using space_Cantor_space_01[OF that(1)] space_Cantor_space_01[OF that(2)] that by(auto simp: product_dist_def dist_real_def) have 1: "\Cantor_to_01 x - Cantor_to_01 y\ \ pd.product_dist x y" (is "?lhs \ ?rhs") if "x \ space Cantor_space" "y \ space Cantor_space" for x y proof - have "?lhs = \(\n. (1/3)^(Suc n)* x n - (1/3)^(Suc n)* y n)\" using that by(simp add: suminf_diff Cantor_to_01_def) also have "... = \\n. (1/3)^(Suc n) * (x n - y n) \" by (simp add: right_diff_distrib) also have "... \ (\n. \(1/3)^(Suc n) * (x n - y n)\)" proof(rule summable_rabs) have "(\n. \(1 / 3) ^ Suc n * (x n - y n)\) = (\n. (1 / 3) ^ Suc n * \(x n - y n)\)" by (simp add: abs_mult_pos mult.commute) moreover have "summable ..." using Cantor_minus_abs_cantor[OF that] by simp ultimately show "summable (\n. \(1 / 3) ^ Suc n * (x n - y n)\)" by simp qed also have "... = (\n. (1/3)^(Suc n) * \x n - y n\)" by (simp add: abs_mult_pos mult.commute) also have "... \ pd.product_dist x y" unfolding pd_def[OF that] by(rule suminf_le) (use Cantor_minus_abs_cantor[OF that] in auto) finally show ?thesis . qed have 2:"\Cantor_to_01 x - Cantor_to_01 y\ \ 1 / 9 * pd.product_dist x y" (is "?lhs \ ?rhs") if "x \ space Cantor_space" "y \ space Cantor_space" for x y proof(cases "x = y") case True then show ?thesis using pd.Product_metric.zero[of x y] that by(simp add: space_Cantor_space) next case False then obtain n' where "x n' \ y n'" by auto define n where "n \ Min {n. n \ n' \ x n \ y n}" have "n \ n'" using \x n' \ y n'\ n_def by fastforce have "x n \ y n" using \x n' \ y n'\ linorder_class.Min_in[of "{n. n \ n' \ x n \ y n}"] by(auto simp: n_def) have "\i y i" then have "i \ {n. n \ n' \ x n \ y n}" using \n \ n'\ \i < n\ by auto thus False using \i < n\ linorder_class.Min_gr_iff[of "{n. n \ n' \ x n \ y n}" i] \x n' \ y n'\ by(auto simp: n_def) qed qed have u1: "(1/3) ^ (Suc n) * (1/2) \ \Cantor_to_01 x - Cantor_to_01 y\" proof - have "(1/3) ^ (Suc n) * (1/2) \ \(\m. (1/3)^(Suc (m + Suc n)) * (x (m + Suc n) - y (m + Suc n))) + (1 / 3) ^ Suc n * (x n - y n)\" proof - have "(1 / 3) ^ Suc n - (1/3)^(n + 2) * 3/2 \ (1 / 3) ^ Suc n - \(\m. (1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n)))\" proof - have "\(\m. (1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n)))\ \ (1/3)^(n + 2) * 3/2" (is "?lhs \ _") proof - have "?lhs \ (\m. \(1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n))\)" apply(rule summable_rabs,rule summable_ignore_initial_segment[of _ "Suc n"]) using Cantor_minus_abs_cantor[OF that(2,1)] by(simp add: abs_mult) also have "... = (\m. (1 / 3) ^ Suc (m + Suc n) * \y (m + Suc n) - x (m + Suc n)\)" by(simp add: abs_mult) also have "... \ (\m. (1 / 3) ^ Suc (m + Suc n))" apply(rule suminf_le) using space_Cantor_space_01[OF Cantor_minus_abs_cantor[OF that(2,1)]] apply simp apply(rule summable_ignore_initial_segment[of _ "Suc n"]) using Cantor_minus_abs_cantor[OF that(2,1)] by auto also have "... = (\m. (1 / 3) ^ (m + Suc (Suc n)) * 1)" by simp also have "... = (1/3)^(n + 2) * 3/(2::real)" by(simp only: pd.nsum_of_rK[of "Suc (Suc n)"],simp) finally show ?thesis . qed thus ?thesis by simp qed also have "... = \(1 / 3) ^ Suc n * (x n - y n)\ - \\m. (1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n))\" using \x n \ y n\ space_Cantor_space_01[OF Cantor_minus_abs_cantor[OF that],of n] by(simp add: abs_mult) also have "... \ \(1 / 3) ^ Suc n * (x n - y n) - (\m. (1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n)))\" by simp also have "... = \(1 / 3) ^ Suc n * (x n - y n) + (\m. (1 / 3) ^ Suc (m + Suc n) * (x (m + Suc n) - y (m + Suc n)))\" proof - have "(\m. (1 / 3) ^ Suc (m + Suc n) * (x (m + Suc n) - y (m + Suc n))) = (\m. - ((1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n))))" proof - { fix nn :: nat have "\r ra rb. - ((- (r::real) + ra) / (1 / rb)) = (- ra + r) / (1 / rb)" by (simp add: left_diff_distrib) then have "- ((y (Suc (n + nn)) + - x (Suc (n + nn))) * (1 / 3) ^ Suc (Suc (n + nn))) = (x (Suc (n + nn)) + - y (Suc (n + nn))) * (1 / 3) ^ Suc (Suc (n + nn))" by fastforce then have "- ((1 / 3) ^ Suc (nn + Suc n) * (y (nn + Suc n) - x (nn + Suc n))) = (1 / 3) ^ Suc (nn + Suc n) * (x (nn + Suc n) - y (nn + Suc n))" by (simp add: add.commute mult.commute) } then show ?thesis by presburger qed also have "... = - (\m. (1 / 3) ^ Suc (m + Suc n) * (y (m + Suc n) - x (m + Suc n)))" apply(rule suminf_minus) apply(rule summable_ignore_initial_segment[of _ "Suc n"]) using that by simp finally show ?thesis by simp qed also have "... = \(\m. (1 / 3) ^ Suc (m + Suc n) * (x (m + Suc n) - y (m + Suc n))) + (1 / 3) ^ Suc n * (x n - y n)\" using 1 by simp finally show ?thesis by simp qed also have "... = \(\m. (1/3)^(Suc (m + Suc n)) * (x (m + Suc n) - y (m + Suc n))) + (\m" using \\i by auto also have "... = \\n. (1/3)^(Suc n) * (x n - y n)\" proof - have "(\n. (1 / 3) ^ Suc n * (x n - y n)) = (\m. (1 / 3) ^ Suc (m + Suc n) * (x (m + Suc n) - y (m + Suc n))) + (\m(\n. (1/3)^(Suc n)* x n - (1/3)^(Suc n)* y n)\" by (simp add: right_diff_distrib) also have "... = \Cantor_to_01 x - Cantor_to_01 y\" using that by(simp add: suminf_diff Cantor_to_01_def) finally show ?thesis . qed have u2: "(1/9) * pd.product_dist x y \ (1/3) ^ (Suc n) * (1/2)" proof - have "pd.product_dist x y = (\m. (1/3)^m * \x m - y m\)" by(simp add: that pd_def) also have "... = (\m. (1/3)^(m + n) * \x (m + n) - y (m + n)\) + (\mx m - y m\)" using Cantor_minus_abs_cantor[OF that] by(auto intro!: suminf_split_initial_segment) also have "... = (\m. (1/3)^(m + n) * \x (m + n) - y (m + n)\)" using \\i by simp also have "... \ (\m. (1/3)^(m + n))" using space_Cantor_space_01[OF Cantor_minus_abs_cantor[OF that]] Cantor_minus_abs_cantor[OF that] by(auto intro!: suminf_le summable_ignore_initial_segment[of _ n]) also have "... = (1 / 3) ^ n * (3 / 2)" using pd.nsum_of_rK[of n] by auto finally show ?thesis by auto qed from u1 u2 show ?thesis by simp qed have inj: "inj_on Cantor_to_01 (space Cantor_space)" proof fix x y assume h:"x \ space Cantor_space" "y \ space Cantor_space" "Cantor_to_01 x = Cantor_to_01 y" then have "pd.product_dist x y = 0" using 2[OF h(1,2)] pd.Product_metric.nonneg[of x y] by simp thus "x = y" using pd.Product_metric.zero[of x y] h(1,2) by(simp add: space_Cantor_space) qed have closed: "closedin m01.sub.mtopology (Cantor_to_01 ` (space Cantor_space))" unfolding m01.sub.metric_closedin_iff_sequentially_closed proof safe show "a \ space Cantor_space \ Cantor_to_01 a \ {0..1}" for a using Cantor_to_01_image by auto next fix xn x assume h:"range xn \ Cantor_to_01 ` space Cantor_space" "limitin m01.sub.mtopology xn x sequentially" have "\n. xn n \ {0..1}" using h(1) measurable_space[OF Cantor_to_01_measurable] by (metis (no_types, lifting) UNIV_I atLeastAtMost_borel image_subset_iff space_restrict_space2 subsetD) with h(2) have xnC:"m01.sub.MCauchy xn" by(auto intro!: m01.sub.convergent_imp_MCauchy) have "\n. \x\space Cantor_space. xn n = Cantor_to_01 x" using h(1) by auto then obtain yn where hx:"\n. yn n \ space Cantor_space" "\n. xn n = Cantor_to_01 (yn n)" by metis have "pd.Product_metric.MCauchy yn" unfolding pd.Product_metric.MCauchy_def proof safe fix \ assume "(0 :: real) < \" hence "0 < \ / 9" by auto then obtain N' where "\n m. n \ N' \ m \ N' \ \xn n - xn m\ < \ / 9" using xnC m01.sub.MCauchy_def xnC unfolding dist_real_def by blast thus "\N. \n n'. N \ n \ N \ n' \ pd.product_dist (yn n) (yn n') < \" using order.strict_trans1[OF 2[OF hx(1) hx(1)],of _ _ "\/9"] hx(1) by(auto intro!: exI[where x=N'] simp: hx(2) space_Cantor_space) qed(use hx space_Cantor_space in auto) then obtain y where y:"limitin pd.Product_metric.mtopology yn y sequentially" using pd_mcomplete pd.Product_metric.mcomplete_def by blast hence "y \ space Cantor_space" by (simp add: pd.Product_metric.limitin_mspace space_Cantor_space) have "limitin m01.sub.mtopology xn (Cantor_to_01 y) sequentially" unfolding m01.sub.limit_metric_sequentially proof safe show "Cantor_to_01 y \ {0..1}" using h(1) funcset_image[OF Cantor_to_01_image] \y \ space Cantor_space\ by blast next fix \ assume "(0 :: real) < \" then obtain N where "\n. n \ N \ pd.product_dist (yn n) y < \" "\n. n \ N \ yn n \ UNIV \\<^sub>E {0, 1}" using y by(fastforce simp: pd.Product_metric.limit_metric_sequentially) with \\n. xn n \ {0..1}\ show "\N. \n\N. xn n \ {0..1} \ dist (xn n) (Cantor_to_01 y) < \" by(auto intro!: exI[where x=N] order.strict_trans1[OF 1[OF hx(1) \y \ space Cantor_space\]] simp: submetric_def \0 < \\ hx(2) dist_real_def) qed hence "Cantor_to_01 y = x" using h(2) by(auto dest: m01.sub.limitin_metric_unique) with \y \ space Cantor_space\ show "x \ Cantor_to_01 ` space Cantor_space" by auto qed have open_map:"open_map pd.Product_metric.mtopology (subtopology m01.sub.mtopology (Cantor_to_01 ` (space Cantor_space))) Cantor_to_01" proof - have "open_map (mtopology_of pd.Product_metric.Self) (subtopology (mtopology_of m01.sub.Self) (Cantor_to_01 ` mspace pd.Product_metric.Self)) Cantor_to_01" proof(rule Metric_space_open_map_from_dist) fix x \ assume "(0 :: real) < \" " x \ mspace pd.Product_metric.Self" then have "x \ (UNIV :: nat set) \\<^sub>E {0, 1::real}" by simp show "\\>0. \y\mspace pd.Product_metric.Self. mdist m01.sub.Self (Cantor_to_01 x) (Cantor_to_01 y) < \ \ mdist pd.Product_metric.Self x y < \" unfolding pd.Product_metric.mspace_Self pd.Product_metric.mdist_Self m01.sub.mdist_Self proof(safe intro!: exI[where x="\/9"]) fix y assume h:"y \ (UNIV :: nat set) \\<^sub>E {0, 1::real}" "dist (Cantor_to_01 x) (Cantor_to_01 y) < \ / 9" then have sc:"x \ space Cantor_space" "y \ space Cantor_space" using \x \ UNIV \\<^sub>E {0, 1}\ by(simp_all add: space_Cantor_space) have "\Cantor_to_01 x - Cantor_to_01 y\ < \ / 9" using h by(simp add: dist_real_def) with 2[OF sc] show "pd.product_dist x y < \ " by simp qed (use \\ > 0\ in auto) qed(use Cantor_to_01_image space_Cantor_space in auto) thus ?thesis by (simp add: mtopology_of_def space_Cantor_space) qed have "Cantor_to_01 ` A \ sets (restrict_space borel {0..1})" if "A \ sets Cantor_space" for A using open_map_preserves_sets'[of pd.Product_metric.mtopology m01.sub.mtopology Cantor_to_01 A] borel_of_closed[OF closed] inj sets_Cantor open_map that mpd_top \restrict_space borel {0..1} = borel_of m01.sub.mtopology\ by (simp add: space_Cantor_space) with inj show "inj_on Cantor_to_01 (space Cantor_space)" and"A \ sets Cantor_space \ Cantor_to_01 ` A \ sets (restrict_space borel {0..1})" by simp_all qed text \ Next, we construct measurable embedding from $[0,1]$ to ${0,1}^{\mathbb{N}}$.\ define to_Cantor_from_01 :: "real \ nat \ real" where "to_Cantor_from_01 \ (\r n. if r = 1 then 1 else real_of_int (\2^(Suc n) * r\ mod 2))" text \ @{term to_Cantor_from_01} is a measurable injective embedding into Cantor space.\ have to_Cantor_from_01_image': "to_Cantor_from_01 r n \ {0,1}" for r n unfolding to_Cantor_from_01_def by auto have to_Cantor_from_01_image'': "\r n. 0 \ to_Cantor_from_01 r n" "\r n. to_Cantor_from_01 r n \ 1" by (auto simp add: to_Cantor_from_01_def) have to_Cantor_from_01_image: "to_Cantor_from_01 \ {0..1} \ space Cantor_space" using to_Cantor_from_01_image' by(auto simp: space_Cantor_space) have to_Cantor_from_01_measurable: "to_Cantor_from_01 \ restrict_space borel {0..1} \\<^sub>M Cantor_space" unfolding to_Cantor_from_01_def Cantor_space_def by(auto intro!: measurable_restrict_space3 measurable_abs_UNIV) have to_Cantor_from_01_summable[simp]: "summable (\n. (1/2)^n * to_Cantor_from_01 r n)" for r proof(rule summable_comparison_test'[where g="\n. (1/2)^ n"]) show "norm ((1 / 2) ^ n * to_Cantor_from_01 r n) \ (1 / 2) ^ n" for n using to_Cantor_from_01_image'[of r n] by auto qed simp have to_Cantor_from_sumn': "(\i r" "r - (\i (1/2)^(Suc n) \ r - (\i r - (\i {0..<1}" for r n proof - let ?f = "to_Cantor_from_01 r" have f_simp: "?f l = real_of_int (\ 2^(Suc l) * r\ mod 2)" for l using assms by(simp add: to_Cantor_from_01_def) define S where "S = (\n. \i2^(Suc m) * (l - S m)\ mod 2 = \2^(Suc m) * l\ mod 2" for l m proof - have "\z. 2^(Suc m) * ((1/2)^(Suc k) * ?f k) = 2*real_of_int z" if "k < m" for k proof - have 0:"(2::real) ^ m * (1 / 2) ^ k = 2 * 2^(m-k-1)" using that by (simp add: power_diff_conv_inverse) consider "?f k = 0" | "?f k = 1" using to_Cantor_from_01_image'[of r k] by auto thus ?thesis apply cases using that 0 by auto qed then obtain z where "\k. k < m \ 2^(Suc m) * ((1/2)^(Suc k) * ?f k) = 2*real_of_int (z k)" by metis hence Sm: "2^(Suc m) * S m = real_of_int (2 * (\k2^(Suc m) * (l - S m)\ mod 2 = \2^(Suc m) * l - 2^(Suc m) * S m\ mod 2" by (simp add: right_diff_distrib) also have "... = \2^(Suc m) * l\ mod 2" unfolding Sm by(simp only: floor_diff_of_int) presburger finally show ?thesis . qed have "S n \ r \ r - S n < (1/2)^n \ (?f n = 1 \ (1/2)^(Suc n) \ r - S n) \ (?f n = 0 \ r - S n < (1/2)^(Suc n))" proof(induction n) case 0 then show ?case using assms by(auto simp: S_def to_Cantor_from_01_def) linarith+ next case (Suc n) hence ih: "S n \ r" "r - S n < (1 / 2) ^ n" "?f n = 1 \ (1 / 2) ^ Suc n \ r - S n" "?f n = 0 \ r - S n < (1 / 2) ^ Suc n" by simp_all have SSuc':"?f n = 0 \ S (Suc n) = S n \ ?f n = 1 \ S (Suc n) = S n + (1/2)^(Suc n)" using to_Cantor_from_01_image'[of r n] by(simp add: SSuc) have goal1:"S (Suc n) \ r" using SSuc' ih(1) ih(3) by auto have goal2: "r - S (Suc n) < (1 / 2) ^ Suc n" using SSuc' ih(4) ih(2) by auto have goal3_1: "(1 / 2) ^ Suc (Suc n) \ r - S (Suc n)" if "?f (Suc n) = 1" proof(rule ccontr) assume "\ (1 / 2) ^ Suc (Suc n) \ r - S (Suc n)" then have "r - S (Suc n) < (1 / 2) ^ Suc (Suc n)" by simp hence h:"2 ^ Suc (Suc n) * (r - S (Suc n)) < 1" using mult_less_cancel_left_pos[of "2 ^ Suc (Suc n)" "r - S (Suc n)" "(1 / 2) ^ Suc (Suc n)"] by (simp add: power_one_over) moreover have "0 \ 2 ^ Suc (Suc n) * (r - S (Suc n))" using goal1 by simp ultimately have "\2 ^ Suc (Suc n) * (r - S (Suc n))\ = 0" by linarith thus False using that[simplified f_simp] Sfloor[of "Suc n" r] by fastforce qed have goal3_2: "?f (Suc n) = 1" if "(1 / 2) ^ Suc (Suc n) \ r - S (Suc n)" proof - have "1 \ 2 ^ Suc (Suc n) * (r - S (Suc n))" using that[simplified f_simp] mult_le_cancel_left_pos[of "2 ^ Suc (Suc n)" "(1 / 2) ^ Suc (Suc n)" "r - S (Suc n)"] by (simp add: power_one_over) moreover have "2 ^ Suc (Suc n) * (r - S (Suc n)) < 2" using mult_less_cancel_left_pos[of "2 ^ Suc (Suc n)" "r - S (Suc n)" "(1 / 2) ^ Suc n"] goal2 by (simp add: power_one_over) ultimately have "\2 ^ Suc (Suc n) * (r - S (Suc n))\ = 1" by linarith thus ?thesis using Sfloor[of "Suc n" r] by(auto simp: f_simp) qed have goal4_1: "r - S (Suc n) < (1 / 2) ^ Suc (Suc n)" if "?f (Suc n) = 0" proof(rule ccontr) assume "\ r - S (Suc n) < (1 / 2) ^ Suc (Suc n)" then have "(1 / 2) ^ Suc (Suc n) \ r - S (Suc n)" by simp hence "1 \ 2 ^ Suc (Suc n) * (r - S (Suc n))" using mult_le_cancel_left_pos[of "2 ^ Suc (Suc n)" "(1 / 2) ^ Suc (Suc n)" "r - S (Suc n)"] by (simp add: power_one_over) moreover have "2 ^ Suc (Suc n) * (r - S (Suc n)) < 2" using mult_less_cancel_left_pos[of "2 ^ Suc (Suc n)" "r - S (Suc n)" "(1 / 2) ^ Suc n"] goal2 by (simp add: power_one_over) ultimately have "\2 ^ Suc (Suc n) * (r - S (Suc n))\ = 1" by linarith thus False using that Sfloor[of "Suc n" r] by(auto simp: f_simp) qed have goal4_2: "?f (Suc n) = 0" if "r - S (Suc n) < (1 / 2) ^ Suc (Suc n)" proof - have h:"2 ^ Suc (Suc n) * (r - S (Suc n)) < 1" using mult_less_cancel_left_pos[of "2 ^ Suc (Suc n)" "r - S (Suc n)" "(1 / 2) ^ Suc (Suc n)"] that by (simp add: power_one_over) moreover have "0 \ 2 ^ Suc (Suc n) * (r - S (Suc n))" using goal1 by simp ultimately have "\2 ^ Suc (Suc n) * (r - S (Suc n))\ = 0" by linarith thus ?thesis using Sfloor[of "Suc n" r] by(auto simp: f_simp) qed show ?case using goal1 goal2 goal3_1 goal3_2 goal4_1 goal4_2 by blast qed thus "(\i r" and "r - (\i (1/2)^(Suc n) \ r - (\i r - (\ii r" "r - (\i (1/2)^n" "to_Cantor_from_01 r n = 1 \ (1/2)^(Suc n) \ r - (\i r - (\i {0..1}" for r n proof - have nsum:"(\i{0..<1}" using assms by fastforce hence "(\i r \ r - (\i (1/2)^n \ (to_Cantor_from_01 r n = 1 \ (1/2)^(Suc n) \ r - (\i (to_Cantor_from_01 r n = 0 \ r - (\ii r" and "r - (\i (1/2)^n" and "to_Cantor_from_01 r n = 1 \ (1/2)^(Suc n) \ r - (\i r - (\in. (1/2)^(Suc n)*to_Cantor_from_01 r n) = r" if assms:"r \ {0..1}" for r proof - have 1:"r \ (\n. (1/2)^(Suc n)*to_Cantor_from_01 r n)" proof - have 0:"r \ (1 / 2) ^ n + (\n. (1/2)^(Suc n)*to_Cantor_from_01 r n)" for n proof - have "r \ (1 / 2) ^ n + (\i (1 / 2) ^ n + (\n. (1/2)^(Suc n)*to_Cantor_from_01 r n)" using to_Cantor_from_01_image''[of r] by(auto intro!: sum_le_suminf) finally show ?thesis . qed have 00:"\no. \n\no. (1 / 2) ^ n < r" if "r>0" for r :: real proof - obtain n0 where "(1 / 2) ^ n0 < r" using reals_power_lt_ex[of _ "2 :: real",OF \r>0\] by auto thus ?thesis using order.strict_trans1[OF power_decreasing[of n0 _ "1/2::real"]] by(auto intro!: exI[where x=n0]) qed show ?thesis apply(rule Lim_bounded2[where f="\n. (1 / 2) ^ n + (\n. (1/2)^(Suc n)*to_Cantor_from_01 r n)" and N=0]) using 0 00 by(auto simp: LIMSEQ_iff) qed have 2:"(\n. (1/2)^(Suc n)*to_Cantor_from_01 r n) \ r" using to_Cantor_from_sumn[OF assms] by(auto intro!: suminf_le_const) show ?thesis using 1 2 by simp qed have to_Cantor_from_sum': "(\im. (1/2)^(Suc (m + n))*to_Cantor_from_01 r (m + n))" if assms:"r \ {0..1}" for r n using suminf_minus_initial_segment[of "\n. (1 / 2) ^ Suc n * to_Cantor_from_01 r n" n] to_Cantor_from_sum[OF assms] by auto have to_Cantor_from_01_exist0: "\n.\k\n. to_Cantor_from_01 r k = 0" if assms:"r \ {0..<1}" for r proof(rule ccontr) assume "\ (\n.\k\n. to_Cantor_from_01 r k = 0)" then obtain n0 where hn0: "\k. k \ n0 \ to_Cantor_from_01 r k = 1" using to_Cantor_from_01_image'[of r] by auto define n where "n = Min {i. i \ n0 \ (\k\i. to_Cantor_from_01 r k = 1)}" have n0in: "n0 \ {i. i \ n0 \ (\k\i. to_Cantor_from_01 r k = 1)}" using hn0 by auto have hn:"n \ n0" "\k. k \ n \ to_Cantor_from_01 r k = 1" using n0in Min_in[of "{i. i \ n0 \ (\k\i. to_Cantor_from_01 r k = 1)}"] by(auto simp: n_def) show False proof(cases n) case 0 then have "r = (\n. (1 / 2) ^ Suc n)" using to_Cantor_from_sum[of r] assms hn(2) by simp also have "... = 1" using nsum_of_r'[of "1/2" 1 1] by auto finally show ?thesis using assms by auto next case eqn:(Suc n') have "to_Cantor_from_01 r n' = 0" proof(rule ccontr) assume "to_Cantor_from_01 r n' \ 0" then have "to_Cantor_from_01 r n' = 1" using to_Cantor_from_01_image'[of r n'] by auto hence "n' \ {i. i \ n0 \ (\k\i. to_Cantor_from_01 r k = 1)}" using hn eqn not_less_eq_eq order_antisym_conv by fastforce hence "n \ n'" using Min.coboundedI[of "{i. i \ n0 \ (\k\i. to_Cantor_from_01 r k = 1)}" n'] by(simp add: n_def) thus False using eqn by simp qed hence le1:"r - (\iim. (1/2)^(m + Suc n')*to_Cantor_from_01 r (m + n'))" using to_Cantor_from_sum'[of r n'] assms by simp also have "... = (\m. (1/2)^(m + Suc n)*to_Cantor_from_01 r (m + n))" proof - have "(\n. (1 / 2) ^ (Suc n + Suc n') * to_Cantor_from_01 r (Suc n + n')) = (\m. (1 / 2) ^ (m + Suc n') * to_Cantor_from_01 r (m + n')) - (1 / 2) ^ (0 + Suc n') * to_Cantor_from_01 r (0 + n')" by(rule suminf_split_head) (auto intro!: summable_ignore_initial_segment) thus ?thesis using \to_Cantor_from_01 r n' = 0\ by(simp add: eqn) qed also have "... = (\m. (1/2)^(m + Suc n))" using hn by simp also have "... = (1 / 2) ^ n" using nsum_of_r'[of "1/2" "Suc n" 1,simplified] by simp finally show ?thesis . qed with le1 show False by simp qed qed have to_Cantor_from_01_if_exist0: "to_Cantor_from_01 (\n. (1 / 2) ^ Suc n * a n) = a" if assms:"\n. a n \ {0,1}" "\n.\k\n. a k = 0" for a proof fix n have [simp]: "summable (\n. (1 / 2) ^ n * a n)" proof(rule summable_comparison_test'[where g="\n. (1/2)^ n"]) show "norm ((1 / 2) ^ n * a n) \ (1 / 2) ^ n" for n using assms(1)[of n] by auto qed simp let ?r = "\n. (1 / 2) ^ Suc n * a n" have "?r \ {0..1}" using assms(1) space_Cantor_space_01[of a,simplified space_Cantor_space] nsum_of_r_leq[of "1/2" a 1 1 0] by auto show "to_Cantor_from_01 ?r n = a n" proof(rule less_induct) fix x assume ih:"y < x \ to_Cantor_from_01 ?r y = a y" for y have eq1:"?r - (\in. (1/2)^(Suc (n + x))* a (n + x))" (is "?lhs = ?rhs") proof - have "?lhs = (\n. (1 / 2) ^ Suc (n + x) * a (n + x)) + (\iin. (1 / 2) ^ (Suc n) * a n" x] by simp also have "... = (\n. (1 / 2) ^ Suc (n + x) * a (n + x)) + (\iin. (1/2)^(Suc (n + x))* a (n + x))" define Sn' where "Sn' = (\n. (1/2)^(Suc (n + (Suc x)))* a (n + (Suc x)))" have SnSn':"Sn = (1/2)^(Suc x) * a x + Sn'" using suminf_split_head[of "\n. (1/2)^(Suc (n + x))* a (n + x)",OF summable_ignore_initial_segment] by(auto simp: Sn_def Sn'_def) have hsn:"0 \ Sn'" "Sn' < (1/2)^(Suc x)" proof - show "0 \ Sn'" unfolding Sn'_def by(rule suminf_nonneg,rule summable_ignore_initial_segment) (use assms(1) space_Cantor_space_01[of a,simplified space_Cantor_space] in fastforce)+ next have "\n'\Suc x. a n' < 1" using assms by fastforce thus "Sn' < (1/2)^(Suc x)" using nsum_of_r_le[of "1/2" a 1 "Suc x" "Suc (Suc x)"] assms(1) space_Cantor_space_01[of a,simplified space_Cantor_space] by(auto simp: Sn'_def) qed have goal1: "to_Cantor_from_01 ?r x = 1 \ a x = 1" proof - have "to_Cantor_from_01 ?r x = 1 \ (1 / 2) ^ Suc x \ Sn" using to_Cantor_from_sumn(3)[OF \?r \ {0..1}\] eq1 by(fastforce simp: Sn_def) also have "... \ (1 / 2) ^ Suc x \ (1/2)^(Suc x) * a x + Sn'" by(simp add: SnSn') also have "... \ a x = 1" proof - have "a x = 1" if "(1 / 2) ^ Suc x \ (1/2)^(Suc x) * a x + Sn'" proof(rule ccontr) assume "a x \ 1" then have "a x = 0" using assms(1) by auto hence "(1 / 2) ^ Suc x \ Sn'" using that by simp thus False using hsn by auto qed thus ?thesis by(auto simp: hsn) qed finally show ?thesis . qed have goal2: "to_Cantor_from_01 ?r x = 0 \ a x = 0" proof - have "to_Cantor_from_01 ?r x = 0 \ Sn < (1 / 2) ^ Suc x" using to_Cantor_from_sumn(4)[OF \?r \ {0..1}\] eq1 by(fastforce simp: Sn_def) also have "... \ (1/2)^(Suc x) * a x + Sn' < (1 / 2) ^ Suc x" by(simp add: SnSn') also have "... \ a x = 0" proof - have "a x = 0" if "(1/2)^(Suc x) * a x + Sn' < (1 / 2) ^ Suc x" proof(rule ccontr) assume "a x \ 0" then have "a x = 1" using assms(1) by auto thus False using that hsn by auto qed thus ?thesis using hsn by auto qed finally show ?thesis . qed show "to_Cantor_from_01 ?r x = a x" using goal1 goal2 to_Cantor_from_01_image'[of ?r x] by auto qed qed have to_Cantor_from_01_sum_of_to_Cantor_from_01: "to_Cantor_from_01 (\n. (1 / 2) ^ Suc n * to_Cantor_from_01 r n) = to_Cantor_from_01 r" if assms:"r \ {0..1}" for r proof - consider "r = 1" | "r \ {0..<1}" using assms by fastforce then show ?thesis proof cases case 1 then show ?thesis using nsum_of_r'[of "1/2" 1 1] by(auto simp: to_Cantor_from_01_def) next case 2 from to_Cantor_from_01_if_exist0[OF to_Cantor_from_01_image' to_Cantor_from_01_exist0[OF this]] show ?thesis . qed qed have to_Cantor_from_01_inj: "inj_on to_Cantor_from_01 (space (restrict_space borel {0..1}))" proof fix x y :: real assume "x \ space (restrict_space borel {0..1})" "y \ space (restrict_space borel {0..1})" and h:"to_Cantor_from_01 x = to_Cantor_from_01 y" then have xyin:"x \ {0..1}" "y \ {0..1}" by simp_all show "x = y" using to_Cantor_from_sum[OF xyin(1)] to_Cantor_from_sum[OF xyin(2)] h by simp qed have to_Cantor_from_01_preserves_sets: "to_Cantor_from_01 ` A \ sets Cantor_space" if assms: "A \ sets (restrict_space borel {0..1})" for A proof - define f :: "(nat \ real) \ real" where "f \ (\x. \n. (1/2)^(Suc n)* x n)" have f_meas:"f \ Cantor_space \\<^sub>M restrict_space borel {0..1}" proof - have "f \ borel_measurable Cantor_space" unfolding Cantor_to_01_def f_def proof(rule borel_measurable_suminf) fix n have "(\x. x n) \ Cantor_space \\<^sub>M restrict_space borel {0, 1}" by(simp add: Cantor_space_def) hence "(\x. x n) \ borel_measurable Cantor_space" by(simp add: measurable_restrict_space2_iff) thus "(\x. (1 / 2) ^ Suc n * x n) \ borel_measurable Cantor_space" by simp qed moreover have "0 \ f x" "f x \ 1" if "x \ space Cantor_space" for x proof - have [simp]:"summable (\n. (1/2)^n* x n)" proof(rule summable_comparison_test'[where g="\n. (1/2)^ n"]) show "norm ((1 / 2) ^ n * x n) \ (1 / 2) ^ n" for n using that by simp qed simp show "0 \ f x" using that by(auto intro!: suminf_nonneg simp: f_def) show "f x \ 1" proof - have "f x \ (\n. (1/2)^(Suc n))" using that by(auto intro!: suminf_le simp: f_def) also have "... = 1" using nsum_of_r'[of "1/2" 1 1] by simp finally show ?thesis . qed qed ultimately show ?thesis by(auto intro!: measurable_restrict_space2) qed have image_sets:"to_Cantor_from_01 ` (space (restrict_space borel {0..1})) \ sets Cantor_space" (is "?A \ _") proof - have "?A \ space Cantor_space" using to_Cantor_from_01_image by auto have comple_sets:"(\\<^sub>E i\ UNIV. {0,1}) - ?A \ sets Cantor_space" proof - have eq1:"?A = {\n. 1} \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 0)}" proof show "?A \ {\n. 1} \ {x. (\n. x n \ {0, 1}) \ (\n. \k\n. x k = 0)}" proof fix x assume "x \ ?A" then obtain r where hr:"r \ {0..1}" "x = to_Cantor_from_01 r" by auto then consider "r = 1" | "r \ {0..<1}" by fastforce thus "x \ {\n. 1} \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 0)}" proof cases case 1 then show ?thesis by(simp add: hr(2) to_Cantor_from_01_def) next case 2 from to_Cantor_from_01_exist0[OF this] to_Cantor_from_01_image' show ?thesis by(auto simp: hr(2)) qed qed next show "{\n. 1} \ {x. (\n. x n \ {0, 1}) \ (\n. \k\n. x k = 0)} \ ?A" proof fix x :: "nat \ real" assume "x \ {\n. 1} \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 0)}" then consider "x = (\n. 1)" | "(\n. x n \ {0,1}) \ (\n. \k\n. x k = 0)" by auto thus "x \ ?A" proof cases case 1 then show ?thesis by(auto intro!: image_eqI[where x=1] simp: to_Cantor_from_01_def) next case 2 hence "\n. 0 \ x n" "\n. x n \ 1" by (metis dual_order.refl empty_iff insert_iff zero_less_one_class.zero_le_one)+ with 2 to_Cantor_from_01_if_exist0[of x] nsum_of_r_leq[of "1/2" x 1 1 0] show ?thesis by(auto intro!: image_eqI[where x="\n. (1 / 2) ^ Suc n * x n"]) qed qed qed have "(\\<^sub>E i\ UNIV. {0,1}) - ?A = {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 1)} - {\n. 1}" proof show "(\\<^sub>E i\ UNIV. {0,1}) - ?A \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 1)} - {\n. 1}" proof fix x :: "nat \ real" assume "x \ (\\<^sub>E i\ UNIV. {0,1}) - ?A" then have "\n. x n \ {0,1}" "\ (\n. \k\n. x k = 0)" "x \ (\n. 1)" using eq1 by blast+ thus "x \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 1)} - {\n. 1}" by blast qed next show "(\\<^sub>E i\ UNIV. {0,1}) - ?A \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 1)} - {\n. 1}" proof fix x :: "nat \ real" assume h:"x \ {x. (\n. x n \ {0,1}) \ (\n. \k\n. x k = 1)} - {\n. 1}" then have "\n. x n \ {0,1}" "\n. \k\n. x k = 1" "x \ (\n. 1)" by blast+ hence "\ (\n. \k\n. x k = 0)" by fastforce with \\n. x n \ {0,1}\ \x \ (\n. 1)\ show "x \ (\\<^sub>E i\ UNIV. {0,1}) - ?A" using eq1 by blast qed qed also have "... = (\ ((\n. {x. (\n. x n \ {0,1}) \ (\k\n. x k = 1)}) ` UNIV)) - {\n. 1}" by blast also have "... \ sets Cantor_space" (is "?B \ _") proof - have "countable ?B" proof - have "countable {x :: nat \ real. (\n. x n = 0 \ x n = 1) \ (\k\m. x k = 1)}" for m :: nat proof - let ?C = "{x::nat \ real. (\n. x n = 0 \ x n = 1) \ (\k\m. x k = 1)}" define g where "g = (\(x::nat \ real) n. if n < m then x n else undefined)" have 1:"g ` ?C = (\\<^sub>E i \{.. g ` ?C" then show "x \ (\\<^sub>E i \{.. (\\<^sub>E i \{..n. if n < m then x n else 1)" by(auto simp add: g_def PiE_def extensional_def) moreover have "(\n. if n < m then x n else 1) \ ?C" using h by auto ultimately show "x \ g ` ?C" by auto qed have 2:"inj_on g ?C" proof fix x y assume hxyg:"x \ ?C" "y : ?C" "g x = g y" show "x = y" proof fix n consider "n < m" | "m \ n" by fastforce thus "x n = y n" proof cases case 1 then show ?thesis using fun_cong[OF hxyg(3),of n] by(simp add: g_def) next case 2 then show ?thesis using hxyg(1,2) by auto qed qed qed show "countable {x::nat \ real. (\n. x n = 0 \ x n = 1) \ (\k\m. x k = 1)}" by(rule countable_image_inj_on[OF _ 2]) (auto intro!: countable_PiE simp: 1) qed thus ?thesis by auto qed moreover have "?B \ space Cantor_space" by(auto simp: space_Cantor_space) ultimately show ?thesis using Cantor_space_standard_ne by(simp add: standard_borel.countable_sets standard_borel_ne_def) qed finally show ?thesis . qed moreover have "space Cantor_space - ((\\<^sub>E i\ UNIV. {0,1}) - ?A) = ?A" using \?A \ space Cantor_space\ space_Cantor_space by blast ultimately show ?thesis using sets.compl_sets[OF comple_sets] by auto qed have "to_Cantor_from_01 ` A = f -` A \ to_Cantor_from_01 ` (space (restrict_space borel {0..1}))" proof show "to_Cantor_from_01 ` A \ f -` A \ to_Cantor_from_01 ` space (restrict_space borel {0..1})" proof fix x assume "x \ to_Cantor_from_01 ` A" then obtain a where ha:"a \ A" "x = to_Cantor_from_01 a" by auto hence "a \ {0..1}" using sets.sets_into_space[OF assms] by auto have "f x = a" using to_Cantor_from_sum[OF \a \ {0..1}\] by(simp add: f_def ha(2)) thus " x \ f -` A \ to_Cantor_from_01 ` space (restrict_space borel {0..1})" using sets.sets_into_space[OF assms] ha by auto qed next show "to_Cantor_from_01 ` A \ f -` A \ to_Cantor_from_01 ` space (restrict_space borel {0..1})" proof fix x assume h:"x \ f -` A \ to_Cantor_from_01 ` space (restrict_space borel {0..1})" then obtain r where "r \ {0..1}" "x = to_Cantor_from_01 r" by auto from h have "f x \ A" by simp hence "to_Cantor_from_01 (f x) = x" using to_Cantor_from_01_sum_of_to_Cantor_from_01[OF \r \ {0..1}\] by(simp add: f_def \x = to_Cantor_from_01 r\) with \f x \ A\ show "x \ to_Cantor_from_01 ` A" by (simp add: rev_image_eqI) qed qed also have "... \ sets Cantor_space" proof - have " f -` A \ space Cantor_space \ to_Cantor_from_01 ` space (restrict_space borel {0..1}) = f -` A \ to_Cantor_from_01 ` (space (restrict_space borel {0..1}))" using to_Cantor_from_01_image sets.sets_into_space[OF assms,simplified] by auto thus ?thesis using sets.Int[OF measurable_sets[OF f_meas assms] image_sets] by fastforce qed finally show ?thesis . qed show ?thesis using Schroeder_Bernstein_measurable[OF Cantor_to_01_measurable Cantor_to_01_preserves_sets Cantor_to_01_inj to_Cantor_from_01_measurable to_Cantor_from_01_preserves_sets to_Cantor_from_01_inj] by(simp add: measurable_isomorphic_def) qed have 1:"Cantor_space measurable_isomorphic (\\<^sub>M (i::nat,j::nat)\ UNIV \ UNIV. restrict_space borel {0,1::real})" unfolding Cantor_space_def by(auto intro!: measurable_isomorphic_sym[OF countable_infinite_isomorphisc_to_nat_index] simp: split_beta' finite_prod) have 2:"(\\<^sub>M (i::nat,j::nat)\ UNIV \ UNIV. restrict_space borel {0,1::real}) measurable_isomorphic (\\<^sub>M (i::nat)\ UNIV. Cantor_space)" unfolding Cantor_space_def by(rule measurable_isomorphic_sym[OF PiM_PiM_isomorphic_to_PiM]) have 3:"(\\<^sub>M (i::nat)\ UNIV. Cantor_space) measurable_isomorphic Hilbert_cube" unfolding Hilbert_cube_def by(rule measurable_isomorphic_lift_product[OF Cantor_space_isomorphic_to_01closed]) show ?thesis by(rule measurable_isomorphic_trans[OF measurable_isomorphic_trans[OF 1 2] 3]) qed subsection \ Final Results \ lemma(in standard_borel) embedding_into_Hilbert_cube: "\A \ sets Hilbert_cube. M measurable_isomorphic (restrict_space Hilbert_cube A)" proof - obtain S where S:"polish_space S" "sets (borel_of S) = sets M" using polish_space by blast obtain A where A:"gdelta_in Hilbert_cube_topology A" "S homeomorphic_space subtopology Hilbert_cube_topology A" using embedding_into_Hilbert_cube_gdelta_in[OF S(1)] by blast show ?thesis using borel_of_gdelta_in[OF A(1)] homeomorphic_space_measurable_isomorphic[OF A(2)] measurable_isomorphic_sets_cong[OF S(2),of "borel_of (subtopology Hilbert_cube_topology A)" "restrict_space Hilbert_cube A"] Hilbert_cube_borel sets_restrict_space_cong[OF Hilbert_cube_borel] by(auto intro!: bexI[where x=A] simp: borel_of_subtopology) qed lemma(in standard_borel) embedding_from_Cantor_space: assumes "uncountable (space M)" shows "\A \ sets M. Cantor_space measurable_isomorphic (restrict_space M A)" proof - obtain S where S:"polish_space S" "sets (borel_of S) = sets M" using polish_space by blast then obtain A where A:"gdelta_in S A" "Cantor_space_topology homeomorphic_space subtopology S A" using embedding_from_Cantor_space[of S] assms sets_eq_imp_space_eq[OF S(2)] by(auto simp: space_borel_of) show ?thesis using borel_of_gdelta_in[OF A(1)] S(2) homeomorphic_space_measurable_isomorphic[OF A(2)] measurable_isomorphic_sets_cong[OF Cantor_space_borel restrict_space_sets_cong[OF refl S(2)],of A] by(auto intro!: bexI[where x=A] simp: borel_of_subtopology) qed corollary(in standard_borel) uncountable_isomorphic_to_Hilbert_cube: assumes "uncountable (space M)" shows "Hilbert_cube measurable_isomorphic M" proof - obtain A B where AB: "M measurable_isomorphic (restrict_space Hilbert_cube A)" "Cantor_space measurable_isomorphic (restrict_space M B)" "A \ sets Hilbert_cube""B \ sets M" using embedding_into_Hilbert_cube embedding_from_Cantor_space[OF assms] by auto show ?thesis by(rule measurable_isomorphic_antisym[OF AB measurable_isomorphic_sym[OF Cantor_space_isomorphic_to_Hilbert_cube]]) qed corollary(in standard_borel) uncountable_isomorphic_to_real: assumes "uncountable (space M)" shows "M measurable_isomorphic (borel :: real measure)" proof - interpret r: standard_borel_ne "borel :: real measure" by simp show ?thesis by(auto intro!: measurable_isomorphic_trans[OF measurable_isomorphic_sym[OF uncountable_isomorphic_to_Hilbert_cube[OF assms]] r.uncountable_isomorphic_to_Hilbert_cube] simp: uncountable_UNIV_real) qed theorem Borel_isomorphism_theorem: assumes "standard_borel M" "standard_borel N" shows "space M \ space N \ M measurable_isomorphic N" proof assume h:"space M \ space N" interpret M: standard_borel M by fact interpret N: standard_borel N by fact consider "countable (space M)" "countable (space N)" | "uncountable (space M)" "uncountable (space N)" by (meson countable_eqpoll eqpoll_sym h) thus "M measurable_isomorphic N" proof cases case 1 then have 2:"sets M = sets (count_space (space M))" "sets N = sets (count_space (space N))" by (simp_all add: M.countable_discrete_space N.countable_discrete_space) show ?thesis by(simp add: measurable_isomorphic_sets_cong[OF 2] measurable_isomorphic_count_spaces h) next case 2 then have "M measurable_isomorphic (borel :: real measure)" "N measurable_isomorphic (borel :: real measure)" by(simp_all add: M.uncountable_isomorphic_to_real N.uncountable_isomorphic_to_real) thus ?thesis using measurable_isomorphic_sym measurable_isomorphic_trans by blast qed qed(rule measurable_isomorphic_cardinality_eq) definition to_real_on :: "'a measure \ 'a \ real" where "to_real_on M \ (if uncountable (space M) then (SOME f. measurable_isomorphic_map M (borel :: real measure) f) else (real \ to_nat_on (space M)))" definition from_real_into :: "'a measure \ real \ 'a" where "from_real_into M \ (if uncountable (space M) then the_inv_into (space M) (to_real_on M) else (\r. from_nat_into (space M) (nat \r\)))" context standard_borel begin abbreviation "to_real \ to_real_on M" abbreviation "from_real \ from_real_into M" lemma to_real_def_countable: assumes "countable (space M)" shows "to_real = (\r. real (to_nat_on (space M) r))" using assms by(auto simp: to_real_on_def) lemma from_real_def_countable: assumes "countable (space M)" shows "from_real = (\r. from_nat_into (space M) (nat \r\))" using assms by(simp add: from_real_into_def) lemma from_real_to_real[simp]: assumes "x \ space M" shows "from_real (to_real x) = x" proof - have [simp]: "space M \ {}" using assms by auto consider "countable (space M)" | "uncountable (space M)" by auto then show ?thesis proof cases case 1 then show ?thesis by(simp add: to_real_def_countable from_real_def_countable assms) next case 2 then obtain f where f: "measurable_isomorphic_map M (borel :: real measure) f" using uncountable_isomorphic_to_real by(auto simp: measurable_isomorphic_def) have 1:"to_real = Eps (measurable_isomorphic_map M borel)" "from_real = the_inv_into (space M) (Eps (measurable_isomorphic_map M borel))" by(simp_all add: to_real_on_def 2 from_real_into_def) show ?thesis unfolding 1 by(rule someI2[of "measurable_isomorphic_map M (borel :: real measure)" f,OF f]) (meson assms bij_betw_imp_inj_on measurable_isomorphic_map_def the_inv_into_f_f) qed qed lemma to_real_measurable[measurable]: "to_real \ M \\<^sub>M borel" proof(cases "countable (space M)") case 1:True then have "sets M = Pow (space M)" by(rule countable_discrete_space) then show ?thesis by(simp add: to_real_def_countable 1 borel_measurableI_le) next case 1:False then obtain f where f: "measurable_isomorphic_map M (borel :: real measure) f" using uncountable_isomorphic_to_real by(auto simp: measurable_isomorphic_def) have 2:"to_real = Eps (measurable_isomorphic_map M borel)" by(simp add: to_real_on_def 1 from_real_into_def) show ?thesis unfolding 2 by(rule someI2[of "measurable_isomorphic_map M (borel :: real measure)" f,OF f],simp add: measurable_isomorphic_map_def) qed lemma from_real_measurable': assumes "space M \ {}" shows "from_real \ borel \\<^sub>M M" proof(cases "countable (space M)") case 1:True then have 2:"sets M = Pow (space M)" by(rule countable_discrete_space) have [measurable]:"from_nat_into (space M) \ count_space UNIV \\<^sub>M M" using from_nat_into[OF assms] by auto show ?thesis by(simp add: from_real_def_countable 1 borel_measurableI_le) next case 2:False then obtain f where f: "measurable_isomorphic_map M (borel :: real measure) f" using uncountable_isomorphic_to_real by(auto simp: measurable_isomorphic_def) have 1: "from_real = the_inv_into (space M) (Eps (measurable_isomorphic_map M borel))" by(simp add: to_real_on_def 2 from_real_into_def) show ?thesis unfolding 1 by(rule someI2[of "measurable_isomorphic_map M (borel :: real measure)" f,OF f],simp add: measurable_isomorphic_map_def) qed lemma countable_isomorphic_to_subset_real: assumes "countable (space M)" obtains A :: "real set" where "countable A" "A \ sets borel" "M measurable_isomorphic (restrict_space borel A)" proof(cases "space M = {}") case True then show ?thesis by (meson countable_empty measurable_isomorphic_empty sets.empty_sets space_restrict_space2 that) next case nin:False define A where "A \ to_real ` (space M)" have "countable A" "A \ sets borel" "M measurable_isomorphic (restrict_space borel A)" proof - show "countable A" "A \ sets borel" using assms(1) standard_borel.countable_sets[of borel A] standard_borel_ne_borel by(auto simp: A_def standard_borel_ne_def) show "M measurable_isomorphic restrict_space borel A" using from_real_to_real A_def \A \ sets borel\ by(auto intro!: measurable_isomorphic_byWitness[OF measurable_restrict_space2[OF _ to_real_measurable] _ measurable_restrict_space1[OF from_real_measurable'[OF nin]]]) qed with that show ?thesis by auto qed lemma to_real_from_real: assumes "uncountable (space M)" shows "to_real (from_real r) = r" proof - obtain f where f: "measurable_isomorphic_map M (borel :: real measure) f" using assms uncountable_isomorphic_to_real by(auto simp: measurable_isomorphic_def) have 1:"to_real = Eps (measurable_isomorphic_map M borel)" "from_real = the_inv_into (space M) (Eps (measurable_isomorphic_map M borel))" by(simp_all add: to_real_on_def assms from_real_into_def) show ?thesis unfolding 1 by(rule someI2[of "measurable_isomorphic_map M (borel :: real measure)" f,OF f]) (metis UNIV_I f_the_inv_into_f_bij_betw measurable_isomorphic_map_def space_borel) qed end lemma(in standard_borel_ne) from_real_measurable[measurable]: "from_real \ borel \\<^sub>M M" by(simp add: from_real_measurable' space_ne) end \ No newline at end of file