# HG changeset patch # User hoelzl # Date 1416933005 -3600 # Node ID 43e07797269b37c08f098698ed7d7aeede6f3cf6 # Parent a05c8305781e3275fa173f6148c5f5782d967321 tuned proof that pmfs are bnfs diff -r a05c8305781e -r 43e07797269b src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 25 17:29:39 2014 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 25 17:30:05 2014 +0100 @@ -297,6 +297,9 @@ lemma map_pmf_id[simp]: "map_pmf id = id" by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI) +lemma map_pmf_ident[simp]: "map_pmf (\x. x) = (\x. x)" + using map_pmf_id unfolding id_def . + lemma map_pmf_compose: "map_pmf (f \ g) = map_pmf f \ map_pmf g" by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) @@ -844,6 +847,24 @@ lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B" by (simp add: pair_pmf_def bind_return_pmf''[symmetric] bind_assoc_pmf bind_return_pmf bind_return_pmf') +lemma nn_integral_pmf': + "inj_on f A \ (\\<^sup>+x. pmf p (f x) \count_space A) = emeasure p (f ` A)" + by (subst nn_integral_bij_count_space[where g=f and B="f`A"]) + (auto simp: bij_betw_def nn_integral_pmf) + +lemma pmf_le_0_iff[simp]: "pmf M p \ 0 \ pmf M p = 0" + using pmf_nonneg[of M p] by simp + +lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0" + using pmf_nonneg[of M p] by simp_all + +lemma pmf_eq_0_set_pmf: "pmf M p = 0 \ p \ set_pmf M" + unfolding set_pmf_iff by simp + +lemma pmf_map_inj: "inj_on f (set_pmf M) \ x \ set_pmf M \ pmf (map_pmf f M) (f x) = pmf M x" + by (auto simp: pmf.rep_eq map_pmf.rep_eq measure_distr AE_measure_pmf_iff inj_onD + intro!: measure_pmf.finite_measure_eq_AE) + inductive rel_pmf :: "('a \ 'b \ bool) \ 'a pmf \ 'b pmf \ bool" for R p q where @@ -864,7 +885,7 @@ { fix p :: "'s pmf" have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \ ordLeq" by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"]) - (auto intro: countable_set_pmf inj_on_to_nat_on) + (auto intro: countable_set_pmf) also have "(card_of (UNIV :: nat set), natLeq) \ ordLeq" by (metis Field_natLeq card_of_least natLeq_Well_order) finally show "(card_of (set_pmf p), natLeq) \ ordLeq" . } @@ -888,118 +909,80 @@ 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) + note pmf_nonneg[intro, simp] - let ?A = "\y. {x. (x, y) \ set_pmf pq}" - and ?B = "\y. {z. (y, z) \ set_pmf qr}" - + def A \ "\y. {x. (x, y) \ set_pmf pq}" + then 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 A: "\x y. (x, y) \ set_pmf pq \ x \ A y" + by (simp add: A_def) - 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))" + let ?P = "\y. to_nat_on (A y)" + def pp \ "map_pmf (\(x, y). (y, ?P y x)) pq" + let ?pp = "\y x. pmf pp (y, x)" + { fix x y have "x \ A y \ pmf pp (y, ?P y x) = pmf pq (x, y)" + unfolding pp_def + by (intro pmf_map_inj[of "\(x, y). (y, ?P y x)" pq "(x, y)", simplified]) + (auto simp: inj_on_def A) } + note pmf_pp = this - 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 + def B \ "\y. {z. (y, z) \ set_pmf qr}" + then 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 + have B: "\y z. (y, z) \ set_pmf qr \ z \ B y" + by (simp add: B_def) - 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 + let ?R = "\y. to_nat_on (B y)" + def rr \ "map_pmf (\(y, z). (y, ?R y z)) qr" + let ?rr = "\y z. pmf rr (y, z)" + { fix y z have "z \ B y \ pmf rr (y, ?R y z) = pmf qr (y, z)" + unfolding rr_def + by (intro pmf_map_inj[of "\(y, z). (y, ?R y z)" qr "(y, z)", simplified]) + (auto simp: inj_on_def B) } + note pmf_rr = this - 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" + 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) + have "(\\<^sup>+ x. ?pp y x \count_space UNIV) = pmf q y" + by (simp add: nn_integral_pmf' inj_on_def pp_def q) + (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) + also have "\ = (\\<^sup>+ x. ?rr y x \count_space UNIV)" + by (simp add: nn_integral_pmf' inj_on_def rr_def q') + (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) finally show "?thesis y" . qed 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.. 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" + 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)" + 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 0" by(simp add: setsum_nonneg) + 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" + 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}" + "\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" + 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 + 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})" + 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 } @@ -1056,41 +1039,41 @@ 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. \z\<^sup>+ z. ?rr y z \count_space UNIV" + also have "(\\<^sup>+ z. ?rr y z \count_space UNIV) = (SUP n. \zzz 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) } + by(auto simp add: find_start_def)(meson atMost_iff finite_atMost lessThan_iff less_imp_le order_trans pmf_nonneg 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}" + 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) + 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) + 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) {.. "\y x z. let used = setsum (?pp y) {.. 0 - | Some start \ assign_aux y (setsum (rr y) {..start} - used) start (pp y x) z" + | 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)" + | 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) + by(simp add: assign_def diff_le_iff find_start_eq_Some_above Let_def 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 Let_def 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) {.. ?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) {.. m" + by(auto dest: find_start_mono[where w'2="setsum (?pp y) {.. setsum (rr y) {.. setsum (?rr y) {.. \ setsum (pp y) {.. \ setsum (?pp y) {.. max 0 (setsum (pp y) {.. max 0 (setsum (?pp y) {.. z" - have "setsum (pp y) {.. setsum (rr y) {..m}" + 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.. \ 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}" + 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) {..m < z\ by(intro setsum_mono3) auto + finally have "max 0 (setsum (?pp y) {.. setsum (?rr y) {.. setsum (rr y) {.. setsum (?rr y) {..y z. (\\<^sup>+ x. assign y x z \count_space UNIV) = rr y z" + 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" + also have "\ = ?rr y z" proof(rule antisym) - show "(SUP n. ereal (\x rr y z" + show "(SUP n. ereal (\x ?rr y z" proof(rule SUP_least) fix n - show "ereal (\x rr y z" + 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") + 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) {..k. k < z \ setsum (?rr y) {..k} \ setsum (?rr y) {.. (\x (\x \ (SUP n. ereal (\x\<^sup>+ z. rr y z \count_space {..z}" + have "setsum (?rr y) {..z} = \\<^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" + 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>+ 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 n. ereal (\x\<^sup>+ x. pp y x \count_space {..\<^sup>+ x. ?pp y x \count_space {.. \ \\<^sup>+ x. pp y x \count_space UNIV" + 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 (rr y) {..z}" using * by simp - finally obtain k where k: "find_start y (setsum (pp y) {.. = 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}\ + with \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) + then have "setsum (?pp y) {.. ereal (\xereal (setsum (?pp y) {.. setsum (?rr y) {..z}\ + apply (subst setsum_assign) + apply (auto simp add: field_simps max_def k ivl_disj_un(2)[symmetric]) + apply (meson add_increasing le_cases pmf_nonneg) + done also have "\ \ (SUP n. ereal (\xx \ + setsum (rr y) {..x \ + 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") + have "(\\<^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 {.. \ \\<^sup>+ x. pp y x \count_space UNIV" + 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) 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 + 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. ?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" + also have "\ = \\<^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 del: 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" + 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)" + 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})" + 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)" + 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}" + 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" + 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) {.. = 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) + def a \ "embed_pmf (\(y, x, z). assign y x z)" { 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 "assign y x z = pmf a (y, x, z)" + unfolding a_def + proof (subst pmf_embed_pmf) + have "(\\<^sup>+ x. ereal ((\(y, x, z). assign y x z) x) \count_space UNIV) = + (\\<^sup>+ x. ereal ((\(y, x, z). assign y x z) x) \(count_space ((\((y, x), z). (y, x, z)) ` (pp \ UNIV))))" + by (force simp add: nn_integral_count_space_indicator pmf_eq_0_set_pmf split: split_indicator + intro!: nn_integral_cong assign_eq_0_outside) + also have "\ = (\\<^sup>+ x. ereal ((\((y, x), z). assign y x z) x) \(count_space (pp \ UNIV)))" + by (subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric]) + (auto simp: inj_on_def intro!: nn_integral_cong) + also have "\ = (\\<^sup>+ y. \\<^sup>+z. ereal ((\((y, x), z). assign y x z) (y, z)) \count_space UNIV \count_space pp)" + by (subst sigma_finite_measure.nn_integral_fst) + (auto simp: pair_measure_countable sigma_finite_measure_count_space_countable) + also have "\ = (\\<^sup>+ z. ?pp (fst z) (snd z) \count_space pp)" + by (subst nn_integral_assign2[symmetric]) (auto intro!: nn_integral_cong) + finally show "(\\<^sup>+ x. ereal ((\(y, x, z). assign y x z) x) \count_space UNIV) = 1" + by (simp add: nn_integral_pmf emeasure_pmf) + qed auto } + note a = this - 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) + def pr \ "map_pmf (\(y, x, z). (from_nat_into (A y) x, from_nat_into (B y) z)) a" - 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" + have "rel_pmf (R OO S) p r" 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" .. } + have pp_eq: "pp = map_pmf (\(y, x, z). (y, x)) a" + proof (rule pmf_eqI) + fix i + show "pmf pp i = pmf (map_pmf (\(y, x, z). (y, x)) a) i" + using nn_integral_assign2[of "fst i" "snd i", symmetric] + by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map + simp del: ereal.inject intro!: arg_cong2[where f=emeasure]) + qed + moreover have pq_eq: "pq = map_pmf (\(y, x). (from_nat_into (A y) x, y)) pp" + by (simp add: pp_def map_pmf_comp split_beta A[symmetric] cong: map_pmf_cong) + ultimately show "map_pmf fst pr = p" + unfolding p pr_def by (simp add: map_pmf_comp split_beta) + + have rr_eq: "rr = map_pmf (\(y, x, z). (y, z)) a" + proof (rule pmf_eqI) + fix i show "pmf rr i = pmf (map_pmf (\(y, x, z). (y, z)) a) i" + using nn_integral_assign1[of "fst i" "snd i", symmetric] + by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map + simp del: ereal.inject intro!: arg_cong2[where f=emeasure]) + qed + moreover have qr_eq: "qr = map_pmf (\(y, z). (y, from_nat_into (B y) z)) rr" + by (simp add: rr_def map_pmf_comp split_beta B[symmetric] cong: map_pmf_cong) + ultimately show "map_pmf snd pr = r" + unfolding r pr_def by (simp add: map_pmf_comp split_beta) + + fix x z assume "(x, z) \ set_pmf pr" + then have "\y. (x, y) \ set_pmf pq \ (y, z) \ set_pmf qr" + by (force simp add: pp_eq pq_eq rr_eq qr_eq set_map_pmf pr_def image_image) + with pq qr show "(R OO S) x z" + by blast + qed } 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)+