# HG changeset patch # User Andreas Lochbihler # Date 1420791411 -3600 # Node ID 922d31f5c3f5836562dc5e17380a9acfddd7464c # Parent 3ef6b0b6232edcde9240bcf7abc7588b07641a2c simplify construction for distribution of rel_pmf over op OO diff -r 3ef6b0b6232e -r 922d31f5c3f5 src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Wed Jan 07 18:09:11 2015 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Fri Jan 09 09:16:51 2015 +0100 @@ -1041,11 +1041,19 @@ 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)" + { fix x y have "x \ A y \ ?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 pp_0: "\y x. pmf q y = 0 \ ?pp y x = 0" + proof(erule contrapos_pp) + fix y x + assume "?pp y x \ 0" + hence "(y, x) \ set_pmf pp" by(simp add: set_pmf_iff) + hence "y \ set_pmf q" by(auto simp add: pp_def q set_map_pmf intro: rev_image_eqI) + thus "pmf q y \ 0" by(simp add: set_pmf_iff) + qed 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) @@ -1056,318 +1064,50 @@ 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)" + { fix y z have "z \ B y \ ?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 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) = 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" . + have rr_0: "\y z. pmf q y = 0 \ ?rr y z = 0" + proof(erule contrapos_pp) + fix y z + assume "?rr y z \ 0" + hence "(y, z) \ set_pmf rr" by(simp add: set_pmf_iff) + hence "y \ set_pmf q" by(auto simp add: rr_def q' set_map_pmf intro: rev_image_eqI) + thus "pmf q y \ 0" by(simp add: set_pmf_iff) 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.. 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 + have nn_integral_pp2: "\y. (\\<^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]) + have nn_integral_rr1: "\y. (\\<^sup>+ x. ?rr y x \count_space UNIV) = pmf q y" + by (simp add: nn_integral_pmf' inj_on_def rr_def q') + (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) + have eq: "\y. (\\<^sup>+ x. ?pp y x \count_space UNIV) = (\\<^sup>+ z. ?rr y z \count_space UNIV)" + by(simp add: nn_integral_pp2 nn_integral_rr1) - 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 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}" - 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 Let_def split: option.split) + def assign \ "\y x z. ?pp y x * ?rr y z / pmf q y" + have assign_nonneg [simp]: "\y x z. 0 \ assign y x z" by(simp add: assign_def) 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) {.. 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) {..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}\ - 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) {..\<^sup>+ x. assign y x z \count_space UNIV) = + (\\<^sup>+ x. ?pp y x \count_space UNIV) * (?rr y z / pmf q y)" + by(simp add: assign_def nn_integral_multc times_ereal.simps(1)[symmetric] divide_real_def mult.assoc del: times_ereal.simps(1)) + also have "\ = ?rr y z" by(simp add: rr_0 nn_integral_pp2) finally show "?thesis y z" . qed - - { fix y x - 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 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 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" - 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. (\\<^sup>+ z. assign y x z \count_space UNIV) = ?pp y x" + proof - + fix x y + have "(\\<^sup>+ z. assign y x z \count_space UNIV) = (\\<^sup>+ z. ?rr y z \count_space UNIV) * (?pp y x / pmf q y)" + by(simp add: assign_def divide_real_def mult.commute[where a="?pp y x"] mult.assoc nn_integral_multc times_ereal.simps(1)[symmetric] del: times_ereal.simps(1)) + also have "\ = ?pp y x" by(simp add: nn_integral_rr1 pp_0) + finally show "?thesis y x" . + qed def a \ "embed_pmf (\(y, x, z). assign y x z)" { fix y x z