# HG changeset patch # User Andreas Lochbihler # Date 1416568304 -3600 # Node ID 4999a616336c420c7afc139f2991c26221d576df # Parent fa7c419f04b48acc371100a7ff79e9c95b6e1bb2 register pmf as BNF diff -r fa7c419f04b4 -r 4999a616336c src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Thu Nov 20 17:29:18 2014 +0100 +++ b/src/HOL/Library/Extended_Real.thy Fri Nov 21 12:11:44 2014 +0100 @@ -1178,6 +1178,11 @@ a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) +lemma ereal_add_le_add_iff2: + fixes a b c :: ereal + shows "a + c \ b + c \ a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" +by(cases rule: ereal3_cases[of a b c])(simp_all add: field_simps) + lemma ereal_mult_le_mult_iff: fixes a b c :: ereal shows "\c\ \ \ \ c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" diff -r fa7c419f04b4 -r 4999a616336c src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Thu Nov 20 17:29:18 2014 +0100 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Fri Nov 21 12:11:44 2014 +0100 @@ -1951,6 +1951,20 @@ shows "(\\<^sup>+x. f x \count_space X) = (\\<^sup>+x. f x * indicator X x \count_space UNIV)" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) +lemma nn_integral_ge_point: + assumes "x \ A" + shows "p x \ \\<^sup>+ x. p x \count_space A" +proof - + from assms have "p x \ \\<^sup>+ x. p x \count_space {x}" + by(auto simp add: nn_integral_count_space_finite max_def) + also have "\ = \\<^sup>+ x'. p x' * indicator {x} x' \count_space A" + using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) + also have "\ \ \\<^sup>+ x. max 0 (p x) \count_space A" + by(rule nn_integral_mono)(simp add: indicator_def) + also have "\ = \\<^sup>+ x. p x \count_space A" by(simp add: nn_integral_def o_def) + finally show ?thesis . +qed + subsubsection {* Measure spaces with an associated density *} definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where diff -r fa7c419f04b4 -r 4999a616336c src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Thu Nov 20 17:29:18 2014 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Fri Nov 21 12:11:44 2014 +0100 @@ -1,5 +1,7 @@ (* Title: HOL/Probability/Probability_Mass_Function.thy - Author: Johannes Hölzl, TU München *) + Author: Johannes Hölzl, TU München + Author: Andreas Lochbihler, ETH Zurich +*) section \ Probability mass function \ @@ -133,7 +135,7 @@ declare [[coercion set_pmf]] -lemma countable_set_pmf: "countable (set_pmf p)" +lemma countable_set_pmf [simp]: "countable (set_pmf p)" by transfer (metis prob_space.finite_measure finite_measure.countable_support) lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV" @@ -193,6 +195,10 @@ lemma emeasure_measure_pmf_finite: "finite S \ emeasure (measure_pmf M) S = (\s\S. pmf M s)" by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single) +lemma measure_measure_pmf_finite: "finite S \ measure (measure_pmf M) S = setsum (pmf M) S" +using emeasure_measure_pmf_finite[of S M] +by(simp add: measure_pmf.emeasure_eq_measure) + lemma nn_integral_measure_pmf_support: fixes f :: "'a \ ereal" assumes f: "finite A" and nn: "\x. x \ A \ 0 \ f x" "\x. x \ set_pmf M \ x \ A \ f x = 0" @@ -234,7 +240,7 @@ then have "integrable (count_space X) (pmf M) = integrable (count_space (M \ X)) (pmf M)" by (simp add: integrable_iff_bounded pmf_nonneg) then show ?thesis - by (simp add: pmf.rep_eq measure_pmf.integrable_measure countable_set_pmf disjoint_family_on_def) + by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def) qed lemma integral_pmf: "(\x. pmf M x \count_space X) = measure M X" @@ -266,6 +272,11 @@ using measure_pmf.emeasure_space_1 by simp qed +lemma in_null_sets_measure_pmfI: + "A \ set_pmf p = {} \ A \ null_sets (measure_pmf p)" +using emeasure_eq_0_AE[where ?P="\x. x \ A" and M="measure_pmf p"] +by(auto simp add: null_sets_def AE_measure_pmf_iff) + lemma map_pmf_id[simp]: "map_pmf id = id" by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI) @@ -287,6 +298,16 @@ lemma nn_integral_map_pmf[simp]: "(\\<^sup>+x. f x \map_pmf g M) = (\\<^sup>+x. f (g x) \M)" unfolding map_pmf.rep_eq by (intro nn_integral_distr) auto +lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\\<^sup>+ y. indicator (f -` {x}) y \measure_pmf p)" +proof(transfer fixing: f x) + fix p :: "'b measure" + presume "prob_space p" + then interpret prob_space p . + presume "sets p = UNIV" + then show "ereal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))" + by(simp add: measure_distr measurable_def emeasure_eq_measure) +qed simp_all + lemma pmf_set_map: fixes f :: "'a \ 'b" shows "set_pmf \ map_pmf f = op ` f \ set_pmf" @@ -317,6 +338,19 @@ lemma set_map_pmf: "set_pmf (map_pmf f M) = f`set_pmf M" using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff) +lemma nn_integral_pmf: "(\\<^sup>+ x. pmf p x \count_space A) = emeasure (measure_pmf p) A" +proof - + have "(\\<^sup>+ x. pmf p x \count_space A) = (\\<^sup>+ x. pmf p x \count_space (A \ set_pmf p))" + by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong) + also have "\ = emeasure (measure_pmf p) (\x\A \ set_pmf p. {x})" + by(subst emeasure_UN_countable)(auto simp add: emeasure_pmf_single disjoint_family_on_def) + also have "\ = emeasure (measure_pmf p) ((\x\A \ set_pmf p. {x}) \ {x. x \ A \ x \ set_pmf p})" + by(rule emeasure_Un_null_set[symmetric])(auto intro: in_null_sets_measure_pmfI) + also have "\ = emeasure (measure_pmf p) A" + by(auto intro: arg_cong2[where f=emeasure]) + finally show ?thesis . +qed + subsection {* PMFs as function *} context @@ -667,7 +701,7 @@ also have "\ = (\ y. \ x. pmf (C x y) i \restrict_space A A \restrict_space B B)" by (rule AB.Fubini_integral[symmetric]) (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2 - simp: pmf_nonneg pmf_le_1 countable_set_pmf measurable_restrict_space1) + simp: pmf_nonneg pmf_le_1 measurable_restrict_space1) also have "\ = (\ y. \ x. pmf (C x y) i \restrict_space A A \B)" by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2 countable_set_pmf borel_measurable_count_space) @@ -783,18 +817,19 @@ done qed - -(* +inductive rel_pmf :: "('a \ 'b \ bool) \ 'a pmf \ 'b pmf \ bool" +for R p q +where + "\ \x y. (x, y) \ set_pmf pq \ R x y; + map_pmf fst pq = p; map_pmf snd pq = q \ + \ rel_pmf R p q" -definition - "rel_pmf P d1 d2 \ (\p3. (\(x, y) \ set_pmf p3. P x y) \ map_pmf fst p3 = d1 \ map_pmf snd p3 = d2)" - -bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: pmf_rel +bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf proof - show "map_pmf id = id" by (rule map_pmf_id) show "\f g. map_pmf (f \ g) = map_pmf f \ map_pmf g" by (rule map_pmf_compose) show "\f g::'a \ 'b. \p. (\x. x \ set_pmf p \ f x = g x) \ map_pmf f p = map_pmf g p" - by (intro map_pmg_cong refl) + by (intro map_pmf_cong refl) show "\f::'a \ 'b. set_pmf \ map_pmf f = op ` f \ set_pmf" by (rule pmf_set_map) @@ -807,46 +842,595 @@ by (metis Field_natLeq card_of_least natLeq_Well_order) finally show "(card_of (set_pmf p), natLeq) \ ordLeq" . } - show "\R. pmf_rel R = - (BNF_Util.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf fst))\\ OO - BNF_Util.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf snd)" - by (auto simp add: fun_eq_iff pmf_rel_def BNF_Util.Grp_def OO_def) + show "\R. rel_pmf R = + (BNF_Def.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf fst))\\ OO + BNF_Def.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf snd)" + by (auto simp add: fun_eq_iff BNF_Def.Grp_def OO_def rel_pmf.simps) + + { fix p :: "'a pmf" and f :: "'a \ 'b" and g x + assume p: "\z. z \ set_pmf p \ f z = g z" + and x: "x \ set_pmf p" + thus "f x = g x" by simp } + + fix R :: "'a => 'b \ bool" and S :: "'b \ 'c \ bool" + { fix p q r + assume pq: "rel_pmf R p q" + and qr:"rel_pmf S q r" + from pq obtain pq where pq: "\x y. (x, y) \ set_pmf pq \ R x y" + and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto + from qr obtain qr where qr: "\y z. (y, z) \ set_pmf qr \ S y z" + and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto + + have support_subset: "set_pmf pq O set_pmf qr \ set_pmf p \ set_pmf r" + by(auto simp add: p r set_map_pmf intro: rev_image_eqI) + + let ?A = "\y. {x. (x, y) \ set_pmf pq}" + and ?B = "\y. {z. (y, z) \ set_pmf qr}" + + + def ppp \ "\A. \f :: 'a \ real. \n. if n \ to_nat_on A ` A then f (from_nat_into A n) else 0" + have [simp]: "\A f n. (\x. x \ A \ 0 \ f x) \ 0 \ ppp A f n" + "\A f n x. \ x \ A; countable A \ \ ppp A f (to_nat_on A x) = f x" + "\A f n. n \ to_nat_on A ` A \ ppp A f n = 0" + by(auto simp add: ppp_def intro: from_nat_into) + def rrr \ "\A. \f :: 'c \ real. \n. if n \ to_nat_on A ` A then f (from_nat_into A n) else 0" + have [simp]: "\A f n. (\x. x \ A \ 0 \ f x) \ 0 \ rrr A f n" + "\A f n x. \ x \ A; countable A \ \ rrr A f (to_nat_on A x) = f x" + "\A f n. n \ to_nat_on A ` A \ rrr A f n = 0" + by(auto simp add: rrr_def intro: from_nat_into) + + def pp \ "\y. ppp (?A y) (\x. pmf pq (x, y))" + and rr \ "\y. rrr (?B y) (\z. pmf qr (y, z))" + + have pos_p [simp]: "\y n. 0 \ pp y n" + and pos_r [simp]: "\y n. 0 \ rr y n" + by(simp_all add: pmf_nonneg pp_def rr_def) + { fix y n + have "pp y n \ 0 \ pp y n = 0" "\ 0 < pp y n \ pp y n = 0" + and "min (pp y n) 0 = 0" "min 0 (pp y n) = 0" + using pos_p[of y n] by(auto simp del: pos_p) } + note pp_convs [simp] = this + { fix y n + have "rr y n \ 0 \ rr y n = 0" "\ 0 < rr y n \ rr y n = 0" + and "min (rr y n) 0 = 0" "min 0 (rr y n) = 0" + using pos_r[of y n] by(auto simp del: pos_r) } + note r_convs [simp] = this + + have "\y. ?A y \ set_pmf p" by(auto simp add: p set_map_pmf intro: rev_image_eqI) + then have [simp]: "\y. countable (?A y)" by(rule countable_subset) simp + + have "\y. ?B y \ set_pmf r" by(auto simp add: r set_map_pmf intro: rev_image_eqI) + then have [simp]: "\y. countable (?B y)" by(rule countable_subset) simp + + let ?P = "\y. to_nat_on (?A y)" + and ?R = "\y. to_nat_on (?B y)" + + have eq: "\y. (\\<^sup>+ x. pp y x \count_space UNIV) = \\<^sup>+ z. rr y z \count_space UNIV" + proof - + fix y + have "(\\<^sup>+ x. pp y x \count_space UNIV) = (\\<^sup>+ x. pp y x \count_space (?P y ` ?A y))" + by(auto simp add: pp_def nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) + also have "\ = (\\<^sup>+ x. pp y (?P y x) \count_space (?A y))" + by(intro nn_integral_bij_count_space[symmetric] inj_on_imp_bij_betw inj_on_to_nat_on) simp + also have "\ = (\\<^sup>+ x. pmf pq (x, y) \count_space (?A y))" + by(rule nn_integral_cong)(simp add: pp_def) + also have "\ = \\<^sup>+ x. emeasure (measure_pmf pq) {(x, y)} \count_space (?A y)" + by(simp add: emeasure_pmf_single) + also have "\ = emeasure (measure_pmf pq) (\x\?A y. {(x, y)})" + by(subst emeasure_UN_countable)(simp_all add: disjoint_family_on_def) + also have "\ = emeasure (measure_pmf pq) ((\x\?A y. {(x, y)}) \ {(x, y'). x \ ?A y \ y' = y})" + by(rule emeasure_Un_null_set[symmetric])+ + (auto simp add: q set_map_pmf split_beta intro!: in_null_sets_measure_pmfI intro: rev_image_eqI) + also have "\ = emeasure (measure_pmf pq) (snd -` {y})" + by(rule arg_cong2[where f=emeasure])+auto + also have "\ = pmf q y" by(simp add: q ereal_pmf_map) + also have "\ = emeasure (measure_pmf qr) (fst -` {y})" + by(simp add: q' ereal_pmf_map) + also have "\ = emeasure (measure_pmf qr) ((\z\?B y. {(y, z)}) \ {(y', z). z \ ?B y \ y' = y})" + by(rule arg_cong2[where f=emeasure])+auto + also have "\ = emeasure (measure_pmf qr) (\z\?B y. {(y, z)})" + by(rule emeasure_Un_null_set) + (auto simp add: q' set_map_pmf split_beta intro!: in_null_sets_measure_pmfI intro: rev_image_eqI) + also have "\ = \\<^sup>+ z. emeasure (measure_pmf qr) {(y, z)} \count_space (?B y)" + by(subst emeasure_UN_countable)(simp_all add: disjoint_family_on_def) + also have "\ = (\\<^sup>+ z. pmf qr (y, z) \count_space (?B y))" + by(simp add: emeasure_pmf_single) + also have "\ = (\\<^sup>+ z. rr y (?R y z) \count_space (?B y))" + by(rule nn_integral_cong)(simp add: rr_def) + also have "\ = (\\<^sup>+ z. rr y z \count_space (?R y ` ?B y))" + by(intro nn_integral_bij_count_space inj_on_imp_bij_betw inj_on_to_nat_on) simp + also have "\ = \\<^sup>+ z. rr y z \count_space UNIV" + by(auto simp add: rr_def nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) + finally show "?thesis y" . + qed - { let ?f = "map_pmf fst" and ?s = "map_pmf snd" - fix R :: "'a \ 'b \ bool" and A assume "\x y. (x, y) \ set_pmf A \ R x y" - fix S :: "'b \ 'c \ bool" and B assume "\y z. (y, z) \ set_pmf B \ S y z" - assume "?f B = ?s A" - have "\C. (\(x, z)\set_pmf C. \y. R x y \ S y z) \ ?f C = ?f A \ ?s C = ?s B" - sorry } -oops - then show "\R::'a \ 'b \ bool. \S::'b \ 'c \ bool. pmf_rel R OO pmf_rel S \ pmf_rel (R OO S)" - by (auto simp add: subset_eq pmf_rel_def fun_eq_iff OO_def Ball_def) -qed (fact natLeq_card_order natLeq_cinfinite)+ + def assign_aux \ "\y remainder start weight z. + if z < start then 0 + else if z = start then min weight remainder + else if remainder + setsum (rr y) {Suc start ..y remainder start weight z. assign_aux y remainder start weight z = + (if z < start then 0 + else if z = start then min weight remainder + else if remainder + setsum (rr y) {Suc start .. weight" + let ?assign_aux = "assign_aux y remainder start weight" + { fix z + have "setsum ?assign_aux {.. start then 0 else if remainder + setsum (rr y) {Suc start.. remainder" + have [simp]: "\z. 0 \ ?assign_aux z" + by(simp add: assign_aux_def weight_nonneg remainder_nonneg) + moreover have "\z. \ rr y z = 0; remainder \ rr y start \ \ ?assign_aux z = 0" + using remainder_nonneg weight_nonneg + by(auto simp add: assign_aux_def min_def) + moreover have "(\\<^sup>+ z. ?assign_aux z \count_space UNIV) = + min weight (\\<^sup>+ z. (if z < start then 0 else if z = start then remainder else rr y z) \count_space UNIV)" + (is "?lhs = ?rhs" is "_ = min _ (\\<^sup>+ y. ?f y \_)") + proof - + have "?lhs = (SUP n. \z = (SUP n. min weight (\zz start then remainder else 0) + setsum ?f {Suc start.. = min weight (setsum ?f {start.. = min weight (setsum ?f {..z" . + qed + also have "\ = min weight (SUP n. setsum ?f {.. = ?rhs" + by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP remainder_nonneg) + finally show ?thesis . + qed + moreover note calculation } + moreover note calculation } + note setsum_start_assign_aux = this(1) + and assign_aux_nonneg [simp] = this(2) + and assign_aux_eq_0_outside = this(3) + and nn_integral_assign_aux = this(4) + { fix y and remainder :: real and start target + have "setsum (rr y) {Suc start.. 0" by(simp add: setsum_nonneg) + moreover assume "0 \ remainder" + ultimately have "assign_aux y remainder start 0 target = 0" + by(auto simp add: assign_aux_def min_def) } + note assign_aux_weight_0 [simp] = this + + def find_start \ "\y weight. if \n. weight \ setsum (rr y) {..n} then Some (LEAST n. weight \ setsum (rr y) {..n}) else None" + have find_start_eq_Some_above: + "\y weight n. find_start y weight = Some n \ weight \ setsum (rr y) {..n}" + by(drule sym)(auto simp add: find_start_def split: split_if_asm intro: LeastI) + { fix y weight n + assume find_start: "find_start y weight = Some n" + and weight: "0 \ weight" + have "setsum (rr y) {..n} \ rr y n + weight" + proof(rule ccontr) + assume "\ ?thesis" + hence "rr y n + weight < setsum (rr y) {..n}" by simp + moreover with weight obtain n' where "n = Suc n'" by(cases n) auto + ultimately have "weight \ setsum (rr y) {..n'}" by simp + hence "(LEAST n. weight \ setsum (rr y) {..n}) \ n'" by(rule Least_le) + moreover from find_start have "n = (LEAST n. weight \ setsum (rr y) {..n})" + by(auto simp add: find_start_def split: split_if_asm) + ultimately show False using \n = Suc n'\ by auto + qed } + note find_start_eq_Some_least = this + have find_start_0 [simp]: "\y. find_start y 0 = Some 0" + by(auto simp add: find_start_def intro!: exI[where x=0]) + { fix y and weight :: real + assume "weight < \\<^sup>+ z. rr y z \count_space UNIV" + also have "(\\<^sup>+ z. rr y z \count_space UNIV) = (SUP n. \zz dom (find_start y)" + by(auto simp add: find_start_def)(meson atMost_iff finite_atMost lessThan_iff less_imp_le order_trans pos_r setsum_mono3 subsetI) } + note in_dom_find_startI = this + { fix y and w w' :: real and m + let ?m' = "LEAST m. w' \ setsum (rr y) {..m}" + assume "w' \ w" + also assume "find_start y w = Some m" + hence "w \ setsum (rr y) {..m}" by(rule find_start_eq_Some_above) + finally have "find_start y w' = Some ?m'" by(auto simp add: find_start_def) + moreover from \w' \ setsum (rr y) {..m}\ have "?m' \ m" by(rule Least_le) + ultimately have "\m'. find_start y w' = Some m' \ m' \ m" by blast } + note find_start_mono = this[rotated] + + def assign \ "\y x z. let used = setsum (pp y) {.. 0 + | Some start \ assign_aux y (setsum (rr y) {..start} - used) start (pp y x) z" + hence assign_alt_def: "\y x z. assign y x z = + (let used = setsum (pp y) {.. 0 + | Some start \ assign_aux y (setsum (rr y) {..start} - used) start (pp y x) z)" + by simp + have assign_nonneg [simp]: "\y x z. 0 \ assign y x z" + by(simp add: assign_def diff_le_iff find_start_eq_Some_above split: option.split) + have assign_eq_0_outside: "\y x z. \ pp y x = 0 \ rr y z = 0 \ \ assign y x z = 0" + by(auto simp add: assign_def assign_aux_eq_0_outside diff_le_iff find_start_eq_Some_above find_start_eq_Some_least setsum_nonneg split: option.split) + + { fix y x z + have "(\n rr y z + | Some m \ if z < m then rr y z + else min (rr y z) (max 0 (setsum (pp y) {.. m" + by(auto dest: find_start_mono[where w'2="setsum (pp y) {.. setsum (rr y) {.. \ setsum (pp y) {.. max 0 (setsum (pp y) {.. z" + have "setsum (pp y) {.. setsum (rr y) {..m}" + using Some by(rule find_start_eq_Some_above) + also have "\ \ setsum (rr y) {..m \ z\ by(intro setsum_mono3) auto + finally have "max 0 (setsum (pp y) {.. rr y z" by simp + moreover have "z \ m \ setsum (rr y) {..m} + setsum (rr y) {Suc m..m \ z\ + by(subst ivl_disj_un(8)[where l="Suc m", symmetric]) + (simp_all add: setsum_Un ivl_disj_un(2)[symmetric] setsum.neutral) + moreover note calculation + } moreover { + assume "m < z" + have "setsum (pp y) {.. setsum (rr y) {..m}" + using Some by(rule find_start_eq_Some_above) + also have "\ \ setsum (rr y) {..m < z\ by(intro setsum_mono3) auto + finally have "max 0 (setsum (pp y) {.. setsum (rr y) {.. setsum (rr y) {.. real" - def IJz \ "rec_nat ((0, 0), \_. 0) (\n ((I, J), z). - let a = x I - (\ji b then I + 1 else I, if b \ a then J + 1 else J), z((I, J) := min a b)))" - def I == "fst \ fst \ IJz" def J == "snd \ fst \ IJz" def z == "snd \ IJz" - let ?a = "\n. x (I n) - (\jp. z 0 p = 0" "I 0 = 0" "J 0 = 0" - by (simp_all add: I_def J_def z_def IJz_def) - have z_Suc[simp]: "\n. z (Suc n) = (z n)((I n, J n) := min (?a n) (?b n))" - by (simp add: z_def I_def J_def IJz_def Let_def split_beta) - have I_Suc[simp]: "\n. I (Suc n) = (if ?a n \ ?b n then I n + 1 else I n)" - by (simp add: z_def I_def J_def IJz_def Let_def split_beta) - have J_Suc[simp]: "\n. J (Suc n) = (if ?b n \ ?a n then J n + 1 else J n)" - by (simp add: z_def I_def J_def IJz_def Let_def split_beta) - - { fix N have "\p. z N p \ 0 \ \nj. z n (i, j)) = x i" - oops -*) + have nn_integral_assign1: "\y z. (\\<^sup>+ x. assign y x z \count_space UNIV) = rr y z" + proof - + fix y z + have "(\\<^sup>+ x. assign y x z \count_space UNIV) = (SUP n. ereal (\x = rr y z" + proof(rule antisym) + show "(SUP n. ereal (\x rr y z" + proof(rule SUP_least) + fix n + show "ereal (\x rr y z" + using setsum_assign[of y z "n - 1"] + by(cases n)(simp_all split: option.split) + qed + show "rr y z \ (SUP n. ereal (\x\<^sup>+ x. pp y x \count_space UNIV") + case True + then obtain n where "setsum (rr y) {..z} < setsum (pp y) {..k. k < z \ setsum (rr y) {..k} \ setsum (rr y) {.. (\x \ (SUP n. ereal (\x\<^sup>+ z. rr y z \count_space {..z}" + by(simp add: nn_integral_count_space_finite max_def) + also have "\ \ \\<^sup>+ z. rr y z \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) + also have "\ = \\<^sup>+ x. pp y x \count_space UNIV" by(simp add: eq) + finally have *: "setsum (rr y) {..z} = \" using False by simp + also have "\ = (SUP n. ereal (\x \ (SUP n. ereal (\x\<^sup>+ x. pp y x \count_space {.. \ \\<^sup>+ x. pp y x \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) + also have "\ = setsum (rr y) {..z}" using * by simp + finally obtain k where k: "find_start y (setsum (pp y) {..ereal (setsum (pp y) {.. setsum (rr y) {..z}\ + have "k \ z" by(auto simp add: find_start_def split: split_if_asm intro: Least_le) + then have "setsum (pp y) {.. ereal (\xereal (setsum (pp y) {.. setsum (rr y) {..z}\ + by(subst setsum_assign)(auto simp add: field_simps max_def k ivl_disj_un(2)[symmetric], metis le_add_same_cancel2 max.bounded_iff max_def pos_p) + also have "\ \ (SUP n. ereal (\xx \ + setsum (rr y) {..\<^sup>+ z. assign y x z \count_space UNIV) = pp y x" + proof(cases "setsum (pp y) {..\<^sup>+ x. pp y x \count_space UNIV") + case False + let ?used = "setsum (pp y) {..\<^sup>+ x. pp y x \count_space {.. \ \\<^sup>+ x. pp y x \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_mono) + finally have "?used < \" using False by auto + also note eq finally have "?used \ dom (find_start y)" by(rule in_dom_find_startI) + then obtain k where k: "find_start y ?used = Some k" by auto + let ?f = "\z. if z < k then 0 else if z = k then setsum (rr y) {..k} - ?used else rr y z" + let ?g = "\x'. if x' < x then 0 else pp y x'" + have "pp y x = ?g x" by simp + also have "?g x \ \\<^sup>+ x'. ?g x' \count_space UNIV" by(rule nn_integral_ge_point) simp + also { + have "?used = \\<^sup>+ x. pp y x \count_space {.. = \\<^sup>+ x'. (if x' < x then pp y x' else 0) \count_space UNIV" + by(simp add: nn_integral_count_space_indicator indicator_def if_distrib zero_ereal_def cong: if_cong) + also have "(\\<^sup>+ x'. ?g x' \count_space UNIV) + \ = \\<^sup>+ x. pp y x \count_space UNIV" + by(subst nn_integral_add[symmetric])(auto intro: nn_integral_cong) + also note calculation } + ultimately have "ereal (pp y x) + ?used \ \\<^sup>+ x. pp y x \count_space UNIV" + by (metis (no_types, lifting) ereal_add_mono order_refl) + also note eq + also have "(\\<^sup>+ z. rr y z \count_space UNIV) = (\\<^sup>+ z. ?f z \count_space UNIV) + (\\<^sup>+ z. (if z < k then rr y z else if z = k then ?used - setsum (rr y) {..count_space UNIV)" + using k by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong simp add: ivl_disj_un(2)[symmetric] setsum_nonneg dest: find_start_eq_Some_least find_start_eq_Some_above) + also have "(\\<^sup>+ z. (if z < k then rr y z else if z = k then ?used - setsum (rr y) {..count_space UNIV) = + (\\<^sup>+ z. (if z < k then rr y z else if z = k then ?used - setsum (rr y) {..count_space {..k})" + by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_cong) + also have "\ = ?used" + using k by(auto simp add: nn_integral_count_space_finite max_def ivl_disj_un(2)[symmetric] diff_le_iff setsum_nonneg dest: find_start_eq_Some_least) + finally have "pp y x \ (\\<^sup>+ z. ?f z \count_space UNIV)" + by(cases "\\<^sup>+ z. ?f z \count_space UNIV") simp_all + then show ?thesis using k + by(simp add: assign_def nn_integral_assign_aux diff_le_iff find_start_eq_Some_above min_def) + next + case True + have "setsum (pp y) {..x} = \\<^sup>+ x. pp y x \count_space {..x}" + by(simp add: nn_integral_count_space_finite max_def) + also have "\ \ \\<^sup>+ x. pp y x \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) + also have "\ = setsum (pp y) {.. "\y x z. ereal (?f y x z)" + + have pos: "\y x z. 0 \ f y x z" by(simp add: f_def) + { fix y x z + have "f y x z \ 0 \ f y x z = 0" using pos[of y x z] by simp } + note f [simp] = this + have support: + "\x y z. (x, y) \ set_pmf pq \ f y x z = 0" + "\x y z. (y, z) \ set_pmf qr \ f y x z = 0" + by(auto simp add: f_def) + + from pos support have support': + "\x z. x \ set_pmf p \ (\\<^sup>+ y. f y x z \count_space UNIV) = 0" + "\x z. z \ set_pmf r \ (\\<^sup>+ y. f y x z \count_space UNIV) = 0" + and support'': + "\x y z. x \ set_pmf p \ f y x z = 0" + "\x y z. y \ set_pmf q \ f y x z = 0" + "\x y z. z \ set_pmf r \ f y x z = 0" + by(auto simp add: nn_integral_0_iff_AE AE_count_space p q r set_map_pmf image_iff)(metis fst_conv snd_conv)+ + + have f_x: "\y z. (\\<^sup>+ x. f y x z \count_space (set_pmf p)) = pmf qr (y, z)" + proof(case_tac "z \ ?B y") + fix y z + assume z: "z \ ?B y" + have "(\\<^sup>+ x. f y x z \count_space (set_pmf p)) = (\\<^sup>+ x. ?f y x z \count_space (?A y))" + using support''(1)[of _ y z] + by(fastforce simp add: f_def nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) + also have "\ = \\<^sup>+ x. assign y (?P y x) (?R y z) \count_space (?A y)" + using z by(intro nn_integral_cong) simp + also have "\ = \\<^sup>+ x. assign y x (?R y z) \count_space (?P y ` ?A y)" + by(intro nn_integral_bij_count_space inj_on_imp_bij_betw inj_on_to_nat_on) simp + also have "\ = \\<^sup>+ x. assign y x (?R y z) \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def assign_eq_0_outside pp_def intro!: nn_integral_cong) + also have "\ = rr y (?R y z)" by(rule nn_integral_assign1) + also have "\ = pmf qr (y, z)" using z by(simp add: rr_def) + finally show "?thesis y z" . + qed(auto simp add: f_def zero_ereal_def[symmetric] set_pmf_iff) + + have f_z: "\x y. (\\<^sup>+ z. f y x z \count_space (set_pmf r)) = pmf pq (x, y)" + proof(case_tac "x \ ?A y") + fix x y + assume x: "x \ ?A y" + have "(\\<^sup>+ z. f y x z \count_space (set_pmf r)) = (\\<^sup>+ z. ?f y x z \count_space (?B y))" + using support''(3)[of _ y x] + by(fastforce simp add: f_def nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) + also have "\ = \\<^sup>+ z. assign y (?P y x) (?R y z) \count_space (?B y)" + using x by(intro nn_integral_cong) simp + also have "\ = \\<^sup>+ z. assign y (?P y x) z \count_space (?R y ` ?B y)" + by(intro nn_integral_bij_count_space inj_on_imp_bij_betw inj_on_to_nat_on) simp + also have "\ = \\<^sup>+ z. assign y (?P y x) z \count_space UNIV" + by(auto simp add: nn_integral_count_space_indicator indicator_def assign_eq_0_outside rr_def intro!: nn_integral_cong) + also have "\ = pp y (?P y x)" by(rule nn_integral_assign2) + also have "\ = pmf pq (x, y)" using x by(simp add: pp_def) + finally show "?thesis x y" . + qed(auto simp add: f_def zero_ereal_def[symmetric] set_pmf_iff) + + let ?pr = "\(x, z). \\<^sup>+ y. f y x z \count_space UNIV" + + have pr_pos: "\xz. 0 \ ?pr xz" + by(auto simp add: nn_integral_nonneg) + + have pr': "?pr = (\(x, z). \\<^sup>+ y. f y x z \count_space (set_pmf q))" + by(auto simp add: fun_eq_iff nn_integral_count_space_indicator indicator_def support'' intro: nn_integral_cong) + + have "(\\<^sup>+ xz. ?pr xz \count_space UNIV) = (\\<^sup>+ xz. ?pr xz * indicator (set_pmf p \ set_pmf r) xz \count_space UNIV)" + by(rule nn_integral_cong)(auto simp add: indicator_def support' intro: ccontr) + also have "\ = (\\<^sup>+ xz. ?pr xz \count_space (set_pmf p \ set_pmf r))" + by(simp add: nn_integral_count_space_indicator) + also have "\ = (\\<^sup>+ xz. ?pr xz \(count_space (set_pmf p) \\<^sub>M count_space (set_pmf r)))" + by(simp add: pair_measure_countable) + also have "\ = (\\<^sup>+ (x, z). \\<^sup>+ y. f y x z \count_space (set_pmf q) \(count_space (set_pmf p) \\<^sub>M count_space (set_pmf r)))" + by(simp add: pr') + also have "\ = (\\<^sup>+ x. \\<^sup>+ z. \\<^sup>+ y. f y x z \count_space (set_pmf q) \count_space (set_pmf r) \count_space (set_pmf p))" + by(subst sigma_finite_measure.nn_integral_fst[symmetric, OF sigma_finite_measure_count_space_countable])(simp_all add: pair_measure_countable) + also have "\ = (\\<^sup>+ x. \\<^sup>+ y. \\<^sup>+ z. f y x z \count_space (set_pmf r) \count_space (set_pmf q) \count_space (set_pmf p))" + by(subst (2) pair_sigma_finite.Fubini')(simp_all add: pair_sigma_finite.intro sigma_finite_measure_count_space_countable pair_measure_countable) + also have "\ = (\\<^sup>+ x. \\<^sup>+ y. pmf pq (x, y) \count_space (set_pmf q) \count_space (set_pmf p))" + by(simp add: f_z) + also have "\ = (\\<^sup>+ y. \\<^sup>+ x. pmf pq (x, y) \count_space (set_pmf p) \count_space (set_pmf q))" + by(subst pair_sigma_finite.Fubini')(simp_all add: pair_sigma_finite.intro sigma_finite_measure_count_space_countable pair_measure_countable) + also have "\ = (\\<^sup>+ y. \\<^sup>+ x. emeasure (measure_pmf pq) {(x, y)} \count_space (set_pmf p) \count_space (set_pmf q))" + by(simp add: emeasure_pmf_single) + also have "\ = (\\<^sup>+ y. emeasure (measure_pmf pq) (\x\set_pmf p. {(x, y)}) \count_space (set_pmf q))" + by(subst emeasure_UN_countable)(simp_all add: disjoint_family_on_def) + also have "\ = (\\<^sup>+ y. emeasure (measure_pmf pq) ((\x\set_pmf p. {(x, y)}) \ {(x, y'). x \ set_pmf p \ y' = y}) \count_space (set_pmf q))" + by(rule nn_integral_cong emeasure_Un_null_set[symmetric])+ + (auto simp add: p set_map_pmf split_beta intro!: in_null_sets_measure_pmfI intro: rev_image_eqI) + also have "\ = (\\<^sup>+ y. emeasure (measure_pmf pq) (snd -` {y}) \count_space (set_pmf q))" + by(rule nn_integral_cong arg_cong2[where f=emeasure])+auto + also have "\ = (\\<^sup>+ y. pmf q y \count_space (set_pmf q))" + by(simp add: ereal_pmf_map q) + also have "\ = (\\<^sup>+ y. pmf q y \count_space UNIV)" + by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong) + also have "\ = 1" + by(subst nn_integral_pmf)(simp add: measure_pmf.emeasure_eq_1_AE) + finally have pr_prob: "(\\<^sup>+ xz. ?pr xz \count_space UNIV) = 1" . + + have pr_bounded: "\xz. ?pr xz \ \" + proof - + fix xz + have "?pr xz \ \\<^sup>+ xz. ?pr xz \count_space UNIV" + by(rule nn_integral_ge_point) simp + also have "\ = 1" by(fact pr_prob) + finally show "?thesis xz" by auto + qed + + def pr \ "embed_pmf (real \ ?pr)" + have pmf_pr: "\xz. pmf pr xz = real (?pr xz)" using pr_pos pr_prob + unfolding pr_def by(subst pmf_embed_pmf)(auto simp add: real_of_ereal_pos ereal_real pr_bounded) + + have set_pmf_pr_subset: "set_pmf pr \ set_pmf pq O set_pmf qr" + proof + fix xz :: "'a \ 'c" + obtain x z where xz: "xz = (x, z)" by(cases xz) + assume "xz \ set_pmf pr" + with xz have "pmf pr (x, z) \ 0" by(simp add: set_pmf_iff) + hence "\y. f y x z \ 0" by(rule contrapos_np)(simp add: pmf_pr) + then obtain y where y: "f y x z \ 0" .. + then have "(x, y) \ set_pmf pq" "(y, z) \ set_pmf qr" + using support by fastforce+ + then show "xz \ set_pmf pq O set_pmf qr" using xz by auto + qed + hence "\x z. (x, z) \ set_pmf pr \ (R OO S) x z" using pq qr by blast + moreover + have "map_pmf fst pr = p" + proof(rule pmf_eqI) + fix x + have "pmf (map_pmf fst pr) x = emeasure (measure_pmf pr) (fst -` {x})" + by(simp add: ereal_pmf_map) + also have "\ = \\<^sup>+ xz. pmf pr xz \count_space (fst -` {x})" + by(simp add: nn_integral_pmf) + also have "\ = \\<^sup>+ xz. ?pr xz \count_space (fst -` {x})" + by(simp add: pmf_pr ereal_real pr_bounded pr_pos) + also have "\ = \\<^sup>+ xz. ?pr xz \count_space {x} \\<^sub>M count_space (set_pmf r)" + by(auto simp add: nn_integral_count_space_indicator indicator_def support' pair_measure_countable intro!: nn_integral_cong) + also have "\ = \\<^sup>+ z. \\<^sup>+ x. ?pr (x, z) \count_space {x} \count_space (set_pmf r)" + by(subst pair_sigma_finite.nn_integral_snd[symmetric])(simp_all add: pair_measure_countable pair_sigma_finite.intro sigma_finite_measure_count_space_countable) + also have "\ = \\<^sup>+ z. ?pr (x, z) \count_space (set_pmf r)" + using pr_pos by(clarsimp simp add: nn_integral_count_space_finite max_def) + also have "\ = \\<^sup>+ z. \\<^sup>+ y. f y x z \count_space (set_pmf q) \count_space (set_pmf r)" + by(simp add: pr') + also have "\ = \\<^sup>+ y. \\<^sup>+ z. f y x z \count_space (set_pmf r) \count_space (set_pmf q)" + by(subst pair_sigma_finite.Fubini')(simp_all add: pair_sigma_finite.intro sigma_finite_measure_count_space_countable pair_measure_countable) + also have "\ = \\<^sup>+ y. pmf pq (x, y) \count_space (set_pmf q)" + by(simp add: f_z) + also have "\ = \\<^sup>+ y. emeasure (measure_pmf pq) {(x, y)} \count_space (set_pmf q)" + by(simp add: emeasure_pmf_single) + also have "\ = emeasure (measure_pmf pq) (\y\set_pmf q. {(x, y)})" + by(subst emeasure_UN_countable)(simp_all add: disjoint_family_on_def) + also have "\ = emeasure (measure_pmf pq) ((\y\set_pmf q. {(x, y)}) \ {(x', y). y \ set_pmf q \ x' = x})" + by(rule emeasure_Un_null_set[symmetric])+ + (auto simp add: q set_map_pmf split_beta intro!: in_null_sets_measure_pmfI intro: rev_image_eqI) + also have "\ = emeasure (measure_pmf pq) (fst -` {x})" + by(rule arg_cong2[where f=emeasure])+auto + also have "\ = pmf p x" by(simp add: ereal_pmf_map p) + finally show "pmf (map_pmf fst pr) x = pmf p x" by simp + qed + moreover + have "map_pmf snd pr = r" + proof(rule pmf_eqI) + fix z + have "pmf (map_pmf snd pr) z = emeasure (measure_pmf pr) (snd -` {z})" + by(simp add: ereal_pmf_map) + also have "\ = \\<^sup>+ xz. pmf pr xz \count_space (snd -` {z})" + by(simp add: nn_integral_pmf) + also have "\ = \\<^sup>+ xz. ?pr xz \count_space (snd -` {z})" + by(simp add: pmf_pr ereal_real pr_bounded pr_pos) + also have "\ = \\<^sup>+ xz. ?pr xz \count_space (set_pmf p) \\<^sub>M count_space {z}" + by(auto simp add: nn_integral_count_space_indicator indicator_def support' pair_measure_countable intro!: nn_integral_cong) + also have "\ = \\<^sup>+ x. \\<^sup>+ z. ?pr (x, z) \count_space {z} \count_space (set_pmf p)" + by(subst sigma_finite_measure.nn_integral_fst[symmetric])(simp_all add: pair_measure_countable sigma_finite_measure_count_space_countable) + also have "\ = \\<^sup>+ x. ?pr (x, z) \count_space (set_pmf p)" + using pr_pos by(clarsimp simp add: nn_integral_count_space_finite max_def) + also have "\ = \\<^sup>+ x. \\<^sup>+ y. f y x z \count_space (set_pmf q) \count_space (set_pmf p)" + by(simp add: pr') + also have "\ = \\<^sup>+ y. \\<^sup>+ x. f y x z \count_space (set_pmf p) \count_space (set_pmf q)" + by(subst pair_sigma_finite.Fubini')(simp_all add: pair_sigma_finite.intro sigma_finite_measure_count_space_countable pair_measure_countable) + also have "\ = \\<^sup>+ y. pmf qr (y, z) \count_space (set_pmf q)" + by(simp add: f_x) + also have "\ = \\<^sup>+ y. emeasure (measure_pmf qr) {(y, z)} \count_space (set_pmf q)" + by(simp add: emeasure_pmf_single) + also have "\ = emeasure (measure_pmf qr) (\y\set_pmf q. {(y, z)})" + by(subst emeasure_UN_countable)(simp_all add: disjoint_family_on_def) + also have "\ = emeasure (measure_pmf qr) ((\y\set_pmf q. {(y, z)}) \ {(y, z'). y \ set_pmf q \ z' = z})" + by(rule emeasure_Un_null_set[symmetric])+ + (auto simp add: q' set_map_pmf split_beta intro!: in_null_sets_measure_pmfI intro: rev_image_eqI) + also have "\ = emeasure (measure_pmf qr) (snd -` {z})" + by(rule arg_cong2[where f=emeasure])+auto + also have "\ = pmf r z" by(simp add: ereal_pmf_map r) + finally show "pmf (map_pmf snd pr) z = pmf r z" by simp + qed + ultimately have "rel_pmf (R OO S) p r" .. } + then show "rel_pmf R OO rel_pmf S \ rel_pmf (R OO S)" + by(auto simp add: le_fun_def) +qed (fact natLeq_card_order natLeq_cinfinite)+ end