| author | blanchet | 
| Wed, 06 Dec 2023 11:08:39 +0100 | |
| changeset 79139 | 359abf434d70 | 
| parent 78517 | 28c1f4f5335f | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1 | (* Author: Andreas Lochbihler, ETH Zurich *) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 3 | section \<open>Discrete subprobability distribution\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 4 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 5 | theory SPMF imports | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 6 | Probability_Mass_Function | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64634diff
changeset | 7 | "HOL-Library.Complete_Partial_Order2" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64634diff
changeset | 8 | "HOL-Library.Rewrite" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 9 | begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 10 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 11 | subsection \<open>Auxiliary material\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 12 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 13 | lemma cSUP_singleton [simp]: "(SUP x\<in>{x}. f x :: _ :: conditionally_complete_lattice) = f x"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 14 | by (metis cSup_singleton image_empty image_insert) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 15 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 16 | subsubsection \<open>More about extended reals\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 17 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 18 | lemma [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 19 | shows ennreal_max_0: "ennreal (max 0 x) = ennreal x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 20 | and ennreal_max_0': "ennreal (max x 0) = ennreal x" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 21 | by(simp_all add: max_def ennreal_eq_0_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 22 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 23 | lemma e2ennreal_0 [simp]: "e2ennreal 0 = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 24 | by(simp add: zero_ennreal_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 25 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 26 | lemma enn2real_bot [simp]: "enn2real \<bottom> = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 27 | by(simp add: bot_ennreal_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 28 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 29 | lemma continuous_at_ennreal[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. ennreal (f x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 30 | unfolding continuous_def by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 31 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 32 | lemma ennreal_Sup: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 33 | assumes *: "(SUP a\<in>A. ennreal a) \<noteq> \<top>" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 34 |   and "A \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 35 | shows "ennreal (Sup A) = (SUP a\<in>A. ennreal a)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 36 | proof (rule continuous_at_Sup_mono) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 37 | obtain r where r: "ennreal r = (SUP a\<in>A. ennreal a)" "r \<ge> 0" | 
| 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 38 | using * by(cases "(SUP a\<in>A. ennreal a)") simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 39 | then show "bdd_above A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 40 | by(auto intro!: SUP_upper bdd_aboveI[of _ r] simp add: ennreal_le_iff[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 41 | qed (auto simp: mono_def continuous_at_imp_continuous_at_within continuous_at_ennreal ennreal_leI assms) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 42 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 43 | lemma ennreal_SUP: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 44 |   "\<lbrakk> (SUP a\<in>A. ennreal (f a)) \<noteq> \<top>; A \<noteq> {} \<rbrakk> \<Longrightarrow> ennreal (SUP a\<in>A. f a) = (SUP a\<in>A. ennreal (f a))"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 45 | using ennreal_Sup[of "f ` A"] by (auto simp: image_comp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 46 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 47 | lemma ennreal_lt_0: "x < 0 \<Longrightarrow> ennreal x = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 48 | by(simp add: ennreal_eq_0_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 49 | |
| 69597 | 50 | subsubsection \<open>More about \<^typ>\<open>'a option\<close>\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 51 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 52 | lemma None_in_map_option_image [simp]: "None \<in> map_option f ` A \<longleftrightarrow> None \<in> A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 53 | by auto | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 54 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 55 | lemma Some_in_map_option_image [simp]: "Some x \<in> map_option f ` A \<longleftrightarrow> (\<exists>y. x = f y \<and> Some y \<in> A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 56 | by (smt (verit, best) imageE imageI map_option_eq_Some) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 57 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 58 | lemma case_option_collapse: "case_option x (\<lambda>_. x) = (\<lambda>_. x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 59 | by(simp add: fun_eq_iff split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 60 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 61 | lemma case_option_id: "case_option None Some = id" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 62 | by(rule ext)(simp split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 63 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 64 | inductive ord_option :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 65 | for ord :: "'a \<Rightarrow> 'b \<Rightarrow> bool" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 66 | where | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 67 | None: "ord_option ord None x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 68 | | Some: "ord x y \<Longrightarrow> ord_option ord (Some x) (Some y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 69 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 70 | inductive_simps ord_option_simps [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 71 | "ord_option ord None x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 72 | "ord_option ord x None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 73 | "ord_option ord (Some x) (Some y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 74 | "ord_option ord (Some x) None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 75 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 76 | inductive_simps ord_option_eq_simps [simp]: | 
| 67399 | 77 | "ord_option (=) None y" | 
| 78 | "ord_option (=) (Some x) y" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 79 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 80 | lemma ord_option_reflI: "(\<And>y. y \<in> set_option x \<Longrightarrow> ord y y) \<Longrightarrow> ord_option ord x x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 81 | by(cases x) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 82 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 83 | lemma reflp_ord_option: "reflp ord \<Longrightarrow> reflp (ord_option ord)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 84 | by(simp add: reflp_def ord_option_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 85 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 86 | lemma ord_option_trans: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 87 | "\<lbrakk> ord_option ord x y; ord_option ord y z; | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 88 | \<And>a b c. \<lbrakk> a \<in> set_option x; b \<in> set_option y; c \<in> set_option z; ord a b; ord b c \<rbrakk> \<Longrightarrow> ord a c \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 89 | \<Longrightarrow> ord_option ord x z" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 90 | by(auto elim!: ord_option.cases) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 91 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 92 | lemma transp_ord_option: "transp ord \<Longrightarrow> transp (ord_option ord)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 93 | unfolding transp_def by(blast intro: ord_option_trans) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 94 | |
| 64634 | 95 | lemma antisymp_ord_option: "antisymp ord \<Longrightarrow> antisymp (ord_option ord)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 96 | by(auto intro!: antisympI elim!: ord_option.cases dest: antisympD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 97 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 98 | lemma ord_option_chainD: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 99 | "Complete_Partial_Order.chain (ord_option ord) Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 100 |   \<Longrightarrow> Complete_Partial_Order.chain ord {x. Some x \<in> Y}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 101 | by(rule chainI)(auto dest: chainD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 102 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 103 | definition lub_option :: "('a set \<Rightarrow> 'b) \<Rightarrow> 'a option set \<Rightarrow> 'b option"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 104 |   where "lub_option lub Y = (if Y \<subseteq> {None} then None else Some (lub {x. Some x \<in> Y}))"
 | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 105 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 106 | lemma map_lub_option: "map_option f (lub_option lub Y) = lub_option (f \<circ> lub) Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 107 | by(simp add: lub_option_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 108 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 109 | lemma lub_option_upper: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 110 | assumes "Complete_Partial_Order.chain (ord_option ord) Y" "x \<in> Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 111 | and lub_upper: "\<And>Y x. \<lbrakk> Complete_Partial_Order.chain ord Y; x \<in> Y \<rbrakk> \<Longrightarrow> ord x (lub Y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 112 | shows "ord_option ord x (lub_option lub Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 113 | using assms(1-2) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 114 | by(cases x)(auto simp: lub_option_def intro: lub_upper[OF ord_option_chainD]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 115 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 116 | lemma lub_option_least: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 117 | assumes Y: "Complete_Partial_Order.chain (ord_option ord) Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 118 | and upper: "\<And>x. x \<in> Y \<Longrightarrow> ord_option ord x y" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 119 | assumes lub_least: "\<And>Y y. \<lbrakk> Complete_Partial_Order.chain ord Y; \<And>x. x \<in> Y \<Longrightarrow> ord x y \<rbrakk> \<Longrightarrow> ord (lub Y) y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 120 | shows "ord_option ord (lub_option lub Y) y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 121 | using Y | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 122 | by(cases y)(auto 4 3 simp add: lub_option_def intro: lub_least[OF ord_option_chainD] dest: upper) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 123 | |
| 67399 | 124 | lemma lub_map_option: "lub_option lub (map_option f ` Y) = lub_option (lub \<circ> (`) f) Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 125 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 126 |   have "\<And>u y. \<lbrakk>Some u \<in> Y; y \<in> Y\<rbrakk> \<Longrightarrow> {f y |y. Some y \<in> Y} = f ` {x. Some x \<in> Y}"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 127 | by blast | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 128 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 129 | by (auto simp: lub_option_def) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 130 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 131 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 132 | lemma ord_option_mono: "\<lbrakk> ord_option A x y; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> ord_option B x y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 133 | by(auto elim: ord_option.cases) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 134 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 135 | lemma ord_option_mono' [mono]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 136 | "(\<And>x y. A x y \<longrightarrow> B x y) \<Longrightarrow> ord_option A x y \<longrightarrow> ord_option B x y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 137 | by(blast intro: ord_option_mono) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 138 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 139 | lemma ord_option_compp: "ord_option (A OO B) = ord_option A OO ord_option B" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 140 | by(auto simp: fun_eq_iff elim!: ord_option.cases intro: ord_option.intros) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 141 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 142 | lemma ord_option_inf: "inf (ord_option A) (ord_option B) = ord_option (inf A B)" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 143 | proof(rule antisym) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 144 | show "?lhs \<le> ?rhs" by(auto elim!: ord_option.cases) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 145 | qed(auto elim: ord_option_mono) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 146 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 147 | lemma ord_option_map2: "ord_option ord x (map_option f y) = ord_option (\<lambda>x y. ord x (f y)) x y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 148 | by(auto elim: ord_option.cases) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 149 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 150 | lemma ord_option_map1: "ord_option ord (map_option f x) y = ord_option (\<lambda>x y. ord (f x) y) x y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 151 | by(auto elim: ord_option.cases) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 152 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 153 | lemma option_ord_Some1_iff: "option_ord (Some x) y \<longleftrightarrow> y = Some x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 154 | by(auto simp: flat_ord_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 155 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 156 | subsubsection \<open>A relator for sets that treats sets like predicates\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 157 | |
| 63343 | 158 | context includes lifting_syntax | 
| 159 | begin | |
| 160 | ||
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 161 | definition rel_pred :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 162 | where "rel_pred R A B = (R ===> (=)) (\<lambda>x. x \<in> A) (\<lambda>y. y \<in> B)" | 
| 67399 | 163 | |
| 164 | lemma rel_predI: "(R ===> (=)) (\<lambda>x. x \<in> A) (\<lambda>y. y \<in> B) \<Longrightarrow> rel_pred R A B" | |
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 165 | by(simp add: rel_pred_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 166 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 167 | lemma rel_predD: "\<lbrakk> rel_pred R A B; R x y \<rbrakk> \<Longrightarrow> x \<in> A \<longleftrightarrow> y \<in> B" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 168 | by(simp add: rel_pred_def rel_fun_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 169 | |
| 67399 | 170 | lemma Collect_parametric: "((A ===> (=)) ===> rel_pred A) Collect Collect" | 
| 63308 | 171 |   \<comment> \<open>Declare this rule as @{attribute transfer_rule} only locally
 | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 172 |       because it blows up the search space for @{method transfer}
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 173 |       (in combination with @{thm [source] Collect_transfer})\<close>
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 174 | by(simp add: rel_funI rel_predI) | 
| 63343 | 175 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 176 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 177 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 178 | subsubsection \<open>Monotonicity rules\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 179 | |
| 67399 | 180 | lemma monotone_gfp_eadd1: "monotone (\<ge>) (\<ge>) (\<lambda>x. x + y :: enat)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 181 | by(auto intro!: monotoneI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 182 | |
| 67399 | 183 | lemma monotone_gfp_eadd2: "monotone (\<ge>) (\<ge>) (\<lambda>y. x + y :: enat)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 184 | by(auto intro!: monotoneI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 185 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 186 | lemma mono2mono_gfp_eadd[THEN gfp.mono2mono2, cont_intro, simp]: | 
| 67399 | 187 | shows monotone_eadd: "monotone (rel_prod (\<ge>) (\<ge>)) (\<ge>) (\<lambda>(x, y). x + y :: enat)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 188 | by(simp add: monotone_gfp_eadd1 monotone_gfp_eadd2) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 189 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 190 | lemma eadd_gfp_partial_function_mono [partial_function_mono]: | 
| 67399 | 191 | "\<lbrakk> monotone (fun_ord (\<ge>)) (\<ge>) f; monotone (fun_ord (\<ge>)) (\<ge>) g \<rbrakk> | 
| 192 | \<Longrightarrow> monotone (fun_ord (\<ge>)) (\<ge>) (\<lambda>x. f x + g x :: enat)" | |
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 193 | by(rule mono2mono_gfp_eadd) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 194 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 195 | lemma mono2mono_ereal[THEN lfp.mono2mono]: | 
| 67399 | 196 | shows monotone_ereal: "monotone (\<le>) (\<le>) ereal" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 197 | by(rule monotoneI) simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 198 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 199 | lemma mono2mono_ennreal[THEN lfp.mono2mono]: | 
| 67399 | 200 | shows monotone_ennreal: "monotone (\<le>) (\<le>) ennreal" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 201 | by(rule monotoneI)(simp add: ennreal_leI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 202 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 203 | subsubsection \<open>Bijections\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 204 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 205 | lemma bi_unique_rel_set_bij_betw: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 206 | assumes unique: "bi_unique R" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 207 | and rel: "rel_set R A B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 208 | shows "\<exists>f. bij_betw f A B \<and> (\<forall>x\<in>A. R x (f x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 209 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 210 | from assms obtain f where f: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" and B: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 211 | by (metis bi_unique_rel_set_lemma image_eqI) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 212 | have "inj_on f A" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 213 | by (metis (no_types, lifting) bi_unique_def f inj_on_def unique) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 214 | moreover have "f ` A = B" using rel | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 215 | by (smt (verit) bi_unique_def bi_unique_rel_set_lemma f image_cong unique) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 216 | ultimately have "bij_betw f A B" unfolding bij_betw_def .. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 217 | thus ?thesis using f by blast | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 218 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 219 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 220 | lemma bij_betw_rel_setD: "bij_betw f A B \<Longrightarrow> rel_set (\<lambda>x y. y = f x) A B" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 221 | by(rule rel_setI)(auto dest: bij_betwE bij_betw_imp_surj_on[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 222 | |
| 63308 | 223 | subsection \<open>Subprobability mass function\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 224 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 225 | type_synonym 'a spmf = "'a option pmf" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 226 | translations (type) "'a spmf" \<leftharpoondown> (type) "'a option pmf" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 227 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 228 | definition measure_spmf :: "'a spmf \<Rightarrow> 'a measure" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 229 | where "measure_spmf p = distr (restrict_space (measure_pmf p) (range Some)) (count_space UNIV) the" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 230 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 231 | abbreviation spmf :: "'a spmf \<Rightarrow> 'a \<Rightarrow> real" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 232 | where "spmf p x \<equiv> pmf p (Some x)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 233 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 234 | lemma space_measure_spmf: "space (measure_spmf p) = UNIV" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 235 | by(simp add: measure_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 236 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 237 | lemma sets_measure_spmf [simp, measurable_cong]: "sets (measure_spmf p) = sets (count_space UNIV)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 238 | by(simp add: measure_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 239 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 240 | lemma measure_spmf_not_bot [simp]: "measure_spmf p \<noteq> \<bottom>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 241 | by (metis empty_not_UNIV space_bot space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 242 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 243 | lemma measurable_the_measure_pmf_Some [measurable, simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 244 | "the \<in> measurable (restrict_space (measure_pmf p) (range Some)) (count_space UNIV)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 245 | by(auto simp: measurable_def sets_restrict_space space_restrict_space integral_restrict_space) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 246 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 247 | lemma measurable_spmf_measure1[simp]: "measurable (measure_spmf M) N = UNIV \<rightarrow> space N" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 248 | by(auto simp: measurable_def space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 249 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 250 | lemma measurable_spmf_measure2[simp]: "measurable N (measure_spmf M) = measurable N (count_space UNIV)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 251 | by(intro measurable_cong_sets) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 252 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 253 | lemma subprob_space_measure_spmf [simp, intro!]: "subprob_space (measure_spmf p)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 254 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 255 | show "emeasure (measure_spmf p) (space (measure_spmf p)) \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 256 | by(simp add: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space measure_pmf.measure_le_1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 257 | qed(simp add: space_measure_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 258 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 259 | interpretation measure_spmf: subprob_space "measure_spmf p" for p | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 260 | by(rule subprob_space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 261 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 262 | lemma finite_measure_spmf [simp]: "finite_measure (measure_spmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 263 | by unfold_locales | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 264 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 265 | lemma spmf_conv_measure_spmf: "spmf p x = measure (measure_spmf p) {x}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 266 | by(auto simp: measure_spmf_def measure_distr measure_restrict_space pmf.rep_eq space_restrict_space intro: arg_cong2[where f=measure]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 267 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 268 | lemma emeasure_measure_spmf_conv_measure_pmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 269 | "emeasure (measure_spmf p) A = emeasure (measure_pmf p) (Some ` A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 270 | by(auto simp: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 271 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 272 | lemma measure_measure_spmf_conv_measure_pmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 273 | "measure (measure_spmf p) A = measure (measure_pmf p) (Some ` A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 274 | using emeasure_measure_spmf_conv_measure_pmf[of p A] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 275 | by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 276 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 277 | lemma emeasure_spmf_map_pmf_Some [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 278 | "emeasure (measure_spmf (map_pmf Some p)) A = emeasure (measure_pmf p) A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 279 | by(auto simp: measure_spmf_def emeasure_distr emeasure_restrict_space space_restrict_space intro: arg_cong2[where f=emeasure]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 280 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 281 | lemma measure_spmf_map_pmf_Some [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 282 | "measure (measure_spmf (map_pmf Some p)) A = measure (measure_pmf p) A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 283 | using emeasure_spmf_map_pmf_Some[of p A] by(simp add: measure_spmf.emeasure_eq_measure measure_pmf.emeasure_eq_measure) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 284 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 285 | lemma nn_integral_measure_spmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_spmf p) = \<integral>\<^sup>+ x. ennreal (spmf p x) * f x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 286 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 287 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 288 | have "?lhs = \<integral>\<^sup>+ x. pmf p x * f (the x) \<partial>count_space (range Some)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 289 | by(simp add: measure_spmf_def nn_integral_distr nn_integral_restrict_space nn_integral_measure_pmf nn_integral_count_space_indicator ac_simps | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 290 | flip: times_ereal.simps [symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 291 | also have "\<dots> = \<integral>\<^sup>+ x. ennreal (spmf p (the x)) * f (the x) \<partial>count_space (range Some)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 292 | by(rule nn_integral_cong) auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 293 | also have "\<dots> = \<integral>\<^sup>+ x. spmf p (the (Some x)) * f (the (Some x)) \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 294 | by(rule nn_integral_bij_count_space[symmetric])(simp add: bij_betw_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 295 | also have "\<dots> = ?rhs" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 296 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 297 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 298 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 299 | lemma integral_measure_spmf: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 300 | assumes "integrable (measure_spmf p) f" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 301 | shows "(\<integral> x. f x \<partial>measure_spmf p) = \<integral> x. spmf p x * f x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 302 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 303 | have "integrable (count_space UNIV) (\<lambda>x. spmf p x * f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 304 | using assms by(simp add: integrable_iff_bounded nn_integral_measure_spmf abs_mult ennreal_mult'') | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 305 | then show ?thesis using assms | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 306 | by(simp add: real_lebesgue_integral_def nn_integral_measure_spmf ennreal_mult'[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 307 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 308 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 309 | lemma emeasure_spmf_single: "emeasure (measure_spmf p) {x} = spmf p x"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 310 | by(simp add: measure_spmf.emeasure_eq_measure spmf_conv_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 311 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 312 | lemma measurable_measure_spmf[measurable]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 313 | "(\<lambda>x. measure_spmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 314 | by (auto simp: space_subprob_algebra) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 315 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 316 | lemma nn_integral_measure_spmf_conv_measure_pmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 317 | assumes [measurable]: "f \<in> borel_measurable (count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 318 | shows "nn_integral (measure_spmf p) f = nn_integral (restrict_space (measure_pmf p) (range Some)) (f \<circ> the)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 319 | by(simp add: measure_spmf_def nn_integral_distr o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 320 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 321 | lemma measure_spmf_in_space_subprob_algebra [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 322 | "measure_spmf p \<in> space (subprob_algebra (count_space UNIV))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 323 | by(simp add: space_subprob_algebra) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 324 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 325 | lemma nn_integral_spmf_neq_top: "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV) \<noteq> \<top>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 326 | using nn_integral_measure_spmf[where f="\<lambda>_. 1", of p, symmetric] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 327 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 328 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 329 | lemma SUP_spmf_neq_top': "(SUP p\<in>Y. ennreal (spmf p x)) \<noteq> \<top>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 330 | by (metis SUP_least ennreal_le_1 ennreal_one_neq_top neq_top_trans pmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 331 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 332 | lemma SUP_spmf_neq_top: "(SUP i. ennreal (spmf (Y i) x)) \<noteq> \<top>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 333 | by (meson SUP_eq_top_iff ennreal_le_1 ennreal_one_less_top linorder_not_le pmf_le_1) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 334 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 335 | lemma SUP_emeasure_spmf_neq_top: "(SUP p\<in>Y. emeasure (measure_spmf p) A) \<noteq> \<top>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 336 | by (metis ennreal_one_less_top less_SUP_iff linorder_not_le measure_spmf.subprob_emeasure_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 337 | |
| 63308 | 338 | subsection \<open>Support\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 339 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 340 | definition set_spmf :: "'a spmf \<Rightarrow> 'a set" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 341 | where "set_spmf p = set_pmf p \<bind> set_option" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 342 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 343 | lemma set_spmf_rep_eq: "set_spmf p = {x. measure (measure_spmf p) {x} \<noteq> 0}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 344 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 345 |   have "\<And>x :: 'a. the -` {x} \<inter> range Some = {Some x}" by auto
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 346 | then show ?thesis | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 347 | unfolding set_spmf_def measure_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 348 | by(auto simp: set_pmf.rep_eq measure_distr measure_restrict_space space_restrict_space) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 349 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 350 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 351 | lemma in_set_spmf: "x \<in> set_spmf p \<longleftrightarrow> Some x \<in> set_pmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 352 | by(simp add: set_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 353 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 354 | lemma AE_measure_spmf_iff [simp]: "(AE x in measure_spmf p. P x) \<longleftrightarrow> (\<forall>x\<in>set_spmf p. P x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 355 | unfolding set_spmf_def measure_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 356 | by(force simp: AE_distr_iff AE_restrict_space_iff AE_measure_pmf_iff cong del: AE_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 357 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 358 | lemma spmf_eq_0_set_spmf: "spmf p x = 0 \<longleftrightarrow> x \<notin> set_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 359 | by(auto simp: pmf_eq_0_set_pmf set_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 360 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 361 | lemma in_set_spmf_iff_spmf: "x \<in> set_spmf p \<longleftrightarrow> spmf p x \<noteq> 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 362 | by(auto simp: set_spmf_def set_pmf_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 363 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 364 | lemma set_spmf_return_pmf_None [simp]: "set_spmf (return_pmf None) = {}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 365 | by(auto simp: set_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 366 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 367 | lemma countable_set_spmf [simp]: "countable (set_spmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 368 | by(simp add: set_spmf_def bind_UNION) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 369 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 370 | lemma spmf_eqI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 371 | assumes "\<And>i. spmf p i = spmf q i" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 372 | shows "p = q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 373 | proof(rule pmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 374 | fix i | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 375 | show "pmf p i = pmf q i" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 376 | proof(cases i) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 377 | case (Some i') | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 378 | thus ?thesis by(simp add: assms) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 379 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 380 | case None | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 381 |     have "ennreal (pmf p i) = measure (measure_pmf p) {i}" by(simp add: pmf_def)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 382 |     also have "{i} = space (measure_pmf p) - range Some"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 383 | by(auto simp: None intro: ccontr) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 384 | also have "measure (measure_pmf p) \<dots> = ennreal 1 - measure (measure_pmf p) (range Some)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 385 | by(simp add: measure_pmf.prob_compl ennreal_minus[symmetric] del: space_measure_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 386 |     also have "range Some = (\<Union>x\<in>set_spmf p. {Some x}) \<union> Some ` (- set_spmf p)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 387 | by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 388 |     also have "measure (measure_pmf p) \<dots> = measure (measure_pmf p) (\<Union>x\<in>set_spmf p. {Some x})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 389 | by(rule measure_pmf.measure_zero_union)(auto simp: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 390 |     also have "ennreal \<dots> = \<integral>\<^sup>+ x. measure (measure_pmf p) {Some x} \<partial>count_space (set_spmf p)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 391 | unfolding measure_pmf.emeasure_eq_measure[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 392 | by(simp_all add: emeasure_UN_countable disjoint_family_on_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 393 | also have "\<dots> = \<integral>\<^sup>+ x. spmf p x \<partial>count_space (set_spmf p)" by(simp add: pmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 394 | also have "\<dots> = \<integral>\<^sup>+ x. spmf q x \<partial>count_space (set_spmf p)" by(simp add: assms) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 395 | also have "set_spmf p = set_spmf q" by(auto simp: in_set_spmf_iff_spmf assms) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 396 |     also have "(\<integral>\<^sup>+ x. spmf q x \<partial>count_space (set_spmf q)) = \<integral>\<^sup>+ x. measure (measure_pmf q) {Some x} \<partial>count_space (set_spmf q)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 397 | by(simp add: pmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 398 |     also have "\<dots> = measure (measure_pmf q) (\<Union>x\<in>set_spmf q. {Some x})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 399 | unfolding measure_pmf.emeasure_eq_measure[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 400 | by(simp_all add: emeasure_UN_countable disjoint_family_on_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 401 |     also have "\<dots> = measure (measure_pmf q) ((\<Union>x\<in>set_spmf q. {Some x}) \<union> Some ` (- set_spmf q))"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 402 | by(rule ennreal_cong measure_pmf.measure_zero_union[symmetric])+(auto simp: measure_pmf.prob_eq_0 AE_measure_pmf_iff in_set_spmf_iff_spmf set_pmf_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 403 |     also have "((\<Union>x\<in>set_spmf q. {Some x}) \<union> Some ` (- set_spmf q)) = range Some" by auto
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 404 | also have "ennreal 1 - measure (measure_pmf q) \<dots> = measure (measure_pmf q) (space (measure_pmf q) - range Some)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 405 | by(simp add: one_ereal_def measure_pmf.prob_compl ennreal_minus[symmetric] del: space_measure_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 406 |     also have "space (measure_pmf q) - range Some = {i}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 407 | by(auto simp: None intro: ccontr) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 408 | also have "measure (measure_pmf q) \<dots> = pmf q i" by(simp add: pmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 409 | finally show ?thesis by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 410 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 411 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 412 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 413 | lemma integral_measure_spmf_restrict: | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 414 |   fixes f ::  "'a \<Rightarrow> 'b :: {banach, second_countable_topology}"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 415 | shows "(\<integral> x. f x \<partial>measure_spmf M) = (\<integral> x. f x \<partial>restrict_space (measure_spmf M) (set_spmf M))" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 416 | by(auto intro!: integral_cong_AE simp add: integral_restrict_space) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 417 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 418 | lemma nn_integral_measure_spmf': | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 419 | "(\<integral>\<^sup>+ x. f x \<partial>measure_spmf p) = \<integral>\<^sup>+ x. ennreal (spmf p x) * f x \<partial>count_space (set_spmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 420 | by(auto simp: nn_integral_measure_spmf nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 421 | |
| 63308 | 422 | subsection \<open>Functorial structure\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 423 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 424 | abbreviation map_spmf :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a spmf \<Rightarrow> 'b spmf"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 425 | where "map_spmf f \<equiv> map_pmf (map_option f)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 426 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 427 | context begin | 
| 63308 | 428 | local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "spmf")\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 429 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 430 | lemma map_comp: "map_spmf f (map_spmf g p) = map_spmf (f \<circ> g) p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 431 | by(simp add: pmf.map_comp o_def option.map_comp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 432 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 433 | lemma map_id0: "map_spmf id = id" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 434 | by(simp add: pmf.map_id option.map_id0) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 435 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 436 | lemma map_id [simp]: "map_spmf id p = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 437 | by(simp add: map_id0) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 438 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 439 | lemma map_ident [simp]: "map_spmf (\<lambda>x. x) p = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 440 | by(simp add: id_def[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 441 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 442 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 443 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 444 | lemma set_map_spmf [simp]: "set_spmf (map_spmf f p) = f ` set_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 445 | by(simp add: set_spmf_def image_bind bind_image o_def Option.option.set_map) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 446 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 447 | lemma map_spmf_cong: | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 448 | "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> map_spmf f p = map_spmf g q" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 449 | by(auto intro: pmf.map_cong option.map_cong simp add: in_set_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 450 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 451 | lemma map_spmf_cong_simp: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 452 | "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q =simp=> f x = g x \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 453 | \<Longrightarrow> map_spmf f p = map_spmf g q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 454 | unfolding simp_implies_def by(rule map_spmf_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 455 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 456 | lemma map_spmf_idI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_spmf f p = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 457 | by(rule map_pmf_idI map_option_idI)+(simp add: in_set_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 458 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 459 | lemma emeasure_map_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 460 | "emeasure (measure_spmf (map_spmf f p)) A = emeasure (measure_spmf p) (f -` A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 461 | by(auto simp: measure_spmf_def emeasure_distr measurable_restrict_space1 space_restrict_space emeasure_restrict_space intro: arg_cong2[where f=emeasure]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 462 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 463 | lemma measure_map_spmf: "measure (measure_spmf (map_spmf f p)) A = measure (measure_spmf p) (f -` A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 464 | using emeasure_map_spmf[of f p A] by(simp add: measure_spmf.emeasure_eq_measure) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 465 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 466 | lemma measure_map_spmf_conv_distr: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 467 | "measure_spmf (map_spmf f p) = distr (measure_spmf p) (count_space UNIV) f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 468 | by(rule measure_eqI)(simp_all add: emeasure_map_spmf emeasure_distr) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 469 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 470 | lemma spmf_map_pmf_Some [simp]: "spmf (map_pmf Some p) i = pmf p i" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 471 | by(simp add: pmf_map_inj') | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 472 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 473 | lemma spmf_map_inj: "\<lbrakk> inj_on f (set_spmf M); x \<in> set_spmf M \<rbrakk> \<Longrightarrow> spmf (map_spmf f M) (f x) = spmf M x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 474 | by (smt (verit) elem_set in_set_spmf inj_on_def option.inj_map_strong option.map(2) pmf_map_inj) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 475 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 476 | lemma spmf_map_inj': "inj f \<Longrightarrow> spmf (map_spmf f M) (f x) = spmf M x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 477 | by(subst option.map(2)[symmetric, where f=f])(rule pmf_map_inj'[OF option.inj_map]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 478 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 479 | lemma spmf_map_outside: "x \<notin> f ` set_spmf M \<Longrightarrow> spmf (map_spmf f M) x = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 480 | unfolding spmf_eq_0_set_spmf by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 481 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 482 | lemma ennreal_spmf_map: "ennreal (spmf (map_spmf f p) x) = emeasure (measure_spmf p) (f -` {x})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 483 | by (metis emeasure_map_spmf emeasure_spmf_single) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 484 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 485 | lemma spmf_map: "spmf (map_spmf f p) x = measure (measure_spmf p) (f -` {x})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 486 | using ennreal_spmf_map[of f p x] by(simp add: measure_spmf.emeasure_eq_measure) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 487 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 488 | lemma ennreal_spmf_map_conv_nn_integral: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 489 |   "ennreal (spmf (map_spmf f p) x) = integral\<^sup>N (measure_spmf p) (indicator (f -` {x}))"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 490 | by (simp add: ennreal_spmf_map) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 491 | |
| 63308 | 492 | subsection \<open>Monad operations\<close> | 
| 493 | ||
| 494 | subsubsection \<open>Return\<close> | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 495 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 496 | abbreviation return_spmf :: "'a \<Rightarrow> 'a spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 497 | where "return_spmf x \<equiv> return_pmf (Some x)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 498 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 499 | lemma pmf_return_spmf: "pmf (return_spmf x) y = indicator {y} (Some x)"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 500 | by(fact pmf_return) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 501 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 502 | lemma measure_spmf_return_spmf: "measure_spmf (return_spmf x) = Giry_Monad.return (count_space UNIV) x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 503 | by(rule measure_eqI)(simp_all add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space indicator_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 504 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 505 | lemma measure_spmf_return_pmf_None [simp]: "measure_spmf (return_pmf None) = null_measure (count_space UNIV)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 506 | by (simp add: emeasure_measure_spmf_conv_measure_pmf measure_eqI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 507 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 508 | lemma set_return_spmf [simp]: "set_spmf (return_spmf x) = {x}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 509 | by(auto simp: set_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 510 | |
| 63308 | 511 | subsubsection \<open>Bind\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 512 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 513 | definition bind_spmf :: "'a spmf \<Rightarrow> ('a \<Rightarrow> 'b spmf) \<Rightarrow> 'b spmf"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 514 | where "bind_spmf x f = bind_pmf x (\<lambda>a. case a of None \<Rightarrow> return_pmf None | Some a' \<Rightarrow> f a')" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 515 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 516 | adhoc_overloading Monad_Syntax.bind bind_spmf | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 517 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 518 | lemma return_None_bind_spmf [simp]: "return_pmf None \<bind> (f :: 'a \<Rightarrow> _) = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 519 | by(simp add: bind_spmf_def bind_return_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 520 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 521 | lemma return_bind_spmf [simp]: "return_spmf x \<bind> f = f x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 522 | by(simp add: bind_spmf_def bind_return_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 523 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 524 | lemma bind_return_spmf [simp]: "x \<bind> return_spmf = x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 525 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 526 | have "\<And>a :: 'a option. (case a of None \<Rightarrow> return_pmf None | Some a' \<Rightarrow> return_spmf a') = return_pmf a" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 527 | by(simp split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 528 | then show ?thesis | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 529 | by(simp add: bind_spmf_def bind_return_pmf') | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 530 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 531 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 532 | lemma bind_spmf_assoc [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 533 | fixes x :: "'a spmf" and f :: "'a \<Rightarrow> 'b spmf" and g :: "'b \<Rightarrow> 'c spmf" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 534 | shows "(x \<bind> f) \<bind> g = x \<bind> (\<lambda>y. f y \<bind> g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 535 | unfolding bind_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 536 | by (smt (verit, best) bind_assoc_pmf bind_pmf_cong bind_return_pmf option.case_eq_if) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 537 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 538 | lemma pmf_bind_spmf_None: "pmf (p \<bind> f) None = pmf p None + \<integral> x. pmf (f x) None \<partial>measure_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 539 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 540 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 541 | let ?f = "\<lambda>x. pmf (case x of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x) None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 542 | have "?lhs = \<integral> x. ?f x \<partial>measure_pmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 543 | by(simp add: bind_spmf_def pmf_bind) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 544 |   also have "\<dots> = \<integral> x. ?f None * indicator {None} x + ?f x * indicator (range Some) x \<partial>measure_pmf p"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 545 | by(rule Bochner_Integration.integral_cong)(auto simp: indicator_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 546 |   also have "\<dots> = (\<integral> x. ?f None * indicator {None} x \<partial>measure_pmf p) + (\<integral> x. ?f x * indicator (range Some) x \<partial>measure_pmf p)"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 547 | by(rule Bochner_Integration.integral_add)(auto 4 3 intro: integrable_real_mult_indicator measure_pmf.integrable_const_bound[where B=1] simp add: AE_measure_pmf_iff pmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 548 | also have "\<dots> = pmf p None + \<integral> x. indicator (range Some) x * pmf (f (the x)) None \<partial>measure_pmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 549 | by(auto simp: measure_measure_pmf_finite indicator_eq_0_iff intro!: Bochner_Integration.integral_cong) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 550 | also have "\<dots> = ?rhs" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 551 | unfolding measure_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 552 | by(subst integral_distr)(auto simp: integral_restrict_space) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 553 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 554 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 555 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 556 | lemma spmf_bind: "spmf (p \<bind> f) y = \<integral> x. spmf (f x) y \<partial>measure_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 557 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 558 | have "\<And>x. spmf (case x of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x) y = | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 559 | indicat_real (range Some) x * spmf (f (the x)) y" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 560 | by (simp add: split: option.split) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 561 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 562 | by (simp add: measure_spmf_def integral_distr bind_spmf_def pmf_bind integral_restrict_space) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 563 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 564 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 565 | lemma ennreal_spmf_bind: "ennreal (spmf (p \<bind> f) x) = \<integral>\<^sup>+ y. spmf (f y) x \<partial>measure_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 566 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 567 | have "\<And>y. ennreal (spmf (case y of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x) x) = | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 568 | ennreal (spmf (f (the y)) x) * indicator (range Some) y" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 569 | by (simp add: split: option.split) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 570 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 571 | by (simp add: bind_spmf_def ennreal_pmf_bind nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 572 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 573 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 574 | lemma measure_spmf_bind_pmf: "measure_spmf (p \<bind> f) = measure_pmf p \<bind> measure_spmf \<circ> f" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 575 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 576 | proof(rule measure_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 577 | show "sets ?lhs = sets ?rhs" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 578 | by (simp add: Giry_Monad.bind_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 579 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 580 | fix A :: "'a set" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 581 | have "emeasure ?lhs A = \<integral>\<^sup>+ x. emeasure (measure_spmf (f x)) A \<partial>measure_pmf p" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 582 | by(simp add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space bind_spmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 583 | also have "\<dots> = emeasure ?rhs A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 584 | by(simp add: emeasure_bind[where N="count_space UNIV"] space_measure_spmf space_subprob_algebra) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 585 | finally show "emeasure ?lhs A = emeasure ?rhs A" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 586 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 587 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 588 | lemma measure_spmf_bind: "measure_spmf (p \<bind> f) = measure_spmf p \<bind> measure_spmf \<circ> f" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 589 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 590 | proof(rule measure_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 591 | show "sets ?lhs = sets ?rhs" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 592 | by(simp add: sets_bind[where N="count_space UNIV"] space_measure_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 593 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 594 | fix A :: "'a set" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 595 | let ?A = "the -` A \<inter> range Some" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 596 | have "emeasure ?lhs A = \<integral>\<^sup>+ x. emeasure (measure_pmf (case x of None \<Rightarrow> return_pmf None | Some x \<Rightarrow> f x)) ?A \<partial>measure_pmf p" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 597 | by(simp add: measure_spmf_def emeasure_distr space_restrict_space emeasure_restrict_space bind_spmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 598 | also have "\<dots> = \<integral>\<^sup>+ x. emeasure (measure_pmf (f (the x))) ?A * indicator (range Some) x \<partial>measure_pmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 599 | by(rule nn_integral_cong)(auto split: option.split simp add: indicator_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 600 | also have "\<dots> = \<integral>\<^sup>+ x. emeasure (measure_spmf (f x)) A \<partial>measure_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 601 | by(simp add: measure_spmf_def nn_integral_distr nn_integral_restrict_space emeasure_distr space_restrict_space emeasure_restrict_space) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 602 | also have "\<dots> = emeasure ?rhs A" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 603 | by(simp add: emeasure_bind[where N="count_space UNIV"] space_measure_spmf space_subprob_algebra) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 604 | finally show "emeasure ?lhs A = emeasure ?rhs A" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 605 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 606 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 607 | lemma map_spmf_bind_spmf: "map_spmf f (bind_spmf p g) = bind_spmf p (map_spmf f \<circ> g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 608 | by(auto simp: bind_spmf_def map_bind_pmf fun_eq_iff split: option.split intro: arg_cong2[where f=bind_pmf]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 609 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 610 | lemma bind_map_spmf: "map_spmf f p \<bind> g = p \<bind> g \<circ> f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 611 | by(simp add: bind_spmf_def bind_map_pmf o_def cong del: option.case_cong_weak) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 612 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 613 | lemma spmf_bind_leI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 614 | assumes "\<And>y. y \<in> set_spmf p \<Longrightarrow> spmf (f y) x \<le> r" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 615 | and "0 \<le> r" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 616 | shows "spmf (bind_spmf p f) x \<le> r" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 617 | proof - | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 618 | have "ennreal (spmf (bind_spmf p f) x) = \<integral>\<^sup>+ y. spmf (f y) x \<partial>measure_spmf p" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 619 | by(rule ennreal_spmf_bind) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 620 | also have "\<dots> \<le> \<integral>\<^sup>+ y. r \<partial>measure_spmf p" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 621 | by(rule nn_integral_mono_AE)(simp add: assms) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 622 | also have "\<dots> \<le> r" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 623 | using assms measure_spmf.emeasure_space_le_1 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 624 | by(auto simp: measure_spmf.emeasure_eq_measure intro!: mult_left_le) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 625 | finally show ?thesis using assms(2) by(simp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 626 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 627 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 628 | lemma map_spmf_conv_bind_spmf: "map_spmf f p = (p \<bind> (\<lambda>x. return_spmf (f x)))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 629 | by(simp add: map_pmf_def bind_spmf_def)(rule bind_pmf_cong, simp_all split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 630 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 631 | lemma bind_spmf_cong: | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 632 | "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> bind_spmf p f = bind_spmf q g" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 633 | by(auto simp: bind_spmf_def in_set_spmf intro: bind_pmf_cong option.case_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 634 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 635 | lemma bind_spmf_cong_simp: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 636 | "\<lbrakk> p = q; \<And>x. x \<in> set_spmf q =simp=> f x = g x \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 637 | \<Longrightarrow> bind_spmf p f = bind_spmf q g" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 638 | by(simp add: simp_implies_def cong: bind_spmf_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 639 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 640 | lemma set_bind_spmf: "set_spmf (M \<bind> f) = set_spmf M \<bind> (set_spmf \<circ> f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 641 | by(auto simp: set_spmf_def bind_spmf_def bind_UNION split: option.splits) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 642 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 643 | lemma bind_spmf_const_return_None [simp]: "bind_spmf p (\<lambda>_. return_pmf None) = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 644 | by(simp add: bind_spmf_def case_option_collapse) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 645 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 646 | lemma bind_commute_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 647 | "bind_spmf p (\<lambda>x. bind_spmf q (f x)) = bind_spmf q (\<lambda>y. bind_spmf p (\<lambda>x. f x y))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 648 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 649 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 650 | let ?f = "\<lambda>x y. case x of None \<Rightarrow> return_pmf None | Some a \<Rightarrow> (case y of None \<Rightarrow> return_pmf None | Some b \<Rightarrow> f a b)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 651 | have "?lhs = p \<bind> (\<lambda>x. q \<bind> ?f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 652 | unfolding bind_spmf_def by(rule bind_pmf_cong[OF refl])(simp split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 653 | also have "\<dots> = q \<bind> (\<lambda>y. p \<bind> (\<lambda>x. ?f x y))" by(rule bind_commute_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 654 | also have "\<dots> = ?rhs" unfolding bind_spmf_def | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 655 | by(rule bind_pmf_cong[OF refl])(auto split: option.split, metis bind_spmf_const_return_None bind_spmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 656 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 657 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 658 | |
| 63308 | 659 | subsection \<open>Relator\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 660 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 661 | abbreviation rel_spmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> 'b spmf \<Rightarrow> bool"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 662 | where "rel_spmf R \<equiv> rel_pmf (rel_option R)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 663 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 664 | lemma rel_spmf_mono: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 665 | "\<lbrakk>rel_spmf A f g; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_spmf B f g" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 666 | by (metis option.rel_sel pmf.rel_mono_strong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 667 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 668 | lemma rel_spmf_mono_strong: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 669 | "\<lbrakk> rel_spmf A f g; \<And>x y. \<lbrakk> A x y; x \<in> set_spmf f; y \<in> set_spmf g \<rbrakk> \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_spmf B f g" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 670 | by (metis elem_set in_set_spmf option.rel_mono_strong pmf.rel_mono_strong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 671 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 672 | lemma rel_spmf_reflI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> P x x) \<Longrightarrow> rel_spmf P p p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 673 | by (metis (mono_tags, lifting) option.rel_eq pmf.rel_eq rel_spmf_mono_strong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 674 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 675 | lemma rel_spmfI [intro?]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 676 | "\<lbrakk> \<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> P x y; map_spmf fst pq = p; map_spmf snd pq = q \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 677 | \<Longrightarrow> rel_spmf P p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 678 | by(rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. case x of None \<Rightarrow> (None, None) | Some (a, b) \<Rightarrow> (Some a, Some b)) pq"]) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 679 | (auto simp: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 680 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 681 | lemma rel_spmfE [elim?, consumes 1, case_names rel_spmf]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 682 | assumes "rel_spmf P p q" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 683 | obtains pq where | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 684 | "\<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> P x y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 685 | "p = map_spmf fst pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 686 | "q = map_spmf snd pq" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 687 | using assms | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 688 | proof(cases rule: rel_pmf.cases[consumes 1, case_names rel_pmf]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 689 | case (rel_pmf pq) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 690 | let ?pq = "map_pmf (\<lambda>(a, b). case (a, b) of (Some x, Some y) \<Rightarrow> Some (x, y) | _ \<Rightarrow> None) pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 691 | have "\<And>x y. (x, y) \<in> set_spmf ?pq \<Longrightarrow> P x y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 692 | by(auto simp: in_set_spmf split: option.split_asm dest: rel_pmf(1)) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 693 | moreover | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 694 | have "\<And>x. (x, None) \<in> set_pmf pq \<Longrightarrow> x = None" by(auto dest!: rel_pmf(1)) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 695 | then have "p = map_spmf fst ?pq" using rel_pmf(2) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 696 | by(auto simp: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 697 | moreover | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 698 | have "\<And>y. (None, y) \<in> set_pmf pq \<Longrightarrow> y = None" by(auto dest!: rel_pmf(1)) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 699 | then have "q = map_spmf snd ?pq" using rel_pmf(3) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 700 | by(auto simp: pmf.map_comp split_beta intro!: pmf.map_cong split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 701 | ultimately show thesis .. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 702 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 703 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 704 | lemma rel_spmf_simps: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 705 | "rel_spmf R p q \<longleftrightarrow> (\<exists>pq. (\<forall>(x, y)\<in>set_spmf pq. R x y) \<and> map_spmf fst pq = p \<and> map_spmf snd pq = q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 706 | by(auto intro: rel_spmfI elim!: rel_spmfE) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 707 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 708 | lemma spmf_rel_map: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 709 | shows spmf_rel_map1: "\<And>R f x. rel_spmf R (map_spmf f x) = rel_spmf (\<lambda>x. R (f x)) x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 710 | and spmf_rel_map2: "\<And>R x g y. rel_spmf R x (map_spmf g y) = rel_spmf (\<lambda>x y. R x (g y)) x y" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 711 | by(simp_all add: fun_eq_iff pmf.rel_map option.rel_map[abs_def]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 712 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 713 | lemma spmf_rel_conversep: "rel_spmf R\<inverse>\<inverse> = (rel_spmf R)\<inverse>\<inverse>" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 714 | by(simp add: option.rel_conversep pmf.rel_conversep) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 715 | |
| 67399 | 716 | lemma spmf_rel_eq: "rel_spmf (=) = (=)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 717 | by(simp add: pmf.rel_eq option.rel_eq) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 718 | |
| 63343 | 719 | context includes lifting_syntax | 
| 720 | begin | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 721 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 722 | lemma bind_spmf_parametric [transfer_rule]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 723 | "(rel_spmf A ===> (A ===> rel_spmf B) ===> rel_spmf B) bind_spmf bind_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 724 | unfolding bind_spmf_def[abs_def] by transfer_prover | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 725 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 726 | lemma return_spmf_parametric: "(A ===> rel_spmf A) return_spmf return_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 727 | by transfer_prover | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 728 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 729 | lemma map_spmf_parametric: "((A ===> B) ===> rel_spmf A ===> rel_spmf B) map_spmf map_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 730 | by transfer_prover | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 731 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 732 | lemma rel_spmf_parametric: | 
| 67399 | 733 | "((A ===> B ===> (=)) ===> rel_spmf A ===> rel_spmf B ===> (=)) rel_spmf rel_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 734 | by transfer_prover | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 735 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 736 | lemma set_spmf_parametric [transfer_rule]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 737 | "(rel_spmf A ===> rel_set A) set_spmf set_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 738 | unfolding set_spmf_def[abs_def] by transfer_prover | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 739 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 740 | lemma return_spmf_None_parametric: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 741 | "(rel_spmf A) (return_pmf None) (return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 742 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 743 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 744 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 745 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 746 | lemma rel_spmf_bindI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 747 | "\<lbrakk> rel_spmf R p q; \<And>x y. R x y \<Longrightarrow> rel_spmf P (f x) (g y) \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 748 | \<Longrightarrow> rel_spmf P (p \<bind> f) (q \<bind> g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 749 | by(fact bind_spmf_parametric[THEN rel_funD, THEN rel_funD, OF _ rel_funI]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 750 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 751 | lemma rel_spmf_bind_reflI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 752 | "(\<And>x. x \<in> set_spmf p \<Longrightarrow> rel_spmf P (f x) (g x)) \<Longrightarrow> rel_spmf P (p \<bind> f) (p \<bind> g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 753 | by(rule rel_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: rel_spmf_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 754 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 755 | lemma rel_pmf_return_pmfI: "P x y \<Longrightarrow> rel_pmf P (return_pmf x) (return_pmf y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 756 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 757 | |
| 63343 | 758 | context includes lifting_syntax | 
| 759 | begin | |
| 760 | ||
| 69597 | 761 | text \<open>We do not yet have a relator for \<^typ>\<open>'a measure\<close>, so we combine \<^const>\<open>measure\<close> and \<^const>\<open>measure_pmf\<close>\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 762 | lemma measure_pmf_parametric: | 
| 67399 | 763 | "(rel_pmf A ===> rel_pred A ===> (=)) (\<lambda>p. measure (measure_pmf p)) (\<lambda>q. measure (measure_pmf q))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 764 | proof(rule rel_funI)+ | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 765 | fix p q X Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 766 | assume "rel_pmf A p q" and "rel_pred A X Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 767 | from this(1) obtain pq where A: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> A x y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 768 | and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 769 | show "measure p X = measure q Y" unfolding p q measure_map_pmf | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 770 | by(rule measure_pmf.finite_measure_eq_AE)(auto simp: AE_measure_pmf_iff dest!: A rel_predD[OF \<open>rel_pred _ _ _\<close>]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 771 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 772 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 773 | lemma measure_spmf_parametric: | 
| 67399 | 774 | "(rel_spmf A ===> rel_pred A ===> (=)) (\<lambda>p. measure (measure_spmf p)) (\<lambda>q. measure (measure_spmf q))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 775 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 776 | have "\<And>x y xa ya. rel_pred A xa ya \<Longrightarrow> rel_pred (rel_option A) (Some ` xa) (Some ` ya)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 777 | by(auto simp: rel_pred_def rel_fun_def elim: option.rel_cases) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 778 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 779 | unfolding measure_measure_spmf_conv_measure_pmf[abs_def] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 780 | by (intro rel_funI) (force elim!: measure_pmf_parametric[THEN rel_funD, THEN rel_funD]) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 781 | qed | 
| 63343 | 782 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 783 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 784 | |
| 69597 | 785 | subsection \<open>From \<^typ>\<open>'a pmf\<close> to \<^typ>\<open>'a spmf\<close>\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 786 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 787 | definition spmf_of_pmf :: "'a pmf \<Rightarrow> 'a spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 788 | where "spmf_of_pmf = map_pmf Some" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 789 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 790 | lemma set_spmf_spmf_of_pmf [simp]: "set_spmf (spmf_of_pmf p) = set_pmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 791 | by(auto simp: spmf_of_pmf_def set_spmf_def bind_image o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 792 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 793 | lemma spmf_spmf_of_pmf [simp]: "spmf (spmf_of_pmf p) x = pmf p x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 794 | by(simp add: spmf_of_pmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 795 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 796 | lemma pmf_spmf_of_pmf_None [simp]: "pmf (spmf_of_pmf p) None = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 797 | using ennreal_pmf_map[of Some p None] by(simp add: spmf_of_pmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 798 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 799 | lemma emeasure_spmf_of_pmf [simp]: "emeasure (measure_spmf (spmf_of_pmf p)) A = emeasure (measure_pmf p) A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 800 | by(simp add: emeasure_measure_spmf_conv_measure_pmf spmf_of_pmf_def inj_vimage_image_eq) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 801 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 802 | lemma measure_spmf_spmf_of_pmf [simp]: "measure_spmf (spmf_of_pmf p) = measure_pmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 803 | by(rule measure_eqI) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 804 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 805 | lemma map_spmf_of_pmf [simp]: "map_spmf f (spmf_of_pmf p) = spmf_of_pmf (map_pmf f p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 806 | by(simp add: spmf_of_pmf_def pmf.map_comp o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 807 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 808 | lemma rel_spmf_spmf_of_pmf [simp]: "rel_spmf R (spmf_of_pmf p) (spmf_of_pmf q) = rel_pmf R p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 809 | by(simp add: spmf_of_pmf_def pmf.rel_map) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 810 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 811 | lemma spmf_of_pmf_return_pmf [simp]: "spmf_of_pmf (return_pmf x) = return_spmf x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 812 | by(simp add: spmf_of_pmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 813 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 814 | lemma bind_spmf_of_pmf [simp]: "bind_spmf (spmf_of_pmf p) f = bind_pmf p f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 815 | by(simp add: spmf_of_pmf_def bind_spmf_def bind_map_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 816 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 817 | lemma set_spmf_bind_pmf: "set_spmf (bind_pmf p f) = Set.bind (set_pmf p) (set_spmf \<circ> f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 818 | unfolding bind_spmf_of_pmf[symmetric] by(subst set_bind_spmf) simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 819 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 820 | lemma spmf_of_pmf_bind: "spmf_of_pmf (bind_pmf p f) = bind_pmf p (\<lambda>x. spmf_of_pmf (f x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 821 | by(simp add: spmf_of_pmf_def map_bind_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 822 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 823 | lemma bind_pmf_return_spmf: "p \<bind> (\<lambda>x. return_spmf (f x)) = spmf_of_pmf (map_pmf f p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 824 | by(simp add: map_pmf_def spmf_of_pmf_bind) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 825 | |
| 63308 | 826 | subsection \<open>Weight of a subprobability\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 827 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 828 | abbreviation weight_spmf :: "'a spmf \<Rightarrow> real" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 829 | where "weight_spmf p \<equiv> measure (measure_spmf p) (space (measure_spmf p))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 830 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 831 | lemma weight_spmf_def: "weight_spmf p = measure (measure_spmf p) UNIV" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 832 | by(simp add: space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 833 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 834 | lemma weight_spmf_le_1: "weight_spmf p \<le> 1" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 835 | by(rule measure_spmf.subprob_measure_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 836 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 837 | lemma weight_return_spmf [simp]: "weight_spmf (return_spmf x) = 1" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 838 | by(simp add: measure_spmf_return_spmf measure_return) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 839 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 840 | lemma weight_return_pmf_None [simp]: "weight_spmf (return_pmf None) = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 841 | by(simp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 842 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 843 | lemma weight_map_spmf [simp]: "weight_spmf (map_spmf f p) = weight_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 844 | by(simp add: weight_spmf_def measure_map_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 845 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 846 | lemma weight_spmf_of_pmf [simp]: "weight_spmf (spmf_of_pmf p) = 1" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 847 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 848 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 849 | lemma weight_spmf_nonneg: "weight_spmf p \<ge> 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 850 | by(fact measure_nonneg) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 851 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 852 | lemma (in finite_measure) integrable_weight_spmf [simp]: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 853 | "(\<lambda>x. weight_spmf (f x)) \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. weight_spmf (f x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 854 | by(rule integrable_const_bound[where B=1])(simp_all add: weight_spmf_nonneg weight_spmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 855 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 856 | lemma weight_spmf_eq_nn_integral_spmf: "weight_spmf p = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 857 | by (metis NO_MATCH_def measure_spmf.emeasure_eq_measure nn_integral_count_space_indicator nn_integral_indicator nn_integral_measure_spmf sets_UNIV sets_measure_spmf space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 858 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 859 | lemma weight_spmf_eq_nn_integral_support: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 860 | "weight_spmf p = \<integral>\<^sup>+ x. spmf p x \<partial>count_space (set_spmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 861 | unfolding weight_spmf_eq_nn_integral_spmf | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 862 | by(auto simp: nn_integral_count_space_indicator in_set_spmf_iff_spmf intro!: nn_integral_cong split: split_indicator) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 863 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 864 | lemma pmf_None_eq_weight_spmf: "pmf p None = 1 - weight_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 865 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 866 | have "weight_spmf p = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV" by(rule weight_spmf_eq_nn_integral_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 867 | also have "\<dots> = \<integral>\<^sup>+ x. ennreal (pmf p x) * indicator (range Some) x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 868 | by(simp add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] nn_integral_embed_measure measurable_embed_measure1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 869 |   also have "\<dots> + pmf p None = \<integral>\<^sup>+ x. ennreal (pmf p x) * indicator (range Some) x + ennreal (pmf p None) * indicator {None} x \<partial>count_space UNIV"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 870 | by(subst nn_integral_add)(simp_all add: max_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 871 | also have "\<dots> = \<integral>\<^sup>+ x. pmf p x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 872 | by(rule nn_integral_cong)(auto split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 873 | also have "\<dots> = 1" by (simp add: nn_integral_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 874 | finally show ?thesis by(simp add: ennreal_plus[symmetric] del: ennreal_plus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 875 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 876 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 877 | lemma weight_spmf_conv_pmf_None: "weight_spmf p = 1 - pmf p None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 878 | by(simp add: pmf_None_eq_weight_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 879 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 880 | lemma weight_spmf_lt_0: "\<not> weight_spmf p < 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 881 | by(simp add: not_less weight_spmf_nonneg) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 882 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 883 | lemma spmf_le_weight: "spmf p x \<le> weight_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 884 | by (simp add: measure_spmf.bounded_measure spmf_conv_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 885 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 886 | lemma weight_spmf_eq_0: "weight_spmf p = 0 \<longleftrightarrow> p = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 887 | by (metis measure_le_0_iff measure_spmf.bounded_measure spmf_conv_measure_spmf spmf_eqI weight_return_pmf_None) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 888 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 889 | lemma weight_bind_spmf: "weight_spmf (x \<bind> f) = lebesgue_integral (measure_spmf x) (weight_spmf \<circ> f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 890 | unfolding weight_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 891 | by(simp add: measure_spmf_bind o_def measure_spmf.measure_bind[where N="count_space UNIV"]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 892 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 893 | lemma rel_spmf_weightD: "rel_spmf A p q \<Longrightarrow> weight_spmf p = weight_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 894 | by(erule rel_spmfE) simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 895 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 896 | lemma rel_spmf_bij_betw: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 897 | assumes f: "bij_betw f (set_spmf p) (set_spmf q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 898 | and eq: "\<And>x. x \<in> set_spmf p \<Longrightarrow> spmf p x = spmf q (f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 899 | shows "rel_spmf (\<lambda>x y. f x = y) p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 900 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 901 | let ?f = "map_option f" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 902 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 903 | have weq: "ennreal (weight_spmf p) = ennreal (weight_spmf q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 904 | unfolding weight_spmf_eq_nn_integral_support | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 905 | by(subst nn_integral_bij_count_space[OF f, symmetric])(rule nn_integral_cong_AE, simp add: eq AE_count_space) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 906 | then have "None \<in> set_pmf p \<longleftrightarrow> None \<in> set_pmf q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 907 | by(simp add: pmf_None_eq_weight_spmf set_pmf_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 908 | with f have "bij_betw (map_option f) (set_pmf p) (set_pmf q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 909 | apply(auto simp: bij_betw_def in_set_spmf inj_on_def intro: option.expand split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 910 | apply(rename_tac [!] x) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 911 | apply(case_tac [!] x) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 912 | apply(auto iff: in_set_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 913 | done | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 914 | then have "rel_pmf (\<lambda>x y. ?f x = y) p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 915 | proof (rule rel_pmf_bij_betw) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 916 | show "pmf p x = pmf q (map_option f x)" if "x \<in> set_pmf p" for x | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 917 | proof (cases x) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 918 | case None | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 919 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 920 | by (metis ennreal_inj measure_nonneg option.map_disc_iff pmf_None_eq_weight_spmf weq) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 921 | qed (use eq in_set_spmf that in force) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 922 | qed | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 923 | thus ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 924 | by (smt (verit, ccfv_SIG) None_eq_map_option_iff option.map_sel option.rel_sel pmf.rel_mono_strong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 925 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 926 | |
| 63308 | 927 | subsection \<open>From density to spmfs\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 928 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 929 | context fixes f :: "'a \<Rightarrow> real" begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 930 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 931 | definition embed_spmf :: "'a spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 932 | where "embed_spmf = embed_pmf (\<lambda>x. case x of None \<Rightarrow> 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) | Some x' \<Rightarrow> max 0 (f x'))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 933 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 934 | context | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 935 | assumes prob: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 936 | begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 937 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 938 | lemma nn_integral_embed_spmf_eq_1: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 939 | "(\<integral>\<^sup>+ x. ennreal (case x of None \<Rightarrow> 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV) | Some x' \<Rightarrow> max 0 (f x')) \<partial>count_space UNIV) = 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 940 | (is "?lhs = _" is "(\<integral>\<^sup>+ x. ?f x \<partial>?M) = _") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 941 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 942 |   have "?lhs = \<integral>\<^sup>+ x. ?f x * indicator {None} x + ?f x * indicator (range Some) x \<partial>?M"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 943 | by(rule nn_integral_cong)(auto split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 944 | also have "\<dots> = (1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV)) + \<integral>\<^sup>+ x. ?f x * indicator (range Some) x \<partial>?M" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 945 | (is "_ = ?None + ?Some") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 946 | by(subst nn_integral_add)(simp_all add: AE_count_space max_def le_diff_eq real_le_ereal_iff one_ereal_def[symmetric] prob split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 947 | also have "?Some = \<integral>\<^sup>+ x. ?f x \<partial>count_space (range Some)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 948 | by(simp add: nn_integral_count_space_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 949 | also have "count_space (range Some) = embed_measure (count_space UNIV) Some" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 950 | by(simp add: embed_measure_count_space) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 951 | also have "(\<integral>\<^sup>+ x. ?f x \<partial>\<dots>) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 952 | by(subst nn_integral_embed_measure)(simp_all add: measurable_embed_measure1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 953 | also have "?None + \<dots> = 1" using prob | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 954 | by(auto simp: ennreal_minus[symmetric] ennreal_1[symmetric] ennreal_enn2real_if top_unique simp del: ennreal_1)(simp add: diff_add_self_ennreal) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 955 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 956 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 957 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 958 | lemma pmf_embed_spmf_None: "pmf embed_spmf None = 1 - enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 959 | unfolding embed_spmf_def | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 960 | by (smt (verit, del_insts) enn2real_leI ennreal_1 nn_integral_cong nn_integral_embed_spmf_eq_1 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 961 | option.case_eq_if pmf_embed_pmf prob) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 962 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 963 | lemma spmf_embed_spmf [simp]: "spmf embed_spmf x = max 0 (f x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 964 | unfolding embed_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 965 | by (smt (verit, best) enn2real_leI ennreal_1 nn_integral_cong nn_integral_embed_spmf_eq_1 option.case_eq_if option.simps(5) pmf_embed_pmf prob) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 966 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 967 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 968 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 969 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 970 | |
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 971 | lemma embed_spmf_K_0[simp]: "embed_spmf (\<lambda>_. 0) = return_pmf None" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 972 | by(rule spmf_eqI)(simp add: zero_ereal_def[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 973 | |
| 63308 | 974 | subsection \<open>Ordering on spmfs\<close> | 
| 975 | ||
| 976 | text \<open> | |
| 69597 | 977 | \<^const>\<open>rel_pmf\<close> does not preserve a ccpo structure. Counterexample by Saheb-Djahromi: | 
| 63308 | 978 | Take prefix order over \<open>bool llist\<close> and | 
| 979 | the set \<open>range (\<lambda>n :: nat. uniform (llist_n n))\<close> where \<open>llist_n\<close> is the set | |
| 980 | of all \<open>llist\<close>s of length \<open>n\<close> and \<open>uniform\<close> returns a uniform distribution over | |
| 981 | the given set. The set forms a chain in \<open>ord_pmf lprefix\<close>, but it has not an upper bound. | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 982 | Any upper bound may contain only infinite lists in its support because otherwise it is not greater | 
| 63308 | 983 | than the \<open>n+1\<close>-st element in the chain where \<open>n\<close> is the length of the finite list. | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 984 | Moreover its support must contain all infinite lists, because otherwise there is a finite list | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 985 | all of whose finite extensions are not in the support - a contradiction to the upper bound property. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 986 | Hence, the support is uncountable, but pmf's only have countable support. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 987 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 988 | However, if all chains in the ccpo are finite, then it should preserve the ccpo structure. | 
| 63308 | 989 | \<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 990 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 991 | abbreviation ord_spmf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf \<Rightarrow> bool"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 992 | where "ord_spmf ord \<equiv> rel_pmf (ord_option ord)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 993 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 994 | locale ord_spmf_syntax begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 995 | notation ord_spmf (infix "\<sqsubseteq>\<index>" 60) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 996 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 997 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 998 | lemma ord_spmf_map_spmf1: "ord_spmf R (map_spmf f p) = ord_spmf (\<lambda>x. R (f x)) p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 999 | by(simp add: pmf.rel_map[abs_def] ord_option_map1[abs_def]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1000 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1001 | lemma ord_spmf_map_spmf2: "ord_spmf R p (map_spmf f q) = ord_spmf (\<lambda>x y. R x (f y)) p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1002 | by(simp add: pmf.rel_map ord_option_map2) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1003 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1004 | lemma ord_spmf_map_spmf12: "ord_spmf R (map_spmf f p) (map_spmf f q) = ord_spmf (\<lambda>x y. R (f x) (f y)) p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1005 | by(simp add: pmf.rel_map ord_option_map1[abs_def] ord_option_map2) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1006 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1007 | lemmas ord_spmf_map_spmf = ord_spmf_map_spmf1 ord_spmf_map_spmf2 ord_spmf_map_spmf12 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1008 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1009 | context fixes ord :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (structure) begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1010 | interpretation ord_spmf_syntax . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1011 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1012 | lemma ord_spmfI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1013 | "\<lbrakk> \<And>x y. (x, y) \<in> set_spmf pq \<Longrightarrow> ord x y; map_spmf fst pq = p; map_spmf snd pq = q \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1014 | \<Longrightarrow> p \<sqsubseteq> q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1015 | by(rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. case x of None \<Rightarrow> (None, None) | Some (a, b) \<Rightarrow> (Some a, Some b)) pq"]) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1016 | (auto simp: pmf.map_comp o_def in_set_spmf split: option.splits intro: pmf.map_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1017 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1018 | lemma ord_spmf_None [simp]: "return_pmf None \<sqsubseteq> x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1019 | by(rule rel_pmf.intros[where pq="map_pmf (Pair None) x"])(auto simp: pmf.map_comp o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1020 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1021 | lemma ord_spmf_reflI: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord x x) \<Longrightarrow> p \<sqsubseteq> p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1022 | by (metis elem_set in_set_spmf ord_option_reflI pmf.rel_refl_strong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1023 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1024 | lemma rel_spmf_inf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1025 | assumes "p \<sqsubseteq> q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1026 | and "q \<sqsubseteq> p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1027 | and refl: "reflp ord" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1028 | and trans: "transp ord" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1029 | shows "rel_spmf (inf ord ord\<inverse>\<inverse>) p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1030 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1031 | from \<open>p \<sqsubseteq> q\<close> \<open>q \<sqsubseteq> p\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1032 | have "rel_pmf (inf (ord_option ord) (ord_option ord)\<inverse>\<inverse>) p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1033 | using local.refl local.trans reflp_ord_option rel_pmf_inf transp_ord_option by blast | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1034 | also have "inf (ord_option ord) (ord_option ord)\<inverse>\<inverse> = rel_option (inf ord ord\<inverse>\<inverse>)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1035 | by(auto simp: fun_eq_iff elim: ord_option.cases option.rel_cases) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1036 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1037 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1038 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1039 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1040 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1041 | lemma ord_spmf_return_spmf2: "ord_spmf R p (return_spmf y) \<longleftrightarrow> (\<forall>x\<in>set_spmf p. R x y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1042 | by(auto simp: rel_pmf_return_pmf2 in_set_spmf ord_option.simps intro: ccontr) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1043 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1044 | lemma ord_spmf_mono: "\<lbrakk> ord_spmf A p q; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> ord_spmf B p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1045 | by(erule pmf.rel_mono_strong)(erule ord_option_mono) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1046 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1047 | lemma ord_spmf_compp: "ord_spmf (A OO B) = ord_spmf A OO ord_spmf B" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1048 | by(simp add: ord_option_compp pmf.rel_compp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1049 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1050 | lemma ord_spmf_bindI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1051 | assumes pq: "ord_spmf R p q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1052 | and fg: "\<And>x y. R x y \<Longrightarrow> ord_spmf P (f x) (g y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1053 | shows "ord_spmf P (p \<bind> f) (q \<bind> g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1054 | unfolding bind_spmf_def using pq | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1055 | by(rule rel_pmf_bindI)(auto split: option.split intro: fg) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1056 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1057 | lemma ord_spmf_bind_reflI: | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1058 | "(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord_spmf R (f x) (g x)) \<Longrightarrow> ord_spmf R (p \<bind> f) (p \<bind> g)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1059 | by(rule ord_spmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_spmf p"])(auto intro: ord_spmf_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1060 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1061 | lemma ord_pmf_increaseI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1062 | assumes le: "\<And>x. spmf p x \<le> spmf q x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1063 | and refl: "\<And>x. x \<in> set_spmf p \<Longrightarrow> R x x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1064 | shows "ord_spmf R p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1065 | proof(rule rel_pmf.intros) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1066 | define pq where "pq = embed_pmf | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1067 | (\<lambda>(x, y). case x of Some x' \<Rightarrow> (case y of Some y' \<Rightarrow> if x' = y' then spmf p x' else 0 | None \<Rightarrow> 0) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1068 | | None \<Rightarrow> (case y of None \<Rightarrow> pmf q None | Some y' \<Rightarrow> spmf q y' - spmf p y'))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1069 | (is "_ = embed_pmf ?f") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1070 | have nonneg: "\<And>xy. ?f xy \<ge> 0" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1071 | by(clarsimp simp add: le field_simps split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1072 | have integral: "(\<integral>\<^sup>+ xy. ?f xy \<partial>count_space UNIV) = 1" (is "nn_integral ?M _ = _") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1073 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1074 | have "(\<integral>\<^sup>+ xy. ?f xy \<partial>count_space UNIV) = | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1075 |       \<integral>\<^sup>+ xy. ennreal (?f xy) * indicator {(None, None)} xy +
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1076 | ennreal (?f xy) * indicator (range (\<lambda>x. (None, Some x))) xy + | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1077 | ennreal (?f xy) * indicator (range (\<lambda>x. (Some x, Some x))) xy \<partial>?M" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1078 | by(rule nn_integral_cong)(auto split: split_indicator option.splits if_split_asm) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1079 |     also have "\<dots> = (\<integral>\<^sup>+ xy. ?f xy * indicator {(None, None)} xy \<partial>?M) +
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1080 | (\<integral>\<^sup>+ xy. ennreal (?f xy) * indicator (range (\<lambda>x. (None, Some x))) xy \<partial>?M) + | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1081 | (\<integral>\<^sup>+ xy. ennreal (?f xy) * indicator (range (\<lambda>x. (Some x, Some x))) xy \<partial>?M)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1082 | (is "_ = ?None + ?Some2 + ?Some") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1083 | by(subst nn_integral_add)(simp_all add: nn_integral_add AE_count_space le_diff_eq le split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1084 | also have "?None = pmf q None" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1085 | also have "?Some2 = \<integral>\<^sup>+ x. ennreal (spmf q x) - spmf p x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1086 | by(simp add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] inj_on_def nn_integral_embed_measure measurable_embed_measure1 ennreal_minus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1087 | also have "\<dots> = (\<integral>\<^sup>+ x. spmf q x \<partial>count_space UNIV) - (\<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1088 | (is "_ = ?Some2' - ?Some2''") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1089 | by(subst nn_integral_diff)(simp_all add: le nn_integral_spmf_neq_top) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1090 | also have "?Some = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1091 | by(simp add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] inj_on_def nn_integral_embed_measure measurable_embed_measure1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1092 | also have "pmf q None + (?Some2' - ?Some2'') + \<dots> = pmf q None + ?Some2'" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1093 | by(auto simp: diff_add_self_ennreal le intro!: nn_integral_mono) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1094 |     also have "\<dots> = \<integral>\<^sup>+ x. ennreal (pmf q x) * indicator {None} x + ennreal (pmf q x) * indicator (range Some) x \<partial>count_space UNIV"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1095 | by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator[symmetric] embed_measure_count_space[symmetric] nn_integral_embed_measure measurable_embed_measure1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1096 | also have "\<dots> = \<integral>\<^sup>+ x. pmf q x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1097 | by(rule nn_integral_cong)(auto split: split_indicator) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1098 | also have "\<dots> = 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1099 | by(simp add: nn_integral_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1100 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1101 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1102 | note f = nonneg integral | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1103 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1104 |   { fix x y
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1105 | assume "(x, y) \<in> set_pmf pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1106 | hence "?f (x, y) \<noteq> 0" unfolding pq_def by(simp add: set_embed_pmf[OF f]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1107 | then show "ord_option R x y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1108 | by(simp add: spmf_eq_0_set_spmf refl split: option.split_asm if_split_asm) } | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1109 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1110 | have weight_le: "weight_spmf p \<le> weight_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1111 | by(subst ennreal_le_iff[symmetric])(auto simp: weight_spmf_eq_nn_integral_spmf intro!: nn_integral_mono le) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1112 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1113 | show "map_pmf fst pq = p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1114 | proof(rule pmf_eqI) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1115 | fix i :: "'a option" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1116 |     have bi: "bij_betw (Pair i) UNIV (fst -` {i})"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1117 | by(auto simp: bij_betw_def inj_on_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1118 | have "ennreal (pmf (map_pmf fst pq) i) = (\<integral>\<^sup>+ y. pmf pq (i, y) \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1119 | unfolding pq_def ennreal_pmf_map | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1120 | apply (simp add: embed_pmf.rep_eq[OF f] o_def emeasure_density flip: nn_integral_count_space_indicator) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1121 | by (smt (verit, best) nn_integral_bij_count_space [OF bi] integral nn_integral_cong nonneg pmf_embed_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1122 | also have "\<dots> = pmf p i" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1123 | proof(cases i) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1124 | case (Some x) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1125 |       have "(\<integral>\<^sup>+ y. pmf pq (Some x, y) \<partial>count_space UNIV) = \<integral>\<^sup>+ y. pmf p (Some x) * indicator {Some x} y \<partial>count_space UNIV"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1126 | by(rule nn_integral_cong)(simp add: pq_def pmf_embed_pmf[OF f] split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1127 | then show ?thesis using Some by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1128 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1129 | case None | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1130 | have "(\<integral>\<^sup>+ y. pmf pq (None, y) \<partial>count_space UNIV) = | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1131 | (\<integral>\<^sup>+ y. ennreal (pmf pq (None, Some (the y))) * indicator (range Some) y + | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1132 |                    ennreal (pmf pq (None, None)) * indicator {None} y \<partial>count_space UNIV)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1133 | by(rule nn_integral_cong)(auto split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1134 | also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (pmf pq (None, Some (the y))) \<partial>count_space (range Some)) + pmf pq (None, None)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1135 | by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1136 | also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (spmf q y) - ennreal (spmf p y) \<partial>count_space UNIV) + pmf q None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1137 | by(simp add: pq_def pmf_embed_pmf[OF f] embed_measure_count_space[symmetric] nn_integral_embed_measure measurable_embed_measure1 ennreal_minus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1138 | also have "(\<integral>\<^sup>+ y. ennreal (spmf q y) - ennreal (spmf p y) \<partial>count_space UNIV) = | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1139 | (\<integral>\<^sup>+ y. spmf q y \<partial>count_space UNIV) - (\<integral>\<^sup>+ y. spmf p y \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1140 | by(subst nn_integral_diff)(simp_all add: AE_count_space le nn_integral_spmf_neq_top split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1141 | also have "\<dots> = pmf p None - pmf q None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1142 | by(simp add: pmf_None_eq_weight_spmf weight_spmf_eq_nn_integral_spmf[symmetric] ennreal_minus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1143 | also have "\<dots> = ennreal (pmf p None) - ennreal (pmf q None)" by(simp add: ennreal_minus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1144 | finally show ?thesis using None weight_le | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1145 | by(auto simp: diff_add_self_ennreal pmf_None_eq_weight_spmf intro: ennreal_leI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1146 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1147 | finally show "pmf (map_pmf fst pq) i = pmf p i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1148 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1149 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1150 | show "map_pmf snd pq = q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1151 | proof(rule pmf_eqI) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1152 | fix i :: "'a option" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1153 |     have bi: "bij_betw (\<lambda>x. (x, i)) UNIV (snd -` {i})"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1154 | by (auto simp: bij_betw_def inj_on_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1155 | have "ennreal (pmf (map_pmf snd pq) i) = (\<integral>\<^sup>+ x. pmf pq (x, i) \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1156 | unfolding pq_def ennreal_pmf_map | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1157 | apply(simp add: embed_pmf.rep_eq[OF f] o_def emeasure_density nn_integral_count_space_indicator[symmetric]) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1158 | by (smt (verit, best) nn_integral_bij_count_space [OF bi] integral nn_integral_cong nonneg pmf_embed_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1159 | also have "\<dots> = ennreal (pmf q i)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1160 | proof(cases i) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1161 | case None | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1162 |       have "(\<integral>\<^sup>+ x. pmf pq (x, None) \<partial>count_space UNIV) = \<integral>\<^sup>+ x. pmf q None * indicator {None :: 'a option} x \<partial>count_space UNIV"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1163 | by(rule nn_integral_cong)(simp add: pq_def pmf_embed_pmf[OF f] split: option.split) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1164 | then show ?thesis using None by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1165 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1166 | case (Some y) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1167 | have "(\<integral>\<^sup>+ x. pmf pq (x, Some y) \<partial>count_space UNIV) = | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1168 | (\<integral>\<^sup>+ x. ennreal (pmf pq (x, Some y)) * indicator (range Some) x + | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1169 |                ennreal (pmf pq (None, Some y)) * indicator {None} x \<partial>count_space UNIV)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1170 | by(rule nn_integral_cong)(auto split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1171 | also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (pmf pq (x, Some y)) * indicator (range Some) x \<partial>count_space UNIV) + pmf pq (None, Some y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1172 | by(subst nn_integral_add)(simp_all) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1173 |       also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (spmf p y) * indicator {Some y} x \<partial>count_space UNIV) + (spmf q y - spmf p y)"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1174 | by(auto simp: pq_def pmf_embed_pmf[OF f] one_ereal_def[symmetric] simp del: nn_integral_indicator_singleton intro!: arg_cong2[where f="(+)"] nn_integral_cong split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1175 | also have "\<dots> = spmf q y" by(simp add: ennreal_minus[symmetric] le) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1176 | finally show ?thesis using Some by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1177 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1178 | finally show "pmf (map_pmf snd pq) i = pmf q i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1179 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1180 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1181 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1182 | lemma ord_spmf_eq_leD: | 
| 67399 | 1183 | assumes "ord_spmf (=) p q" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1184 | shows "spmf p x \<le> spmf q x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1185 | proof(cases "x \<in> set_spmf p") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1186 | case False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1187 | thus ?thesis by(simp add: in_set_spmf_iff_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1188 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1189 | case True | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1190 | from assms obtain pq | 
| 67399 | 1191 | where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> ord_option (=) x y" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1192 | and p: "p = map_pmf fst pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1193 | and q: "q = map_pmf snd pq" by cases auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1194 |   have "ennreal (spmf p x) = integral\<^sup>N pq (indicator (fst -` {Some x}))"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1195 | using p by(simp add: ennreal_pmf_map) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1196 |   also have "\<dots> = integral\<^sup>N pq (indicator {(Some x, Some x)})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1197 | by(rule nn_integral_cong_AE)(auto simp: AE_measure_pmf_iff split: split_indicator dest: pq) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1198 |   also have "\<dots> \<le> integral\<^sup>N pq (indicator (snd -` {Some x}))"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1199 | by(rule nn_integral_mono) simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1200 | also have "\<dots> = ennreal (spmf q x)" using q by(simp add: ennreal_pmf_map) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1201 | finally show ?thesis by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1202 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1203 | |
| 67399 | 1204 | lemma ord_spmf_eqD_set_spmf: "ord_spmf (=) p q \<Longrightarrow> set_spmf p \<subseteq> set_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1205 | by (metis ord_spmf_eq_leD pmf_le_0_iff spmf_eq_0_set_spmf subsetI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1206 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1207 | lemma ord_spmf_eqD_emeasure: | 
| 67399 | 1208 | "ord_spmf (=) p q \<Longrightarrow> emeasure (measure_spmf p) A \<le> emeasure (measure_spmf q) A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1209 | by(auto intro!: nn_integral_mono split: split_indicator dest: ord_spmf_eq_leD simp add: nn_integral_measure_spmf nn_integral_indicator[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1210 | |
| 67399 | 1211 | lemma ord_spmf_eqD_measure_spmf: "ord_spmf (=) p q \<Longrightarrow> measure_spmf p \<le> measure_spmf q" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1212 | by (subst le_measure) (auto simp: ord_spmf_eqD_emeasure) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1213 | |
| 69597 | 1214 | subsection \<open>CCPO structure for the flat ccpo \<^term>\<open>ord_option (=)\<close>\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1215 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1216 | context fixes Y :: "'a spmf set" begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1217 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1218 | definition lub_spmf :: "'a spmf" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1219 | where "lub_spmf = embed_spmf (\<lambda>x. enn2real (SUP p \<in> Y. ennreal (spmf p x)))" | 
| 69597 | 1220 | \<comment> \<open>We go through \<^typ>\<open>ennreal\<close> to have a sensible definition even if \<^term>\<open>Y\<close> is empty.\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1221 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1222 | lemma lub_spmf_empty [simp]: "SPMF.lub_spmf {} = return_pmf None"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1223 | by(simp add: SPMF.lub_spmf_def bot_ereal_def) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1224 | |
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1225 | context assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1226 | begin | 
| 67399 | 1227 | |
| 1228 | lemma chain_ord_spmf_eqD: "Complete_Partial_Order.chain (\<le>) ((\<lambda>p x. ennreal (spmf p x)) ` Y)" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1229 | (is "Complete_Partial_Order.chain _ (?f ` _)") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1230 | proof(rule chainI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1231 | fix f g | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1232 | assume "f \<in> ?f ` Y" "g \<in> ?f ` Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1233 | then obtain p q where f: "f = ?f p" "p \<in> Y" and g: "g = ?f q" "q \<in> Y" by blast | 
| 67399 | 1234 | from chain \<open>p \<in> Y\<close> \<open>q \<in> Y\<close> have "ord_spmf (=) p q \<or> ord_spmf (=) q p" by(rule chainD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1235 | thus "f \<le> g \<or> g \<le> f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1236 | by (metis ennreal_leI f(1) g(1) le_funI ord_spmf_eq_leD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1237 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1238 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1239 | lemma ord_spmf_eq_pmf_None_eq: | 
| 67399 | 1240 | assumes le: "ord_spmf (=) p q" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1241 | and None: "pmf p None = pmf q None" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1242 | shows "p = q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1243 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1244 | fix i | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1245 | from le have le': "\<And>x. spmf p x \<le> spmf q x" by(rule ord_spmf_eq_leD) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1246 | have "(\<integral>\<^sup>+ x. ennreal (spmf q x) - spmf p x \<partial>count_space UNIV) = | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1247 | (\<integral>\<^sup>+ x. spmf q x \<partial>count_space UNIV) - (\<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1248 | by(subst nn_integral_diff)(simp_all add: AE_count_space le' nn_integral_spmf_neq_top) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1249 | also have "\<dots> = (1 - pmf q None) - (1 - pmf p None)" unfolding pmf_None_eq_weight_spmf | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1250 | by(simp add: weight_spmf_eq_nn_integral_spmf[symmetric] ennreal_minus) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1251 | also have "\<dots> = 0" using None by simp | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1252 | finally have "\<And>x. spmf q x \<le> spmf p x" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1253 | by(simp add: nn_integral_0_iff_AE AE_count_space ennreal_minus ennreal_eq_0_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1254 | with le' show "spmf p i = spmf q i" by(rule antisym) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1255 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1256 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1257 | lemma ord_spmf_eqD_pmf_None: | 
| 67399 | 1258 | assumes "ord_spmf (=) x y" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1259 | shows "pmf x None \<ge> pmf y None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1260 | using assms | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1261 | apply cases | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1262 | apply(clarsimp simp only: ennreal_le_iff[symmetric, OF pmf_nonneg] ennreal_pmf_map) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1263 | apply(fastforce simp: AE_measure_pmf_iff intro!: nn_integral_mono_AE) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1264 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1265 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1266 | text \<open> | 
| 69597 | 1267 | Chains on \<^typ>\<open>'a spmf\<close> maintain countable support. | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1268 | Thanks to Johannes Hölzl for the proof idea. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1269 | \<close> | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1270 | lemma spmf_chain_countable: "countable (\<Union>p\<in>Y. set_spmf p)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1271 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1272 | case Y: False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1273 | show ?thesis | 
| 67399 | 1274 | proof(cases "\<exists>x\<in>Y. \<forall>y\<in>Y. ord_spmf (=) y x") | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1275 | case True | 
| 67399 | 1276 | then obtain x where x: "x \<in> Y" and upper: "\<And>y. y \<in> Y \<Longrightarrow> ord_spmf (=) y x" by blast | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1277 | hence "(\<Union>x\<in>Y. set_spmf x) \<subseteq> set_spmf x" by(auto dest: ord_spmf_eqD_set_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1278 | thus ?thesis by(rule countable_subset) simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1279 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1280 | case False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1281 | define N :: "'a option pmf \<Rightarrow> real" where "N p = pmf p None" for p | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1282 | |
| 67399 | 1283 | have N_less_imp_le_spmf: "\<lbrakk> x \<in> Y; y \<in> Y; N y < N x \<rbrakk> \<Longrightarrow> ord_spmf (=) x y" for x y | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1284 | using chainD[OF chain, of x y] ord_spmf_eqD_pmf_None[of x y] ord_spmf_eqD_pmf_None[of y x] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1285 | by (auto simp: N_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1286 | have N_eq_imp_eq: "\<lbrakk> x \<in> Y; y \<in> Y; N y = N x \<rbrakk> \<Longrightarrow> x = y" for x y | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1287 | using chainD[OF chain, of x y] by(auto simp: N_def dest: ord_spmf_eq_pmf_None_eq) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1288 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1289 |     have NC: "N ` Y \<noteq> {}" "bdd_below (N ` Y)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1290 |       using \<open>Y \<noteq> {}\<close> by(auto intro!: bdd_belowI[of _ 0] simp: N_def)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1291 | have NC_less: "Inf (N ` Y) < N x" if "x \<in> Y" for x unfolding cInf_less_iff[OF NC] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1292 | proof(rule ccontr) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1293 | assume **: "\<not> (\<exists>y\<in>N ` Y. y < N x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1294 |       { fix y
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1295 | assume "y \<in> Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1296 | with ** consider "N x < N y" | "N x = N y" by(auto simp: not_less le_less) | 
| 67399 | 1297 | hence "ord_spmf (=) y x" using \<open>y \<in> Y\<close> \<open>x \<in> Y\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1298 | by cases(auto dest: N_less_imp_le_spmf N_eq_imp_eq intro: ord_spmf_reflI) } | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1299 | with False \<open>x \<in> Y\<close> show False by blast | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1300 | qed | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1301 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1302 | from NC have "Inf (N ` Y) \<in> closure (N ` Y)" by (intro closure_contains_Inf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1303 | then obtain X' where "\<And>n. X' n \<in> N ` Y" and X': "X' \<longlonglongrightarrow> Inf (N ` Y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1304 | unfolding closure_sequential by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1305 | then obtain X where X: "\<And>n. X n \<in> Y" and "X' = (\<lambda>n. N (X n))" unfolding image_iff Bex_def by metis | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1306 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1307 | with X' have seq: "(\<lambda>n. N (X n)) \<longlonglongrightarrow> Inf (N ` Y)" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1308 | have "(\<Union>x \<in> Y. set_spmf x) \<subseteq> (\<Union>n. set_spmf (X n))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1309 | proof(rule UN_least) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1310 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1311 | assume "x \<in> Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1312 | from order_tendstoD(2)[OF seq NC_less[OF \<open>x \<in> Y\<close>]] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1313 | obtain i where "N (X i) < N x" by (auto simp: eventually_sequentially) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1314 | thus "set_spmf x \<subseteq> (\<Union>n. set_spmf (X n))" using X \<open>x \<in> Y\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1315 | by(blast dest: N_less_imp_le_spmf ord_spmf_eqD_set_spmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1316 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1317 | thus ?thesis by(rule countable_subset) simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1318 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1319 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1320 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1321 | lemma lub_spmf_subprob: "(\<integral>\<^sup>+ x. (SUP p \<in> Y. ennreal (spmf p x)) \<partial>count_space UNIV) \<le> 1" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1322 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1323 | case True | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1324 | thus ?thesis by(simp add: bot_ennreal) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1325 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1326 | case False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1327 | let ?B = "\<Union>p\<in>Y. set_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1328 | have countable: "countable ?B" by(rule spmf_chain_countable) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1329 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1330 | have "(\<integral>\<^sup>+ x. (SUP p\<in>Y. ennreal (spmf p x)) \<partial>count_space UNIV) = | 
| 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1331 | (\<integral>\<^sup>+ x. (SUP p\<in>Y. ennreal (spmf p x) * indicator ?B x) \<partial>count_space UNIV)" | 
| 69661 | 1332 | by (intro nn_integral_cong arg_cong [of _ _ Sup]) (auto split: split_indicator simp add: spmf_eq_0_set_spmf) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1333 | also have "\<dots> = (\<integral>\<^sup>+ x. (SUP p\<in>Y. ennreal (spmf p x)) \<partial>count_space ?B)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1334 | unfolding ennreal_indicator[symmetric] using False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1335 | by(subst SUP_mult_right_ennreal[symmetric])(simp add: ennreal_indicator nn_integral_count_space_indicator) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1336 | also have "\<dots> = (SUP p\<in>Y. \<integral>\<^sup>+ x. spmf p x \<partial>count_space ?B)" using False _ countable | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1337 | by(rule nn_integral_monotone_convergence_SUP_countable)(rule chain_ord_spmf_eqD) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1338 | also have "\<dots> \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1339 | proof(rule SUP_least) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1340 | fix p | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1341 | assume "p \<in> Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1342 | have "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space ?B) = \<integral>\<^sup>+ x. ennreal (spmf p x) * indicator ?B x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1343 | by(simp add: nn_integral_count_space_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1344 | also have "\<dots> = \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1345 | by(rule nn_integral_cong)(auto split: split_indicator simp add: spmf_eq_0_set_spmf \<open>p \<in> Y\<close>) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1346 | also have "\<dots> \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1347 | by(simp add: weight_spmf_eq_nn_integral_spmf[symmetric] weight_spmf_le_1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1348 | finally show "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space ?B) \<le> 1" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1349 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1350 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1351 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1352 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1353 | lemma spmf_lub_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1354 |   assumes "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1355 | shows "spmf lub_spmf x = (SUP p \<in> Y. spmf p x)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1356 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1357 | from assms obtain p where "p \<in> Y" by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1358 | have "spmf lub_spmf x = max 0 (enn2real (SUP p\<in>Y. ennreal (spmf p x)))" unfolding lub_spmf_def | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1359 | by(rule spmf_embed_spmf)(simp del: SUP_eq_top_iff Sup_eq_top_iff add: ennreal_enn2real_if SUP_spmf_neq_top' lub_spmf_subprob) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1360 | also have "\<dots> = enn2real (SUP p\<in>Y. ennreal (spmf p x))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1361 | by(rule max_absorb2)(simp) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1362 | also have "\<dots> = enn2real (ennreal (SUP p \<in> Y. spmf p x))" using assms | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1363 | by(subst ennreal_SUP[symmetric])(simp_all add: SUP_spmf_neq_top' del: SUP_eq_top_iff Sup_eq_top_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1364 | also have "0 \<le> (\<Squnion>p\<in>Y. spmf p x)" using assms | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1365 | by(auto intro!: cSUP_upper2 bdd_aboveI[where M=1] simp add: pmf_le_1) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1366 | then have "enn2real (ennreal (SUP p \<in> Y. spmf p x)) = (SUP p \<in> Y. spmf p x)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1367 | by(rule enn2real_ennreal) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1368 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1369 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1370 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1371 | lemma ennreal_spmf_lub_spmf: "Y \<noteq> {} \<Longrightarrow> ennreal (spmf lub_spmf x) = (SUP p\<in>Y. ennreal (spmf p x))"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1372 | by (metis SUP_spmf_neq_top' ennreal_SUP spmf_lub_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1373 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1374 | lemma lub_spmf_upper: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1375 | assumes p: "p \<in> Y" | 
| 67399 | 1376 | shows "ord_spmf (=) p lub_spmf" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1377 | proof(rule ord_pmf_increaseI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1378 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1379 |   from p have [simp]: "Y \<noteq> {}" by auto
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1380 | from p have "ennreal (spmf p x) \<le> (SUP p\<in>Y. ennreal (spmf p x))" by(rule SUP_upper) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1381 | also have "\<dots> = ennreal (spmf lub_spmf x)" using p | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1382 | by(subst spmf_lub_spmf)(auto simp: ennreal_SUP SUP_spmf_neq_top' simp del: SUP_eq_top_iff Sup_eq_top_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1383 | finally show "spmf p x \<le> spmf lub_spmf x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1384 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1385 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1386 | lemma lub_spmf_least: | 
| 67399 | 1387 | assumes z: "\<And>x. x \<in> Y \<Longrightarrow> ord_spmf (=) x z" | 
| 1388 | shows "ord_spmf (=) lub_spmf z" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1389 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1390 | case nonempty: False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1391 | show ?thesis | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1392 | proof(rule ord_pmf_increaseI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1393 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1394 | from nonempty obtain p where p: "p \<in> Y" by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1395 | have "ennreal (spmf lub_spmf x) = (SUP p\<in>Y. ennreal (spmf p x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1396 | by(subst spmf_lub_spmf)(auto simp: ennreal_SUP SUP_spmf_neq_top' nonempty simp del: SUP_eq_top_iff Sup_eq_top_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1397 | also have "\<dots> \<le> ennreal (spmf z x)" by(rule SUP_least)(simp add: ord_spmf_eq_leD z) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1398 | finally show "spmf lub_spmf x \<le> spmf z x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1399 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1400 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1401 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1402 | lemma set_lub_spmf: "set_spmf lub_spmf = (\<Union>p\<in>Y. set_spmf p)" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1403 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1404 | case [simp]: False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1405 | show ?thesis | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1406 | proof(rule set_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1407 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1408 | have "x \<in> ?lhs \<longleftrightarrow> ennreal (spmf lub_spmf x) > 0" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1409 | by(simp_all add: in_set_spmf_iff_spmf less_le) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1410 | also have "\<dots> \<longleftrightarrow> (\<exists>p\<in>Y. ennreal (spmf p x) > 0)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1411 | by(simp add: ennreal_spmf_lub_spmf less_SUP_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1412 | also have "\<dots> \<longleftrightarrow> x \<in> ?rhs" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1413 | by(auto simp: in_set_spmf_iff_spmf less_le) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1414 | finally show "x \<in> ?lhs \<longleftrightarrow> x \<in> ?rhs" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1415 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1416 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1417 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1418 | lemma emeasure_lub_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1419 |   assumes Y: "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1420 | shows "emeasure (measure_spmf lub_spmf) A = (SUP y\<in>Y. emeasure (measure_spmf y) A)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1421 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1422 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1423 | let ?M = "count_space (set_spmf lub_spmf)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1424 | have "?lhs = \<integral>\<^sup>+ x. ennreal (spmf lub_spmf x) * indicator A x \<partial>?M" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1425 | by(auto simp: nn_integral_indicator[symmetric] nn_integral_measure_spmf') | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1426 | also have "\<dots> = \<integral>\<^sup>+ x. (SUP y\<in>Y. ennreal (spmf y x) * indicator A x) \<partial>?M" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1427 | unfolding ennreal_indicator[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1428 | by(simp add: spmf_lub_spmf assms ennreal_SUP[OF SUP_spmf_neq_top'] SUP_mult_right_ennreal) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1429 | also from assms have "\<dots> = (SUP y\<in>Y. \<integral>\<^sup>+ x. ennreal (spmf y x) * indicator A x \<partial>?M)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1430 | proof(rule nn_integral_monotone_convergence_SUP_countable) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1431 | have "(\<lambda>i x. ennreal (spmf i x) * indicator A x) ` Y = (\<lambda>f x. f x * indicator A x) ` (\<lambda>p x. ennreal (spmf p x)) ` Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1432 | by(simp add: image_image) | 
| 67399 | 1433 | also have "Complete_Partial_Order.chain (\<le>) \<dots>" using chain_ord_spmf_eqD | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1434 | by(rule chain_imageI)(auto simp: le_fun_def split: split_indicator) | 
| 67399 | 1435 | finally show "Complete_Partial_Order.chain (\<le>) ((\<lambda>i x. ennreal (spmf i x) * indicator A x) ` Y)" . | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1436 | qed simp | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1437 | also have "\<dots> = (SUP y\<in>Y. \<integral>\<^sup>+ x. ennreal (spmf y x) * indicator A x \<partial>count_space UNIV)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1438 | by(auto simp: nn_integral_count_space_indicator set_lub_spmf spmf_eq_0_set_spmf split: split_indicator intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1439 | also have "\<dots> = ?rhs" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1440 | by(auto simp: nn_integral_indicator[symmetric] nn_integral_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1441 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1442 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1443 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1444 | lemma measure_lub_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1445 |   assumes Y: "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1446 | shows "measure (measure_spmf lub_spmf) A = (SUP y\<in>Y. measure (measure_spmf y) A)" (is "?lhs = ?rhs") | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1447 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1448 | have "ennreal ?lhs = ennreal ?rhs" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1449 | using emeasure_lub_spmf[OF assms] SUP_emeasure_spmf_neq_top[of A Y] Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1450 | unfolding measure_spmf.emeasure_eq_measure by(subst ennreal_SUP) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1451 | moreover have "0 \<le> ?rhs" using Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1452 | by(auto intro!: cSUP_upper2 bdd_aboveI[where M=1] measure_spmf.subprob_measure_le_1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1453 | ultimately show ?thesis by(simp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1454 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1455 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1456 | lemma weight_lub_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1457 |   assumes Y: "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1458 | shows "weight_spmf lub_spmf = (SUP y\<in>Y. weight_spmf y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1459 | by (smt (verit, best) SUP_cong assms measure_lub_spmf space_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1460 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1461 | lemma measure_spmf_lub_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1462 |   assumes Y: "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1463 | shows "measure_spmf lub_spmf = (SUP p\<in>Y. measure_spmf p)" (is "?lhs = ?rhs") | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1464 | proof(rule measure_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1465 | from assms obtain p where p: "p \<in> Y" by auto | 
| 67399 | 1466 | from chain have chain': "Complete_Partial_Order.chain (\<le>) (measure_spmf ` Y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1467 | by(rule chain_imageI)(rule ord_spmf_eqD_measure_spmf) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1468 | show "sets ?lhs = sets ?rhs" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1469 | using Y by (subst sets_SUP) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1470 | show "emeasure ?lhs A = emeasure ?rhs A" for A | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1471 | using chain' Y p by (subst emeasure_SUP_chain) (auto simp: emeasure_lub_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1472 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1473 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1474 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1475 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1476 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1477 | |
| 67399 | 1478 | lemma partial_function_definitions_spmf: "partial_function_definitions (ord_spmf (=)) lub_spmf" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1479 | (is "partial_function_definitions ?R _") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1480 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1481 | fix x show "?R x x" by(simp add: ord_spmf_reflI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1482 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1483 | fix x y z | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1484 | assume "?R x y" "?R y z" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1485 | with transp_ord_option[OF transp_equality] show "?R x z" by(rule transp_rel_pmf[THEN transpD]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1486 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1487 | fix x y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1488 | assume "?R x y" "?R y x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1489 | thus "x = y" | 
| 64634 | 1490 | by(rule rel_pmf_antisym)(simp_all add: reflp_ord_option transp_ord_option antisymp_ord_option) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1491 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1492 | fix Y x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1493 | assume "Complete_Partial_Order.chain ?R Y" "x \<in> Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1494 | then show "?R x (lub_spmf Y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1495 | by(rule lub_spmf_upper) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1496 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1497 | fix Y z | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1498 | assume "Complete_Partial_Order.chain ?R Y" "\<And>x. x \<in> Y \<Longrightarrow> ?R x z" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1499 | then show "?R (lub_spmf Y) z" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1500 |     by(cases "Y = {}")(simp_all add: lub_spmf_least)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1501 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1502 | |
| 67399 | 1503 | lemma ccpo_spmf: "class.ccpo lub_spmf (ord_spmf (=)) (mk_less (ord_spmf (=)))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1504 | by(metis ccpo partial_function_definitions_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1505 | |
| 67399 | 1506 | interpretation spmf: partial_function_definitions "ord_spmf (=)" "lub_spmf" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1507 |   rewrites "lub_spmf {} \<equiv> return_pmf None"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1508 | by(rule partial_function_definitions_spmf) simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1509 | |
| 69597 | 1510 | declaration \<open>Partial_Function.init "spmf" \<^term>\<open>spmf.fixp_fun\<close> | 
| 1511 |   \<^term>\<open>spmf.mono_body\<close> @{thm spmf.fixp_rule_uc} @{thm spmf.fixp_induct_uc}
 | |
| 63308 | 1512 | NONE\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1513 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1514 | declare spmf.leq_refl[simp] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1515 | declare admissible_leI[OF ccpo_spmf, cont_intro] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1516 | |
| 67399 | 1517 | abbreviation "mono_spmf \<equiv> monotone (fun_ord (ord_spmf (=))) (ord_spmf (=))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1518 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1519 | lemma lub_spmf_const [simp]: "lub_spmf {p} = p"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1520 | by(rule spmf_eqI)(simp add: spmf_lub_spmf[OF ccpo.chain_singleton[OF ccpo_spmf]]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1521 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1522 | lemma bind_spmf_mono': | 
| 67399 | 1523 | assumes fg: "ord_spmf (=) f g" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1524 | and hk: "\<And>x :: 'a. ord_spmf (=) (h x) (k x)" | 
| 67399 | 1525 | shows "ord_spmf (=) (f \<bind> h) (g \<bind> k)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1526 | unfolding bind_spmf_def using assms(1) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1527 | by(rule rel_pmf_bindI)(auto split: option.split simp add: hk) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1528 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1529 | lemma bind_spmf_mono [partial_function_mono]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1530 | assumes mf: "mono_spmf B" and mg: "\<And>y. mono_spmf (\<lambda>f. C y f)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1531 | shows "mono_spmf (\<lambda>f. bind_spmf (B f) (\<lambda>y. C y f))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1532 | proof (rule monotoneI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1533 | fix f g :: "'a \<Rightarrow> 'b spmf" | 
| 67399 | 1534 | assume fg: "fun_ord (ord_spmf (=)) f g" | 
| 1535 | with mf have "ord_spmf (=) (B f) (B g)" by (rule monotoneD[of _ _ _ f g]) | |
| 1536 | moreover from mg have "\<And>y'. ord_spmf (=) (C y' f) (C y' g)" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1537 | by (rule monotoneD) (rule fg) | 
| 67399 | 1538 | ultimately show "ord_spmf (=) (bind_spmf (B f) (\<lambda>y. C y f)) (bind_spmf (B g) (\<lambda>y'. C y' g))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1539 | by(rule bind_spmf_mono') | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1540 | qed | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1541 | |
| 67399 | 1542 | lemma monotone_bind_spmf1: "monotone (ord_spmf (=)) (ord_spmf (=)) (\<lambda>y. bind_spmf y g)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1543 | by(rule monotoneI)(simp add: bind_spmf_mono' ord_spmf_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1544 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1545 | lemma monotone_bind_spmf2: | 
| 67399 | 1546 | assumes g: "\<And>x. monotone ord (ord_spmf (=)) (\<lambda>y. g y x)" | 
| 1547 | shows "monotone ord (ord_spmf (=)) (\<lambda>y. bind_spmf p (g y))" | |
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1548 | by(rule monotoneI)(auto intro: bind_spmf_mono' monotoneD[OF g] ord_spmf_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1549 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1550 | lemma bind_lub_spmf: | 
| 67399 | 1551 | assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1552 | shows "bind_spmf (lub_spmf Y) f = lub_spmf ((\<lambda>p. bind_spmf p f) ` Y)" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1553 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1554 | case Y: False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1555 | show ?thesis | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1556 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1557 | fix i | 
| 67399 | 1558 | have chain': "Complete_Partial_Order.chain (\<le>) ((\<lambda>p x. ennreal (spmf p x * spmf (f x) i)) ` Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1559 | using chain by(rule chain_imageI)(auto simp: le_fun_def dest: ord_spmf_eq_leD intro: mult_right_mono) | 
| 67399 | 1560 | have chain'': "Complete_Partial_Order.chain (ord_spmf (=)) ((\<lambda>p. p \<bind> f) ` Y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1561 | using chain by(rule chain_imageI)(auto intro!: monotoneI bind_spmf_mono' ord_spmf_reflI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1562 | let ?M = "count_space (set_spmf (lub_spmf Y))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1563 | have "ennreal (spmf ?lhs i) = \<integral>\<^sup>+ x. ennreal (spmf (lub_spmf Y) x) * ennreal (spmf (f x) i) \<partial>?M" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1564 | by(auto simp: ennreal_spmf_lub_spmf ennreal_spmf_bind nn_integral_measure_spmf') | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1565 | also have "\<dots> = \<integral>\<^sup>+ x. (SUP p\<in>Y. ennreal (spmf p x * spmf (f x) i)) \<partial>?M" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1566 | by(subst ennreal_spmf_lub_spmf[OF chain Y])(subst SUP_mult_right_ennreal, simp_all add: ennreal_mult Y) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1567 | also have "\<dots> = (SUP p\<in>Y. \<integral>\<^sup>+ x. ennreal (spmf p x * spmf (f x) i) \<partial>?M)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1568 | using Y chain' by(rule nn_integral_monotone_convergence_SUP_countable) simp | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1569 | also have "\<dots> = (SUP p\<in>Y. ennreal (spmf (bind_spmf p f) i))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1570 | by(auto simp: ennreal_spmf_bind nn_integral_measure_spmf nn_integral_count_space_indicator set_lub_spmf[OF chain] in_set_spmf_iff_spmf ennreal_mult intro!: arg_cong [of _ _ Sup] image_cong nn_integral_cong split: split_indicator) | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 1571 | also have "\<dots> = ennreal (spmf ?rhs i)" using chain'' by(simp add: ennreal_spmf_lub_spmf Y image_comp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1572 | finally show "spmf ?lhs i = spmf ?rhs i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1573 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1574 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1575 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1576 | lemma map_lub_spmf: | 
| 67399 | 1577 | "Complete_Partial_Order.chain (ord_spmf (=)) Y | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1578 | \<Longrightarrow> map_spmf f (lub_spmf Y) = lub_spmf (map_spmf f ` Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1579 | unfolding map_spmf_conv_bind_spmf[abs_def] by(simp add: bind_lub_spmf o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1580 | |
| 67399 | 1581 | lemma mcont_bind_spmf1: "mcont lub_spmf (ord_spmf (=)) lub_spmf (ord_spmf (=)) (\<lambda>y. bind_spmf y f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1582 | using monotone_bind_spmf1 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1583 | by(intro contI mcontI) (auto simp: bind_lub_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1584 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1585 | lemma bind_lub_spmf2: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1586 | assumes chain: "Complete_Partial_Order.chain ord Y" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1587 | and g: "\<And>y. monotone ord (ord_spmf (=)) (g y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1588 | shows "bind_spmf x (\<lambda>y. lub_spmf (g y ` Y)) = lub_spmf ((\<lambda>p. bind_spmf x (\<lambda>y. g y p)) ` Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1589 | (is "?lhs = ?rhs") | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1590 | proof(cases "Y = {}")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1591 | case Y: False | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1592 | show ?thesis | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1593 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1594 | fix i | 
| 67399 | 1595 | have chain': "\<And>y. Complete_Partial_Order.chain (ord_spmf (=)) (g y ` Y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1596 | using chain g[THEN monotoneD] by(rule chain_imageI) | 
| 67399 | 1597 | have chain'': "Complete_Partial_Order.chain (\<le>) ((\<lambda>p y. ennreal (spmf x y * spmf (g y p) i)) ` Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1598 | using chain by(rule chain_imageI)(auto simp: le_fun_def dest: ord_spmf_eq_leD monotoneD[OF g] intro!: mult_left_mono) | 
| 67399 | 1599 | have chain''': "Complete_Partial_Order.chain (ord_spmf (=)) ((\<lambda>p. bind_spmf x (\<lambda>y. g y p)) ` Y)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1600 | using chain by(rule chain_imageI)(rule monotone_bind_spmf2[OF g, THEN monotoneD]) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1601 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1602 | have "ennreal (spmf ?lhs i) = \<integral>\<^sup>+ y. (SUP p\<in>Y. ennreal (spmf x y * spmf (g y p) i)) \<partial>count_space (set_spmf x)" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 1603 | by(simp add: ennreal_spmf_bind ennreal_spmf_lub_spmf[OF chain'] Y nn_integral_measure_spmf' SUP_mult_left_ennreal ennreal_mult image_comp) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1604 | also have "\<dots> = (SUP p\<in>Y. \<integral>\<^sup>+ y. ennreal (spmf x y * spmf (g y p) i) \<partial>count_space (set_spmf x))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1605 | unfolding nn_integral_measure_spmf' using Y chain'' | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1606 | by(rule nn_integral_monotone_convergence_SUP_countable) simp | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1607 | also have "\<dots> = (SUP p\<in>Y. ennreal (spmf (bind_spmf x (\<lambda>y. g y p)) i))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1608 | by(simp add: ennreal_spmf_bind nn_integral_measure_spmf' ennreal_mult) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1609 | also have "\<dots> = ennreal (spmf ?rhs i)" using chain''' | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1610 | by(auto simp: ennreal_spmf_lub_spmf Y image_comp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1611 | finally show "spmf ?lhs i = spmf ?rhs i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1612 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1613 | qed simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1614 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1615 | lemma mcont_bind_spmf [cont_intro]: | 
| 67399 | 1616 | assumes f: "mcont luba orda lub_spmf (ord_spmf (=)) f" | 
| 1617 | and g: "\<And>y. mcont luba orda lub_spmf (ord_spmf (=)) (g y)" | |
| 1618 | shows "mcont luba orda lub_spmf (ord_spmf (=)) (\<lambda>x. bind_spmf (f x) (\<lambda>y. g y x))" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1619 | proof(rule spmf.mcont2mcont'[OF _ _ f]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1620 | fix z | 
| 67399 | 1621 | show "mcont lub_spmf (ord_spmf (=)) lub_spmf (ord_spmf (=)) (\<lambda>x. bind_spmf x (\<lambda>y. g y z))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1622 | by(rule mcont_bind_spmf1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1623 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1624 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1625 | let ?f = "\<lambda>z. bind_spmf x (\<lambda>y. g y z)" | 
| 67399 | 1626 | have "monotone orda (ord_spmf (=)) ?f" using mcont_mono[OF g] by(rule monotone_bind_spmf2) | 
| 1627 | moreover have "cont luba orda lub_spmf (ord_spmf (=)) ?f" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1628 | proof(rule contI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1629 | fix Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1630 |     assume chain: "Complete_Partial_Order.chain orda Y" and Y: "Y \<noteq> {}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1631 | have "bind_spmf x (\<lambda>y. g y (luba Y)) = bind_spmf x (\<lambda>y. lub_spmf (g y ` Y))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1632 | by(rule bind_spmf_cong)(simp_all add: mcont_contD[OF g chain Y]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1633 | also have "\<dots> = lub_spmf ((\<lambda>p. x \<bind> (\<lambda>y. g y p)) ` Y)" using chain | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1634 | by(rule bind_lub_spmf2)(rule mcont_mono[OF g]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1635 | finally show "bind_spmf x (\<lambda>y. g y (luba Y)) = \<dots>" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1636 | qed | 
| 67399 | 1637 | ultimately show "mcont luba orda lub_spmf (ord_spmf (=)) ?f" by(rule mcontI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1638 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1639 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1640 | lemma bind_pmf_mono [partial_function_mono]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1641 | "(\<And>y. mono_spmf (\<lambda>f. C y f)) \<Longrightarrow> mono_spmf (\<lambda>f. bind_pmf p (\<lambda>x. C x f))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1642 | using bind_spmf_mono[of "\<lambda>_. spmf_of_pmf p" C] by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1643 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1644 | lemma map_spmf_mono [partial_function_mono]: "mono_spmf B \<Longrightarrow> mono_spmf (\<lambda>g. map_spmf f (B g))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1645 | unfolding map_spmf_conv_bind_spmf by(rule bind_spmf_mono) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1646 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1647 | lemma mcont_map_spmf [cont_intro]: | 
| 67399 | 1648 | "mcont luba orda lub_spmf (ord_spmf (=)) g | 
| 1649 | \<Longrightarrow> mcont luba orda lub_spmf (ord_spmf (=)) (\<lambda>x. map_spmf f (g x))" | |
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1650 | unfolding map_spmf_conv_bind_spmf by(rule mcont_bind_spmf) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1651 | |
| 67399 | 1652 | lemma monotone_set_spmf: "monotone (ord_spmf (=)) (\<subseteq>) set_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1653 | by(rule monotoneI)(rule ord_spmf_eqD_set_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1654 | |
| 67399 | 1655 | lemma cont_set_spmf: "cont lub_spmf (ord_spmf (=)) Union (\<subseteq>) set_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1656 | by(rule contI)(subst set_lub_spmf; simp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1657 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1658 | lemma mcont2mcont_set_spmf[THEN mcont2mcont, cont_intro]: | 
| 67399 | 1659 | shows mcont_set_spmf: "mcont lub_spmf (ord_spmf (=)) Union (\<subseteq>) set_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1660 | by(rule mcontI monotone_set_spmf cont_set_spmf)+ | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1661 | |
| 67399 | 1662 | lemma monotone_spmf: "monotone (ord_spmf (=)) (\<le>) (\<lambda>p. spmf p x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1663 | by(rule monotoneI)(simp add: ord_spmf_eq_leD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1664 | |
| 67399 | 1665 | lemma cont_spmf: "cont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. spmf p x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1666 | by(rule contI)(simp add: spmf_lub_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1667 | |
| 67399 | 1668 | lemma mcont_spmf: "mcont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. spmf p x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1669 | by(metis mcontI monotone_spmf cont_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1670 | |
| 67399 | 1671 | lemma cont_ennreal_spmf: "cont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. ennreal (spmf p x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1672 | by(rule contI)(simp add: ennreal_spmf_lub_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1673 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1674 | lemma mcont2mcont_ennreal_spmf [THEN mcont2mcont, cont_intro]: | 
| 67399 | 1675 | shows mcont_ennreal_spmf: "mcont lub_spmf (ord_spmf (=)) Sup (\<le>) (\<lambda>p. ennreal (spmf p x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1676 | by(metis mcontI mono2mono_ennreal monotone_spmf cont_ennreal_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1677 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1678 | lemma nn_integral_map_spmf [simp]: "nn_integral (measure_spmf (map_spmf f p)) g = nn_integral (measure_spmf p) (g \<circ> f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1679 | by(force simp: measure_spmf_def nn_integral_distr nn_integral_restrict_space intro: nn_integral_cong split: split_indicator) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1680 | |
| 69597 | 1681 | subsubsection \<open>Admissibility of \<^term>\<open>rel_spmf\<close>\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1682 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1683 | lemma rel_spmf_measureD: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1684 | assumes "rel_spmf R p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1685 |   shows "measure (measure_spmf p) A \<le> measure (measure_spmf q) {y. \<exists>x\<in>A. R x y}" (is "?lhs \<le> ?rhs")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1686 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1687 | have "?lhs = measure (measure_pmf p) (Some ` A)" by(simp add: measure_measure_spmf_conv_measure_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1688 |   also have "\<dots> \<le> measure (measure_pmf q) {y. \<exists>x\<in>Some ` A. rel_option R x y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1689 | using assms by(rule rel_pmf_measureD) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1690 | also have "\<dots> = ?rhs" unfolding measure_measure_spmf_conv_measure_pmf | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1691 | by(rule arg_cong2[where f=measure])(auto simp: option_rel_Some1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1692 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1693 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1694 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1695 | locale rel_spmf_characterisation = | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1696 | assumes rel_pmf_measureI: | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1697 | "\<And>(R :: 'a option \<Rightarrow> 'b option \<Rightarrow> bool) p q. | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1698 |     (\<And>A. measure (measure_pmf p) A \<le> measure (measure_pmf q) {y. \<exists>x\<in>A. R x y})
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1699 | \<Longrightarrow> rel_pmf R p q" | 
| 63308 | 1700 | \<comment> \<open>This assumption is shown to hold in general in the AFP entry \<open>MFMC_Countable\<close>.\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1701 | begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1702 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1703 | context fixes R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" begin | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1704 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1705 | lemma rel_spmf_measureI: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1706 |   assumes eq1: "\<And>A. measure (measure_spmf p) A \<le> measure (measure_spmf q) {y. \<exists>x\<in>A. R x y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1707 | assumes eq2: "weight_spmf q \<le> weight_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1708 | shows "rel_spmf R p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1709 | proof(rule rel_pmf_measureI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1710 | fix A :: "'a option set" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1711 | define A' where "A' = the ` (A \<inter> range Some)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1712 |   define A'' where "A'' = A \<inter> {None}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1713 |   have A: "A = Some ` A' \<union> A''" "Some ` A' \<inter> A'' = {}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1714 | unfolding A'_def A''_def by(auto simp: image_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1715 | have "measure (measure_pmf p) A = measure (measure_pmf p) (Some ` A') + measure (measure_pmf p) A''" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1716 | by(simp add: A measure_pmf.finite_measure_Union) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1717 | also have "measure (measure_pmf p) (Some ` A') = measure (measure_spmf p) A'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1718 | by(simp add: measure_measure_spmf_conv_measure_pmf) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1719 |   also have "\<dots> \<le> measure (measure_spmf q) {y. \<exists>x\<in>A'. R x y}" by(rule eq1)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1720 | also (ord_eq_le_trans[OF _ add_right_mono]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1721 |   have "\<dots> = measure (measure_pmf q) {y. \<exists>x\<in>A'. rel_option R (Some x) y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1722 | unfolding measure_measure_spmf_conv_measure_pmf | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1723 | by(rule arg_cong2[where f=measure])(auto simp: A'_def option_rel_Some1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1724 | also | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1725 |   { have "weight_spmf p \<le> measure (measure_spmf q) {y. \<exists>x. R x y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1726 | using eq1[of UNIV] unfolding weight_spmf_def by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1727 | also have "\<dots> \<le> weight_spmf q" unfolding weight_spmf_def | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1728 | by(rule measure_spmf.finite_measure_mono) simp_all | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1729 | finally have "weight_spmf p = weight_spmf q" using eq2 by simp } | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1730 |   then have "measure (measure_pmf p) A'' = measure (measure_pmf q) (if None \<in> A then {None} else {})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1731 | unfolding A''_def by(simp add: pmf_None_eq_weight_spmf measure_pmf_single) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1732 |   also have "measure (measure_pmf q) {y. \<exists>x\<in>A'. rel_option R (Some x) y} + \<dots> = measure (measure_pmf q) {y. \<exists>x\<in>A. rel_option R x y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1733 | by(subst measure_pmf.finite_measure_Union[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1734 | (auto 4 3 intro!: arg_cong2[where f=measure] simp add: option_rel_Some1 option_rel_Some2 A'_def intro: rev_bexI elim: option.rel_cases) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1735 | finally show "measure (measure_pmf p) A \<le> \<dots>" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1736 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1737 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1738 | lemma admissible_rel_spmf: | 
| 67399 | 1739 | "ccpo.admissible (prod_lub lub_spmf lub_spmf) (rel_prod (ord_spmf (=)) (ord_spmf (=))) (case_prod (rel_spmf R))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1740 | (is "ccpo.admissible ?lub ?ord ?P") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1741 | proof(rule ccpo.admissibleI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1742 | fix Y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1743 | assume chain: "Complete_Partial_Order.chain ?ord Y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1744 |     and Y: "Y \<noteq> {}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1745 | and R: "\<forall>(p, q) \<in> Y. rel_spmf R p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1746 | from R have R: "\<And>p q. (p, q) \<in> Y \<Longrightarrow> rel_spmf R p q" by auto | 
| 67399 | 1747 | have chain1: "Complete_Partial_Order.chain (ord_spmf (=)) (fst ` Y)" | 
| 1748 | and chain2: "Complete_Partial_Order.chain (ord_spmf (=)) (snd ` Y)" | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1749 | using chain by(rule chain_imageI; clarsimp)+ | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1750 |   from Y have Y1: "fst ` Y \<noteq> {}" and Y2: "snd ` Y \<noteq> {}" by auto
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1751 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1752 | have "rel_spmf R (lub_spmf (fst ` Y)) (lub_spmf (snd ` Y))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1753 | proof(rule rel_spmf_measureI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1754 | show "weight_spmf (lub_spmf (snd ` Y)) \<le> weight_spmf (lub_spmf (fst ` Y))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1755 | by(auto simp: weight_lub_spmf chain1 chain2 Y rel_spmf_weightD[OF R, symmetric] intro!: cSUP_least intro: cSUP_upper2[OF bdd_aboveI2[OF weight_spmf_le_1]]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1756 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1757 | fix A | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1758 | have "measure (measure_spmf (lub_spmf (fst ` Y))) A = (SUP y\<in>fst ` Y. measure (measure_spmf y) A)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1759 | using chain1 Y1 by(rule measure_lub_spmf) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 1760 |     also have "\<dots> \<le> (SUP y\<in>snd ` Y. measure (measure_spmf y) {y. \<exists>x\<in>A. R x y})" using Y1
 | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1761 | by(rule cSUP_least)(auto intro!: cSUP_upper2[OF bdd_aboveI2[OF measure_spmf.subprob_measure_le_1]] rel_spmf_measureD R) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1762 |     also have "\<dots> = measure (measure_spmf (lub_spmf (snd ` Y))) {y. \<exists>x\<in>A. R x y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1763 | using chain2 Y2 by(rule measure_lub_spmf[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1764 | finally show "measure (measure_spmf (lub_spmf (fst ` Y))) A \<le> \<dots>" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1765 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1766 | then show "?P (?lub Y)" by(simp add: prod_lub_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1767 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1768 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1769 | lemma admissible_rel_spmf_mcont [cont_intro]: | 
| 67399 | 1770 | "\<lbrakk> mcont lub ord lub_spmf (ord_spmf (=)) f; mcont lub ord lub_spmf (ord_spmf (=)) g \<rbrakk> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1771 | \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. rel_spmf R (f x) (g x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1772 | by(rule admissible_subst[OF admissible_rel_spmf, where f="\<lambda>x. (f x, g x)", simplified])(rule mcont_Pair) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1773 | |
| 63343 | 1774 | context includes lifting_syntax | 
| 1775 | begin | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1776 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1777 | lemma fixp_spmf_parametric': | 
| 67399 | 1778 | assumes f: "\<And>x. monotone (ord_spmf (=)) (ord_spmf (=)) F" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1779 | and g: "\<And>x. monotone (ord_spmf (=)) (ord_spmf (=)) G" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1780 | and param: "(rel_spmf R ===> rel_spmf R) F G" | 
| 67399 | 1781 | shows "(rel_spmf R) (ccpo.fixp lub_spmf (ord_spmf (=)) F) (ccpo.fixp lub_spmf (ord_spmf (=)) G)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1782 | by(rule parallel_fixp_induct[OF ccpo_spmf ccpo_spmf _ f g])(auto intro: param[THEN rel_funD]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1783 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1784 | lemma fixp_spmf_parametric: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1785 | assumes f: "\<And>x. mono_spmf (\<lambda>f. F f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1786 | and g: "\<And>x. mono_spmf (\<lambda>f. G f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1787 | and param: "((A ===> rel_spmf R) ===> A ===> rel_spmf R) F G" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1788 | shows "(A ===> rel_spmf R) (spmf.fixp_fun F) (spmf.fixp_fun G)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1789 | using f g | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1790 | proof(rule parallel_fixp_induct_1_1[OF partial_function_definitions_spmf partial_function_definitions_spmf _ _ reflexive reflexive, where P="(A ===> rel_spmf R)"]) | 
| 67399 | 1791 | show "ccpo.admissible (prod_lub (fun_lub lub_spmf) (fun_lub lub_spmf)) (rel_prod (fun_ord (ord_spmf (=))) (fun_ord (ord_spmf (=)))) (\<lambda>x. (A ===> rel_spmf R) (fst x) (snd x))" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1792 | unfolding rel_fun_def | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1793 | by(fastforce intro: admissible_all admissible_imp admissible_rel_spmf_mcont) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1794 |   show "(A ===> rel_spmf R) (\<lambda>_. lub_spmf {}) (\<lambda>_. lub_spmf {})" 
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1795 | by auto | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1796 | show "(A ===> rel_spmf R) (F f) (G g)" if "(A ===> rel_spmf R) f g" for f g | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1797 | using that by(rule rel_funD[OF param]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1798 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1799 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1800 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1801 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1802 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1803 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1804 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1805 | |
| 63308 | 1806 | subsection \<open>Restrictions on spmfs\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1807 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1808 | definition restrict_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf" (infixl "\<upharpoonleft>" 110) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1809 | where "p \<upharpoonleft> A = map_pmf (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) p" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1810 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1811 | lemma set_restrict_spmf [simp]: "set_spmf (p \<upharpoonleft> A) = set_spmf p \<inter> A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1812 | by(fastforce simp: restrict_spmf_def set_spmf_def split: bind_splits if_split_asm) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1813 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1814 | lemma restrict_map_spmf: "map_spmf f p \<upharpoonleft> A = map_spmf f (p \<upharpoonleft> (f -` A))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1815 | by(simp add: restrict_spmf_def pmf.map_comp o_def map_option_bind bind_map_option if_distrib cong del: if_weak_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1816 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1817 | lemma restrict_restrict_spmf [simp]: "p \<upharpoonleft> A \<upharpoonleft> B = p \<upharpoonleft> (A \<inter> B)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1818 | by(auto simp: restrict_spmf_def pmf.map_comp o_def intro!: pmf.map_cong bind_option_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1819 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1820 | lemma restrict_spmf_empty [simp]: "p \<upharpoonleft> {} = return_pmf None"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1821 | by(simp add: restrict_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1822 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1823 | lemma restrict_spmf_UNIV [simp]: "p \<upharpoonleft> UNIV = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1824 | by(simp add: restrict_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1825 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1826 | lemma spmf_restrict_spmf_outside [simp]: "x \<notin> A \<Longrightarrow> spmf (p \<upharpoonleft> A) x = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1827 | by(simp add: spmf_eq_0_set_spmf) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1828 | |
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1829 | lemma emeasure_restrict_spmf [simp]: "emeasure (measure_spmf (p \<upharpoonleft> A)) X = emeasure (measure_spmf p) (X \<inter> A)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1830 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1831 | have "(\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) -` the -` X \<inter> | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1832 | (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) -` range Some = | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1833 | the -` X \<inter> the -` A \<inter> range Some" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1834 | by(auto split: bind_splits if_split_asm) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1835 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1836 | by (simp add: restrict_spmf_def measure_spmf_def emeasure_distr emeasure_restrict_space) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1837 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1838 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1839 | lemma measure_restrict_spmf [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1840 | "measure (measure_spmf (p \<upharpoonleft> A)) X = measure (measure_spmf p) (X \<inter> A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1841 | using emeasure_restrict_spmf[of p A X] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1842 | by(simp only: measure_spmf.emeasure_eq_measure ennreal_inj measure_nonneg) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1843 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1844 | lemma spmf_restrict_spmf: "spmf (p \<upharpoonleft> A) x = (if x \<in> A then spmf p x else 0)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1845 | by(simp add: spmf_conv_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1846 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1847 | lemma spmf_restrict_spmf_inside [simp]: "x \<in> A \<Longrightarrow> spmf (p \<upharpoonleft> A) x = spmf p x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1848 | by(simp add: spmf_restrict_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1849 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1850 | lemma pmf_restrict_spmf_None: "pmf (p \<upharpoonleft> A) None = pmf p None + measure (measure_spmf p) (- A)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1851 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1852 | have [simp]: "None \<notin> Some ` (- A)" by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1853 |   have "(\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) -` {None} = {None} \<union> (Some ` (- A))"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1854 | by(auto split: bind_splits if_split_asm) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1855 | then show ?thesis unfolding ereal.inject[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1856 | by(simp add: restrict_spmf_def ennreal_pmf_map emeasure_pmf_single del: ereal.inject) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1857 | (simp add: pmf.rep_eq measure_pmf.finite_measure_Union[symmetric] measure_measure_spmf_conv_measure_pmf measure_pmf.emeasure_eq_measure) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1858 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1859 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1860 | lemma restrict_spmf_trivial: "(\<And>x. x \<in> set_spmf p \<Longrightarrow> x \<in> A) \<Longrightarrow> p \<upharpoonleft> A = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1861 | by(rule spmf_eqI)(auto simp: spmf_restrict_spmf spmf_eq_0_set_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1862 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1863 | lemma restrict_spmf_trivial': "set_spmf p \<subseteq> A \<Longrightarrow> p \<upharpoonleft> A = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1864 | by(rule restrict_spmf_trivial) blast | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1865 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1866 | lemma restrict_return_spmf: "return_spmf x \<upharpoonleft> A = (if x \<in> A then return_spmf x else return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1867 | by(simp add: restrict_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1868 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1869 | lemma restrict_return_spmf_inside [simp]: "x \<in> A \<Longrightarrow> return_spmf x \<upharpoonleft> A = return_spmf x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1870 | by(simp add: restrict_return_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1871 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1872 | lemma restrict_return_spmf_outside [simp]: "x \<notin> A \<Longrightarrow> return_spmf x \<upharpoonleft> A = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1873 | by(simp add: restrict_return_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1874 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1875 | lemma restrict_spmf_return_pmf_None [simp]: "return_pmf None \<upharpoonleft> A = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1876 | by(simp add: restrict_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1877 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1878 | lemma restrict_bind_pmf: "bind_pmf p g \<upharpoonleft> A = p \<bind> (\<lambda>x. g x \<upharpoonleft> A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1879 | by(simp add: restrict_spmf_def map_bind_pmf o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1880 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1881 | lemma restrict_bind_spmf: "bind_spmf p g \<upharpoonleft> A = p \<bind> (\<lambda>x. g x \<upharpoonleft> A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1882 | by(auto simp: bind_spmf_def restrict_bind_pmf cong del: option.case_cong_weak cong: option.case_cong intro!: bind_pmf_cong split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1883 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1884 | lemma bind_restrict_pmf: "bind_pmf (p \<upharpoonleft> A) g = p \<bind> (\<lambda>x. if x \<in> Some ` A then g x else g None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1885 | by(auto simp: restrict_spmf_def bind_map_pmf fun_eq_iff split: bind_split intro: arg_cong2[where f=bind_pmf]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1886 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1887 | lemma bind_restrict_spmf: "bind_spmf (p \<upharpoonleft> A) g = p \<bind> (\<lambda>x. if x \<in> A then g x else return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1888 | by(auto simp: bind_spmf_def bind_restrict_pmf fun_eq_iff intro: arg_cong2[where f=bind_pmf] split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1889 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1890 | lemma spmf_map_restrict: "spmf (map_spmf fst (p \<upharpoonleft> (snd -` {y}))) x = spmf p (x, y)"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1891 | by(subst spmf_map)(auto intro: arg_cong2[where f=measure] simp add: spmf_conv_measure_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1892 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1893 | lemma measure_eqI_restrict_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1894 | assumes "rel_spmf R (restrict_spmf p A) (restrict_spmf q B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1895 | shows "measure (measure_spmf p) A = measure (measure_spmf q) B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1896 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1897 | from assms have "weight_spmf (restrict_spmf p A) = weight_spmf (restrict_spmf q B)" by(rule rel_spmf_weightD) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1898 | thus ?thesis by(simp add: weight_spmf_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1899 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1900 | |
| 63308 | 1901 | subsection \<open>Subprobability distributions of sets\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1902 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1903 | definition spmf_of_set :: "'a set \<Rightarrow> 'a spmf" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1904 | where | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1905 |   "spmf_of_set A = (if finite A \<and> A \<noteq> {} then spmf_of_pmf (pmf_of_set A) else return_pmf None)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1906 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1907 | lemma spmf_of_set: "spmf (spmf_of_set A) x = indicator A x / card A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1908 | by(auto simp: spmf_of_set_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1909 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1910 | lemma pmf_spmf_of_set_None [simp]: "pmf (spmf_of_set A) None = indicator {A. infinite A \<or> A = {}} A"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1911 | by(simp add: spmf_of_set_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1912 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1913 | lemma set_spmf_of_set: "set_spmf (spmf_of_set A) = (if finite A then A else {})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1914 | by(simp add: spmf_of_set_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1915 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1916 | lemma set_spmf_of_set_finite [simp]: "finite A \<Longrightarrow> set_spmf (spmf_of_set A) = A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1917 | by(simp add: set_spmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1918 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1919 | lemma spmf_of_set_singleton: "spmf_of_set {x} = return_spmf x"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1920 | by(simp add: spmf_of_set_def pmf_of_set_singleton) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1921 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1922 | lemma map_spmf_of_set_inj_on [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1923 | "inj_on f A \<Longrightarrow> map_spmf f (spmf_of_set A) = spmf_of_set (f ` A)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1924 | by(auto simp: spmf_of_set_def map_pmf_of_set_inj dest: finite_imageD) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1925 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1926 | lemma spmf_of_pmf_pmf_of_set [simp]: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1927 |   "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> spmf_of_pmf (pmf_of_set A) = spmf_of_set A"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1928 | by(simp add: spmf_of_set_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1929 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1930 | lemma weight_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1931 |   "weight_spmf (spmf_of_set A) = (if finite A \<and> A \<noteq> {} then 1 else 0)"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1932 | by(auto simp only: spmf_of_set_def weight_spmf_of_pmf weight_return_pmf_None split: if_split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1933 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1934 | lemma weight_spmf_of_set_finite [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> weight_spmf (spmf_of_set A) = 1"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1935 | by(simp add: weight_spmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1936 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1937 | lemma weight_spmf_of_set_infinite [simp]: "infinite A \<Longrightarrow> weight_spmf (spmf_of_set A) = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1938 | by(simp add: weight_spmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1939 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1940 | lemma measure_spmf_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1941 |   "measure_spmf (spmf_of_set A) = (if finite A \<and> A \<noteq> {} then measure_pmf (pmf_of_set A) else null_measure (count_space UNIV))"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1942 | by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1943 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1944 | lemma emeasure_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1945 | "emeasure (measure_spmf (spmf_of_set S)) A = card (S \<inter> A) / card S" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1946 | by(auto simp: measure_spmf_spmf_of_set emeasure_pmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1947 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1948 | lemma measure_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1949 | "measure (measure_spmf (spmf_of_set S)) A = card (S \<inter> A) / card S" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1950 | by(auto simp: measure_spmf_spmf_of_set measure_pmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1951 | |
| 64267 | 1952 | lemma nn_integral_spmf_of_set: "nn_integral (measure_spmf (spmf_of_set A)) f = sum f A / card A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1953 | by(cases "finite A")(auto simp: spmf_of_set_def nn_integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1954 | |
| 64267 | 1955 | lemma integral_spmf_of_set: "integral\<^sup>L (measure_spmf (spmf_of_set A)) f = sum f A / card A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1956 | by (metis card.infinite div_0 division_ring_divide_zero integral_null_measure integral_pmf_of_set measure_spmf_spmf_of_set of_nat_0 sum.empty) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1957 | |
| 69597 | 1958 | notepad begin \<comment> \<open>\<^const>\<open>pmf_of_set\<close> is not fully parametric.\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1959 | define R :: "nat \<Rightarrow> nat \<Rightarrow> bool" where "R x y \<longleftrightarrow> (x \<noteq> 0 \<longrightarrow> y = 0)" for x y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1960 |   define A :: "nat set" where "A = {0, 1}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1961 |   define B :: "nat set" where "B = {0, 1, 2}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1962 | have "rel_set R A B" unfolding R_def[abs_def] A_def B_def rel_set_def by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1963 | have "\<not> rel_pmf R (pmf_of_set A) (pmf_of_set B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1964 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1965 | assume "rel_pmf R (pmf_of_set A) (pmf_of_set B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1966 | then obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1967 | and 1: "map_pmf fst pq = pmf_of_set A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1968 | and 2: "map_pmf snd pq = pmf_of_set B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1969 | by cases auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1970 | have "pmf (pmf_of_set B) 1 = 1 / 3" by(simp add: B_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1971 | have "pmf (pmf_of_set B) 2 = 1 / 3" by(simp add: B_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1972 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1973 | have "2 / 3 = pmf (pmf_of_set B) 1 + pmf (pmf_of_set B) 2" by(simp add: B_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1974 |     also have "\<dots> = measure (measure_pmf (pmf_of_set B)) ({1} \<union> {2})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1975 | by(subst measure_pmf.finite_measure_Union)(simp_all add: measure_pmf_single) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1976 |     also have "\<dots> = emeasure (measure_pmf pq) (snd -` {2, 1})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1977 | unfolding 2[symmetric] measure_pmf.emeasure_eq_measure[symmetric] by(simp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1978 |     also have "\<dots> = emeasure (measure_pmf pq) {(0, 2), (0, 1)}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 1979 | by(rule emeasure_eq_AE)(auto simp: AE_measure_pmf_iff R_def dest!: pq) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1980 |     also have "\<dots> \<le> emeasure (measure_pmf pq) (fst -` {0})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1981 | by(rule emeasure_mono) auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1982 |     also have "\<dots> = emeasure (measure_pmf (pmf_of_set A)) {0}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1983 | unfolding 1[symmetric] by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1984 | also have "\<dots> = pmf (pmf_of_set A) 0" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1985 | by(simp add: measure_pmf_single measure_pmf.emeasure_eq_measure) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1986 | also have "pmf (pmf_of_set A) 0 = 1 / 2" by(simp add: A_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1987 | finally show False by(subst (asm) ennreal_le_iff; simp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1988 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1989 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1990 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1991 | lemma rel_pmf_of_set_bij: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1992 | assumes f: "bij_betw f A B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1993 |   and A: "A \<noteq> {}" "finite A"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 1994 |   and B: "B \<noteq> {}" "finite B"
 | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1995 | and R: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1996 | shows "rel_pmf R (pmf_of_set A) (pmf_of_set B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1997 | proof(rule pmf.rel_mono_strong) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1998 | define AB where "AB = (\<lambda>x. (x, f x)) ` A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 1999 | define R' where "R' x y \<longleftrightarrow> (x, y) \<in> AB" for x y | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2000 | have "(x, y) \<in> AB" if "(x, y) \<in> set_pmf (pmf_of_set AB)" for x y | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2001 | using that by(auto simp: AB_def A) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2002 | moreover have "map_pmf fst (pmf_of_set AB) = pmf_of_set A" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2003 | by(simp add: AB_def map_pmf_of_set_inj[symmetric] inj_on_def A pmf.map_comp o_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2004 | moreover | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2005 | from f have [simp]: "inj_on f A" by(rule bij_betw_imp_inj_on) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2006 | from f have [simp]: "f ` A = B" by(rule bij_betw_imp_surj_on) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2007 | have "map_pmf snd (pmf_of_set AB) = pmf_of_set B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2008 | by(simp add: AB_def map_pmf_of_set_inj[symmetric] inj_on_def A pmf.map_comp o_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2009 | (simp add: map_pmf_of_set_inj A) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2010 | ultimately show "rel_pmf (\<lambda>x y. (x, y) \<in> AB) (pmf_of_set A) (pmf_of_set B)" .. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2011 | qed(auto intro: R) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2012 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2013 | lemma rel_spmf_of_set_bij: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2014 | assumes f: "bij_betw f A B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2015 | and R: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2016 | shows "rel_spmf R (spmf_of_set A) (spmf_of_set B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2017 | proof - | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2018 |   obtain "finite A \<longleftrightarrow> finite B" "A = {} \<longleftrightarrow> B = {}"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2019 | using bij_betw_empty1 bij_betw_empty2 bij_betw_finite f by blast | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2020 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2021 | using assms | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2022 | by (metis rel_pmf_of_set_bij rel_spmf_spmf_of_pmf return_spmf_None_parametric spmf_of_set_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2023 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2024 | |
| 63343 | 2025 | context includes lifting_syntax | 
| 2026 | begin | |
| 2027 | ||
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2028 | lemma rel_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2029 | assumes "bi_unique R" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2030 | shows "(rel_set R ===> rel_spmf R) spmf_of_set spmf_of_set" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2031 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2032 | fix A B | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2033 | assume R: "rel_set R A B" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2034 | with assms obtain f where "bij_betw f A B" and f: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2035 | by(auto dest: bi_unique_rel_set_bij_betw) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2036 | then show "rel_spmf R (spmf_of_set A) (spmf_of_set B)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2037 | by(rule rel_spmf_of_set_bij) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2038 | qed | 
| 63343 | 2039 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2040 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2041 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2042 | lemma map_mem_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2043 |   assumes "finite B" "B \<noteq> {}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2044 | shows "map_spmf (\<lambda>x. x \<in> A) (spmf_of_set B) = spmf_of_pmf (bernoulli_pmf (card (A \<inter> B) / card B))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2045 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2046 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2047 | fix i | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2048 |   have "ennreal (spmf ?lhs i) = card (B \<inter> (\<lambda>x. x \<in> A) -` {i}) / (card B)"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2049 | by(subst ennreal_spmf_map)(simp add: measure_spmf_spmf_of_set assms emeasure_pmf_of_set) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2050 | also have "\<dots> = (if i then card (B \<inter> A) / card B else card (B - A) / card B)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2051 | by(auto intro: arg_cong[where f=card]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2052 | also have "\<dots> = (if i then card (B \<inter> A) / card B else (card B - card (B \<inter> A)) / card B)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2053 | by(auto simp: card_Diff_subset_Int assms) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2054 | also have "\<dots> = ennreal (spmf ?rhs i)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2055 | by(simp add: assms card_gt_0_iff field_simps card_mono Int_commute of_nat_diff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2056 | finally show "spmf ?lhs i = spmf ?rhs i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2057 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2058 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2059 | abbreviation coin_spmf :: "bool spmf" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2060 | where "coin_spmf \<equiv> spmf_of_set UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2061 | |
| 67399 | 2062 | lemma map_eq_const_coin_spmf: "map_spmf ((=) c) coin_spmf = coin_spmf" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2063 | proof - | 
| 67399 | 2064 | have "inj ((\<longleftrightarrow>) c)" "range ((\<longleftrightarrow>) c) = UNIV" by(auto intro: inj_onI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2065 | then show ?thesis by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2066 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2067 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2068 | lemma bind_coin_spmf_eq_const: "coin_spmf \<bind> (\<lambda>x :: bool. return_spmf (b = x)) = coin_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2069 | using map_eq_const_coin_spmf unfolding map_spmf_conv_bind_spmf by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2070 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2071 | lemma bind_coin_spmf_eq_const': "coin_spmf \<bind> (\<lambda>x :: bool. return_spmf (x = b)) = coin_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2072 | by(rewrite in "_ = \<hole>" bind_coin_spmf_eq_const[symmetric, of b])(auto intro: bind_spmf_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2073 | |
| 63308 | 2074 | subsection \<open>Losslessness\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2075 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2076 | definition lossless_spmf :: "'a spmf \<Rightarrow> bool" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2077 | where "lossless_spmf p \<longleftrightarrow> weight_spmf p = 1" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2078 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2079 | lemma lossless_iff_pmf_None: "lossless_spmf p \<longleftrightarrow> pmf p None = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2080 | by(simp add: lossless_spmf_def pmf_None_eq_weight_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2081 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2082 | lemma lossless_return_spmf [iff]: "lossless_spmf (return_spmf x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2083 | by(simp add: lossless_iff_pmf_None) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2084 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2085 | lemma lossless_return_pmf_None [iff]: "\<not> lossless_spmf (return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2086 | by(simp add: lossless_iff_pmf_None) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2087 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2088 | lemma lossless_map_spmf [simp]: "lossless_spmf (map_spmf f p) \<longleftrightarrow> lossless_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2089 | by(auto simp: lossless_iff_pmf_None pmf_eq_0_set_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2090 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2091 | lemma lossless_bind_spmf [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2092 | "lossless_spmf (p \<bind> f) \<longleftrightarrow> lossless_spmf p \<and> (\<forall>x\<in>set_spmf p. lossless_spmf (f x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2093 | by(simp add: lossless_iff_pmf_None pmf_bind_spmf_None add_nonneg_eq_0_iff integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_spmf.integrable_const_bound[where B=1] pmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2094 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2095 | lemma lossless_weight_spmfD: "lossless_spmf p \<Longrightarrow> weight_spmf p = 1" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2096 | by(simp add: lossless_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2097 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2098 | lemma lossless_iff_set_pmf_None: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2099 | "lossless_spmf p \<longleftrightarrow> None \<notin> set_pmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2100 | by (simp add: lossless_iff_pmf_None pmf_eq_0_set_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2101 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2102 | lemma lossless_spmf_of_set [simp]: "lossless_spmf (spmf_of_set A) \<longleftrightarrow> finite A \<and> A \<noteq> {}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2103 | by(auto simp: lossless_spmf_def weight_spmf_of_set) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2104 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2105 | lemma lossless_spmf_spmf_of_spmf [simp]: "lossless_spmf (spmf_of_pmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2106 | by(simp add: lossless_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2107 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2108 | lemma lossless_spmf_bind_pmf [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2109 | "lossless_spmf (bind_pmf p f) \<longleftrightarrow> (\<forall>x\<in>set_pmf p. lossless_spmf (f x))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2110 | by(simp add: lossless_iff_pmf_None pmf_bind integral_nonneg_AE integral_nonneg_eq_0_iff_AE measure_pmf.integrable_const_bound[where B=1] AE_measure_pmf_iff pmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2111 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2112 | lemma lossless_spmf_conv_spmf_of_pmf: "lossless_spmf p \<longleftrightarrow> (\<exists>p'. p = spmf_of_pmf p')" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2113 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2114 | assume "lossless_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2115 | hence *: "\<And>y. y \<in> set_pmf p \<Longrightarrow> \<exists>x. y = Some x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2116 | by(case_tac y)(simp_all add: lossless_iff_set_pmf_None) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2117 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2118 | let ?p = "map_pmf the p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2119 | have "p = spmf_of_pmf ?p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2120 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2121 | fix i | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2122 |     have "ennreal (pmf (map_pmf the p) i) = \<integral>\<^sup>+ x. indicator (the -` {i}) x \<partial>p" by(simp add: ennreal_pmf_map)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2123 |     also have "\<dots> = \<integral>\<^sup>+ x. indicator {i} x \<partial>measure_spmf p" unfolding measure_spmf_def
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2124 | by(subst nn_integral_distr)(auto simp: nn_integral_restrict_space AE_measure_pmf_iff simp del: nn_integral_indicator intro!: nn_integral_cong_AE split: split_indicator dest!: * ) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2125 | also have "\<dots> = spmf p i" by(simp add: emeasure_spmf_single) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2126 | finally show "spmf p i = spmf (spmf_of_pmf ?p) i" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2127 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2128 | thus "\<exists>p'. p = spmf_of_pmf p'" .. | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2129 | qed auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2130 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2131 | lemma spmf_False_conv_True: "lossless_spmf p \<Longrightarrow> spmf p False = 1 - spmf p True" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2132 | by(clarsimp simp add: lossless_spmf_conv_spmf_of_pmf pmf_False_conv_True) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2133 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2134 | lemma spmf_True_conv_False: "lossless_spmf p \<Longrightarrow> spmf p True = 1 - spmf p False" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2135 | by(simp add: spmf_False_conv_True) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2136 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2137 | lemma bind_eq_return_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2138 | "bind_spmf p f = return_spmf x \<longleftrightarrow> (\<forall>y\<in>set_spmf p. f y = return_spmf x) \<and> lossless_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2139 | apply (simp add: bind_spmf_def bind_eq_return_pmf split: option.split) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2140 | by (metis in_set_spmf lossless_iff_set_pmf_None not_None_eq) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2141 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2142 | lemma rel_spmf_return_spmf2: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2143 | "rel_spmf R p (return_spmf x) \<longleftrightarrow> lossless_spmf p \<and> (\<forall>a\<in>set_spmf p. R a x)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2144 | apply (simp add: lossless_iff_set_pmf_None rel_pmf_return_pmf2 option_rel_Some2 in_set_spmf) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2145 | by (metis in_set_spmf not_None_eq option.sel) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2146 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2147 | lemma rel_spmf_return_spmf1: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2148 | "rel_spmf R (return_spmf x) p \<longleftrightarrow> lossless_spmf p \<and> (\<forall>a\<in>set_spmf p. R x a)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2149 | using rel_spmf_return_spmf2[of "R\<inverse>\<inverse>"] by(simp add: spmf_rel_conversep) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2150 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2151 | lemma rel_spmf_bindI1: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2152 | assumes f: "\<And>x. x \<in> set_spmf p \<Longrightarrow> rel_spmf R (f x) q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2153 | and p: "lossless_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2154 | shows "rel_spmf R (bind_spmf p f) q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2155 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2156 | fix x :: 'a | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2157 | have "rel_spmf R (bind_spmf p f) (bind_spmf (return_spmf x) (\<lambda>_. q))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2158 | by(rule rel_spmf_bindI[where R="\<lambda>x _. x \<in> set_spmf p"])(simp_all add: rel_spmf_return_spmf2 p f) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2159 | then show ?thesis by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2160 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2161 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2162 | lemma rel_spmf_bindI2: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2163 | "\<lbrakk> \<And>x. x \<in> set_spmf q \<Longrightarrow> rel_spmf R p (f x); lossless_spmf q \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2164 | \<Longrightarrow> rel_spmf R p (bind_spmf q f)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2165 | using rel_spmf_bindI1[of q "conversep R" f p] by(simp add: spmf_rel_conversep) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2166 | |
| 63308 | 2167 | subsection \<open>Scaling\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2168 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2169 | definition scale_spmf :: "real \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2170 | where | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2171 | "scale_spmf r p = embed_spmf (\<lambda>x. min (inverse (weight_spmf p)) (max 0 r) * spmf p x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2172 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2173 | lemma scale_spmf_le_1: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2174 | "(\<integral>\<^sup>+ x. min (inverse (weight_spmf p)) (max 0 r) * spmf p x \<partial>count_space UNIV) \<le> 1" (is "?lhs \<le> _") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2175 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2176 | have "?lhs = min (inverse (weight_spmf p)) (max 0 r) * \<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2177 | by(subst nn_integral_cmult[symmetric])(simp_all add: weight_spmf_nonneg max_def min_def ennreal_mult) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2178 | also have "\<dots> \<le> 1" unfolding weight_spmf_eq_nn_integral_spmf[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2179 | by(simp add: min_def max_def weight_spmf_nonneg order.strict_iff_order field_simps ennreal_mult[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2180 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2181 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2182 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2183 | lemma spmf_scale_spmf: "spmf (scale_spmf r p) x = max 0 (min (inverse (weight_spmf p)) r) * spmf p x" (is "?lhs = ?rhs") | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2184 | unfolding scale_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2185 | apply(subst spmf_embed_spmf[OF scale_spmf_le_1]) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2186 | apply(simp add: max_def min_def measure_le_0_iff field_simps weight_spmf_nonneg not_le order.strict_iff_order) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2187 | apply(metis antisym_conv order_trans weight_spmf_nonneg zero_le_mult_iff zero_le_one) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2188 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2189 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2190 | lemma real_inverse_le_1_iff: fixes x :: real | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2191 | shows "\<lbrakk> 0 \<le> x; x \<le> 1 \<rbrakk> \<Longrightarrow> 1 / x \<le> 1 \<longleftrightarrow> x = 1 \<or> x = 0" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2192 | by auto | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2193 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2194 | lemma spmf_scale_spmf': "r \<le> 1 \<Longrightarrow> spmf (scale_spmf r p) x = max 0 r * spmf p x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2195 | using real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1, of p] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2196 | by(auto simp: spmf_scale_spmf max_def min_def field_simps)(metis pmf_le_0_iff spmf_le_weight) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2197 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2198 | lemma scale_spmf_neg: "r \<le> 0 \<Longrightarrow> scale_spmf r p = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2199 | by(rule spmf_eqI)(simp add: spmf_scale_spmf' max_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2200 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2201 | lemma scale_spmf_return_None [simp]: "scale_spmf r (return_pmf None) = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2202 | by(rule spmf_eqI)(simp add: spmf_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2203 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2204 | lemma scale_spmf_conv_bind_bernoulli: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2205 | assumes "r \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2206 | shows "scale_spmf r p = bind_pmf (bernoulli_pmf r) (\<lambda>b. if b then p else return_pmf None)" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2207 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2208 | fix x | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2209 | have "\<lbrakk>weight_spmf p = 0\<rbrakk> \<Longrightarrow> spmf p x = 0" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2210 | by (metis pmf_le_0_iff spmf_le_weight) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2211 | moreover have "\<lbrakk>weight_spmf p \<noteq> 0; 1 / weight_spmf p < 1\<rbrakk> \<Longrightarrow> weight_spmf p = 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2212 | by (smt (verit) divide_less_eq_1 measure_spmf.subprob_measure_le_1 weight_spmf_lt_0) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2213 | ultimately have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2214 | using assms | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2215 | unfolding spmf_scale_spmf ennreal_pmf_bind nn_integral_measure_pmf UNIV_bool bernoulli_pmf.rep_eq | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2216 | by(auto simp: nn_integral_count_space_finite max_def min_def field_simps | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2217 | real_inverse_le_1_iff[OF weight_spmf_nonneg weight_spmf_le_1] ennreal_mult[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2218 | thus "spmf ?lhs x = spmf ?rhs x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2219 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2220 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2221 | lemma nn_integral_spmf: "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space A) = emeasure (measure_spmf p) A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2222 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2223 | have "bij_betw Some A (the -` A \<inter> range Some)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2224 | by(auto simp: bij_betw_def) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2225 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2226 | by (metis bij_betw_def emeasure_measure_spmf_conv_measure_pmf nn_integral_pmf') | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2227 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2228 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2229 | lemma measure_spmf_scale_spmf: "measure_spmf (scale_spmf r p) = scale_measure (min (inverse (weight_spmf p)) r) (measure_spmf p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2230 | by(rule measure_eqI; simp add: spmf_scale_spmf ennreal_mult' flip: nn_integral_spmf nn_integral_cmult) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2231 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2232 | lemma measure_spmf_scale_spmf': | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2233 | assumes "r \<le> 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2234 | shows "measure_spmf (scale_spmf r p) = scale_measure r (measure_spmf p)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2235 | proof(cases "weight_spmf p > 0") | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2236 | case True | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2237 | with assms show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2238 | by(simp add: measure_spmf_scale_spmf field_simps weight_spmf_le_1 mult_le_one) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2239 | next | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2240 | case False | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2241 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2242 | by (simp add: order_less_le weight_spmf_eq_0) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2243 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2244 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2245 | lemma scale_spmf_1 [simp]: "scale_spmf 1 p = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2246 | by (simp add: spmf_eqI spmf_scale_spmf') | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2247 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2248 | lemma scale_spmf_0 [simp]: "scale_spmf 0 p = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2249 | by (simp add: scale_spmf_neg) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2250 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2251 | lemma bind_scale_spmf: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2252 | assumes r: "r \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2253 | shows "bind_spmf (scale_spmf r p) f = bind_spmf p (\<lambda>x. scale_spmf r (f x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2254 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2255 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2256 | fix x | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2257 | have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2258 | using r | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2259 | by(simp add: ennreal_spmf_bind measure_spmf_scale_spmf' nn_integral_scale_measure spmf_scale_spmf' | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2260 | ennreal_mult nn_integral_cmult) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2261 | thus "spmf ?lhs x = spmf ?rhs x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2262 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2263 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2264 | lemma scale_bind_spmf: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2265 | assumes "r \<le> 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2266 | shows "scale_spmf r (bind_spmf p f) = bind_spmf p (\<lambda>x. scale_spmf r (f x))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2267 | (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2268 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2269 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2270 | have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" using assms | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2271 | unfolding spmf_scale_spmf'[OF assms] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2272 | by(simp add: ennreal_mult ennreal_spmf_bind spmf_scale_spmf' nn_integral_cmult max_def min_def) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2273 | thus "spmf ?lhs x = spmf ?rhs x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2274 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2275 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2276 | lemma bind_spmf_const: "bind_spmf p (\<lambda>x. q) = scale_spmf (weight_spmf p) q" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2277 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2278 | fix x | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2279 | have "ennreal (spmf ?lhs x) = ennreal (spmf ?rhs x)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2280 | using measure_spmf.subprob_measure_le_1[of p "space (measure_spmf p)"] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2281 | by(subst ennreal_spmf_bind)(simp add: spmf_scale_spmf' weight_spmf_le_1 ennreal_mult mult.commute max_def min_def measure_spmf.emeasure_eq_measure) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2282 | thus "spmf ?lhs x = spmf ?rhs x" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2283 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2284 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2285 | lemma map_scale_spmf: "map_spmf f (scale_spmf r p) = scale_spmf r (map_spmf f p)" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2286 | proof(rule spmf_eqI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2287 | fix i | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2288 | show "spmf ?lhs i = spmf ?rhs i" unfolding spmf_scale_spmf | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2289 | by(subst (1 2) spmf_map)(auto simp: measure_spmf_scale_spmf max_def min_def ennreal_lt_0) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2290 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2291 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2292 | lemma set_scale_spmf: "set_spmf (scale_spmf r p) = (if r > 0 then set_spmf p else {})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2293 | apply(auto simp: in_set_spmf_iff_spmf spmf_scale_spmf) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2294 | apply(simp add: min_def weight_spmf_eq_0 split: if_split_asm) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2295 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2296 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2297 | lemma set_scale_spmf' [simp]: "0 < r \<Longrightarrow> set_spmf (scale_spmf r p) = set_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2298 | by(simp add: set_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2299 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2300 | lemma rel_spmf_scaleI: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2301 | assumes "r > 0 \<Longrightarrow> rel_spmf A p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2302 | shows "rel_spmf A (scale_spmf r p) (scale_spmf r q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2303 | proof(cases "r > 0") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2304 | case True | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2305 | from assms[OF True] show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2306 | by(rule rel_spmfE)(auto simp: map_scale_spmf[symmetric] spmf_rel_map True intro: rel_spmf_reflI) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2307 | qed(simp add: not_less scale_spmf_neg) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2308 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2309 | lemma weight_scale_spmf: "weight_spmf (scale_spmf r p) = min 1 (max 0 r * weight_spmf p)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2310 | proof - | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2311 | have "\<lbrakk>1 / weight_spmf p \<le> r; ennreal r * ennreal (weight_spmf p) < 1\<rbrakk> \<Longrightarrow> weight_spmf p = 0" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2312 | by (smt (verit) ennreal_less_one_iff ennreal_mult'' measure_le_0_iff mult_imp_less_div_pos) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2313 | moreover | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2314 | have "\<lbrakk>r < 1 / weight_spmf p; 1 \<le> ennreal r * ennreal (weight_spmf p)\<rbrakk> \<Longrightarrow> weight_spmf p = 0" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2315 | by (smt (verit, ccfv_threshold) ennreal_ge_1 ennreal_mult'' mult_imp_div_pos_le weight_spmf_lt_0) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2316 | ultimately | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2317 | have "ennreal (weight_spmf (scale_spmf r p)) = min 1 (max 0 r * ennreal (weight_spmf p))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2318 | unfolding weight_spmf_eq_nn_integral_spmf | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2319 | apply(simp add: spmf_scale_spmf ennreal_mult zero_ereal_def[symmetric] nn_integral_cmult) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2320 | apply(auto simp: weight_spmf_eq_nn_integral_spmf[symmetric] field_simps min_def max_def not_le weight_spmf_lt_0 ennreal_mult[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2321 | done | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2322 | thus ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2323 | by(auto simp: min_def max_def ennreal_mult[symmetric] split: if_split_asm) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2324 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2325 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2326 | lemma weight_scale_spmf' [simp]: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2327 | "\<lbrakk> 0 \<le> r; r \<le> 1 \<rbrakk> \<Longrightarrow> weight_spmf (scale_spmf r p) = r * weight_spmf p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2328 | by(simp add: weight_scale_spmf max_def min_def)(metis antisym_conv mult_left_le order_trans weight_spmf_le_1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2329 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2330 | lemma pmf_scale_spmf_None: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2331 | "pmf (scale_spmf k p) None = 1 - min 1 (max 0 k * (1 - pmf p None))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2332 | unfolding pmf_None_eq_weight_spmf by(simp add: weight_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2333 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2334 | lemma scale_scale_spmf: | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2335 | "scale_spmf r (scale_spmf r' p) = scale_spmf (r * max 0 (min (inverse (weight_spmf p)) r')) p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2336 | (is "?lhs = ?rhs") | 
| 77434 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2337 | proof(cases "weight_spmf p > 0") | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2338 | case False | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2339 | thus ?thesis | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2340 | by (simp add: weight_spmf_eq_0 zero_less_measure_iff) | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2341 | next | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2342 | case True | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2343 | show ?thesis | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2344 | proof(rule spmf_eqI) | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2345 | fix i | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2346 | have *: "max 0 (min (1 / weight_spmf p) r') * max 0 (min (1 / min 1 (weight_spmf p * max 0 r')) r) = | 
| 77434 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2347 | max 0 (min (1 / weight_spmf p) (r * max 0 (min (1 / weight_spmf p) r')))" | 
| 
da41823d09a7
Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2348 | using True | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2349 | by (simp add: max_def) (auto simp: min_def field_simps zero_le_mult_iff) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2350 | show "spmf ?lhs i = spmf ?rhs i" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2351 | by (simp add: spmf_scale_spmf) (metis * inverse_eq_divide mult.commute weight_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2352 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2353 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2354 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2355 | lemma scale_scale_spmf' [simp]: | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2356 | assumes "0 \<le> r" "r \<le> 1" "0 \<le> r'" "r' \<le> 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2357 | shows "scale_spmf r (scale_spmf r' p) = scale_spmf (r * r') p" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2358 | proof(cases "weight_spmf p > 0") | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2359 | case True | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2360 | with assms have "r' = 1" if "1 \<le> r' * weight_spmf p" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2361 | by (smt (verit, best) measure_spmf.subprob_measure_le_1 mult_eq_1 mult_le_one that) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2362 | with assms True show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2363 | by (smt (verit, best) eq_divide_imp measure_le_0_iff mult.assoc mult_nonneg_nonneg scale_scale_spmf weight_scale_spmf') | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2364 | next | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2365 | case False | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2366 | with assms show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2367 | by (simp add: weight_spmf_eq_0 zero_less_measure_iff) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2368 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2369 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2370 | lemma scale_spmf_eq_same: "scale_spmf r p = p \<longleftrightarrow> weight_spmf p = 0 \<or> r = 1 \<or> r \<ge> 1 \<and> weight_spmf p = 1" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2371 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2372 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2373 | assume ?lhs | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2374 | hence "weight_spmf (scale_spmf r p) = weight_spmf p" by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2375 | hence *: "min 1 (max 0 r * weight_spmf p) = weight_spmf p" by(simp add: weight_scale_spmf) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2376 | hence **: "weight_spmf p = 0 \<or> r \<ge> 1" by(auto simp: min_def max_def split: if_split_asm) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2377 | show ?rhs | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2378 | proof(cases "weight_spmf p = 0") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2379 | case False | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2380 | with ** have "r \<ge> 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2381 | by simp | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2382 | with * False have "r = 1 \<or> weight_spmf p = 1" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2383 | by(simp add: max_def min_def not_le split: if_split_asm) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2384 | with \<open>r \<ge> 1\<close> show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2385 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2386 | qed simp | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2387 | next | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2388 | show "weight_spmf p = 0 \<or> r = 1 \<or> 1 \<le> r \<and> weight_spmf p = 1 \<Longrightarrow> scale_spmf r p = p" | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2389 | by (smt (verit) div_by_1 inverse_eq_divide inverse_positive_iff_positive scale_scale_spmf scale_spmf_1) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2390 | qed | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2391 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2392 | lemma map_const_spmf_of_set: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2393 |   "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> map_spmf (\<lambda>_. c) (spmf_of_set A) = return_spmf c"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2394 | by(simp add: map_spmf_conv_bind_spmf bind_spmf_const) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2395 | |
| 63308 | 2396 | subsection \<open>Conditional spmfs\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2397 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2398 | lemma set_pmf_Int_Some: "set_pmf p \<inter> Some ` A = {} \<longleftrightarrow> set_spmf p \<inter> A = {}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2399 | by(auto simp: in_set_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2400 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2401 | lemma measure_spmf_zero_iff: "measure (measure_spmf p) A = 0 \<longleftrightarrow> set_spmf p \<inter> A = {}"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2402 | unfolding measure_measure_spmf_conv_measure_pmf by(simp add: measure_pmf_zero_iff set_pmf_Int_Some) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2403 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2404 | definition cond_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2405 |   where "cond_spmf p A = (if set_spmf p \<inter> A = {} then return_pmf None else cond_pmf p (Some ` A))"
 | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2406 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2407 | lemma set_cond_spmf [simp]: "set_spmf (cond_spmf p A) = set_spmf p \<inter> A" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2408 | by(auto 4 4 simp add: cond_spmf_def in_set_spmf iff: set_cond_pmf[THEN set_eq_iff[THEN iffD1], THEN spec, rotated]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2409 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2410 | lemma cond_map_spmf [simp]: "cond_spmf (map_spmf f p) A = map_spmf f (cond_spmf p (f -` A))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2411 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2412 | have "map_option f -` Some ` A = Some ` f -` A" by auto | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2413 |   moreover have "set_pmf p \<inter> map_option f -` Some ` A \<noteq> {}" if "Some x \<in> set_pmf p" "f x \<in> A" for x
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2414 | using that by auto | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2415 | ultimately show ?thesis by(auto simp: cond_spmf_def in_set_spmf cond_map_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2416 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2417 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2418 | lemma spmf_cond_spmf [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2419 | "spmf (cond_spmf p A) x = (if x \<in> A then spmf p x / measure (measure_spmf p) A else 0)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2420 | by(auto simp: cond_spmf_def pmf_cond set_pmf_Int_Some[symmetric] measure_measure_spmf_conv_measure_pmf measure_pmf_zero_iff) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2421 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2422 | lemma bind_eq_return_pmf_None: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2423 | "bind_spmf p f = return_pmf None \<longleftrightarrow> (\<forall>x\<in>set_spmf p. f x = return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2424 | by(auto simp: bind_spmf_def bind_eq_return_pmf in_set_spmf split: option.splits) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2425 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2426 | lemma return_pmf_None_eq_bind: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2427 | "return_pmf None = bind_spmf p f \<longleftrightarrow> (\<forall>x\<in>set_spmf p. f x = return_pmf None)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2428 | using bind_eq_return_pmf_None[of p f] by auto | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2429 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2430 | (* Conditional probabilities do not seem to interact nicely with bind. *) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2431 | |
| 63308 | 2432 | subsection \<open>Product spmf\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2433 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2434 | definition pair_spmf :: "'a spmf \<Rightarrow> 'b spmf \<Rightarrow> ('a \<times> 'b) spmf"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2435 | where "pair_spmf p q = bind_pmf (pair_pmf p q) (\<lambda>xy. case xy of (Some x, Some y) \<Rightarrow> return_spmf (x, y) | _ \<Rightarrow> return_pmf None)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2436 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2437 | lemma map_fst_pair_spmf [simp]: "map_spmf fst (pair_spmf p q) = scale_spmf (weight_spmf q) p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2438 | unfolding bind_spmf_const[symmetric] | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2439 | apply(simp add: pair_spmf_def map_bind_pmf pair_pmf_def bind_assoc_pmf option.case_distrib) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2440 | apply(subst bind_commute_pmf) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2441 | apply(force intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2442 | option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong: option.case_cong) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2443 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2444 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2445 | lemma map_snd_pair_spmf [simp]: "map_spmf snd (pair_spmf p q) = scale_spmf (weight_spmf p) q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2446 | unfolding bind_spmf_const[symmetric] | 
| 63566 | 2447 | apply(simp add: pair_spmf_def map_bind_pmf pair_pmf_def bind_assoc_pmf option.case_distrib | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2448 | cong del: option.case_cong_weak) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2449 | apply(auto intro!: bind_pmf_cong[OF refl] simp add: bind_return_pmf bind_spmf_def bind_return_pmf' case_option_collapse | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2450 | option.case_distrib[where h="map_spmf _"] option.case_distrib[symmetric] case_option_id split: option.split cong del: option.case_cong_weak) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2451 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2452 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2453 | lemma set_pair_spmf [simp]: "set_spmf (pair_spmf p q) = set_spmf p \<times> set_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2454 | by(force simp add: pair_spmf_def set_spmf_bind_pmf bind_UNION in_set_spmf split: option.splits) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2455 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2456 | lemma spmf_pair [simp]: "spmf (pair_spmf p q) (x, y) = spmf p x * spmf q y" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2457 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2458 |   have "ennreal ?lhs = \<integral>\<^sup>+ a. \<integral>\<^sup>+ b. indicator {(x, y)} (a, b) \<partial>measure_spmf q \<partial>measure_spmf p"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2459 | unfolding measure_spmf_def pair_spmf_def ennreal_pmf_bind nn_integral_pair_pmf' | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2460 | by(auto simp: zero_ereal_def[symmetric] nn_integral_distr nn_integral_restrict_space nn_integral_multc[symmetric] intro!: nn_integral_cong split: option.split split_indicator) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2461 |   also have "\<dots> = \<integral>\<^sup>+ a. (\<integral>\<^sup>+ b. indicator {y} b \<partial>measure_spmf q) * indicator {x} a \<partial>measure_spmf p"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2462 | by(subst nn_integral_multc[symmetric])(auto intro!: nn_integral_cong split: split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2463 | also have "\<dots> = ennreal ?rhs" by(simp add: emeasure_spmf_single max_def ennreal_mult mult.commute) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2464 | finally show ?thesis by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2465 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2466 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2467 | lemma pair_map_spmf2: "pair_spmf p (map_spmf f q) = map_spmf (apsnd f) (pair_spmf p q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2468 | unfolding pair_spmf_def pair_map_pmf2 bind_map_pmf map_bind_pmf | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2469 | by (intro bind_pmf_cong refl) (auto split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2470 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2471 | lemma pair_map_spmf1: "pair_spmf (map_spmf f p) q = map_spmf (apfst f) (pair_spmf p q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2472 | unfolding pair_spmf_def pair_map_pmf1 bind_map_pmf map_bind_pmf | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2473 | by (intro bind_pmf_cong refl) (auto split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2474 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2475 | lemma pair_map_spmf: "pair_spmf (map_spmf f p) (map_spmf g q) = map_spmf (map_prod f g) (pair_spmf p q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2476 | unfolding pair_map_spmf2 pair_map_spmf1 spmf.map_comp | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2477 | by(simp add: apfst_def apsnd_def o_def prod.map_comp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2478 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2479 | lemma pair_spmf_alt_def: "pair_spmf p q = bind_spmf p (\<lambda>x. bind_spmf q (\<lambda>y. return_spmf (x, y)))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2480 | unfolding pair_spmf_def pair_pmf_def bind_spmf_def bind_assoc_pmf bind_return_pmf | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2481 | by (intro bind_pmf_cong refl) (auto split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2482 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2483 | lemma weight_pair_spmf [simp]: "weight_spmf (pair_spmf p q) = weight_spmf p * weight_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2484 | unfolding pair_spmf_alt_def by(simp add: weight_bind_spmf o_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2485 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2486 | lemma pair_scale_spmf1: (* FIXME: generalise to arbitrary r *) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2487 | "r \<le> 1 \<Longrightarrow> pair_spmf (scale_spmf r p) q = scale_spmf r (pair_spmf p q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2488 | by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2489 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2490 | lemma pair_scale_spmf2: (* FIXME: generalise to arbitrary r *) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2491 | "r \<le> 1 \<Longrightarrow> pair_spmf p (scale_spmf r q) = scale_spmf r (pair_spmf p q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2492 | by(simp add: pair_spmf_alt_def scale_bind_spmf bind_scale_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2493 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2494 | lemma pair_spmf_return_None1 [simp]: "pair_spmf (return_pmf None) p = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2495 | by(rule spmf_eqI)(clarsimp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2496 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2497 | lemma pair_spmf_return_None2 [simp]: "pair_spmf p (return_pmf None) = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2498 | by(rule spmf_eqI)(clarsimp) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2499 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2500 | lemma pair_spmf_return_spmf1: "pair_spmf (return_spmf x) q = map_spmf (Pair x) q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2501 | by(rule spmf_eqI)(auto split: split_indicator simp add: spmf_map_inj' inj_on_def intro: spmf_map_outside) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2502 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2503 | lemma pair_spmf_return_spmf2: "pair_spmf p (return_spmf y) = map_spmf (\<lambda>x. (x, y)) p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2504 | by(rule spmf_eqI)(auto split: split_indicator simp add: inj_on_def intro!: spmf_map_outside spmf_map_inj'[symmetric]) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2505 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2506 | lemma pair_spmf_return_spmf [simp]: "pair_spmf (return_spmf x) (return_spmf y) = return_spmf (x, y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2507 | by(simp add: pair_spmf_return_spmf1) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2508 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2509 | lemma rel_pair_spmf_prod: | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2510 | "rel_spmf (rel_prod A B) (pair_spmf p q) (pair_spmf p' q') \<longleftrightarrow> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2511 | rel_spmf A (scale_spmf (weight_spmf q) p) (scale_spmf (weight_spmf q') p') \<and> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2512 | rel_spmf B (scale_spmf (weight_spmf p) q) (scale_spmf (weight_spmf p') q')" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2513 | (is "?lhs \<longleftrightarrow> ?rhs" is "_ \<longleftrightarrow> ?A \<and> ?B" is "_ \<longleftrightarrow> rel_spmf _ ?p ?p' \<and> rel_spmf _ ?q ?q'") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2514 | proof(intro iffI conjI) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2515 | assume ?rhs | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2516 | then obtain pq pq' where p: "map_spmf fst pq = ?p" and p': "map_spmf snd pq = ?p'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2517 | and q: "map_spmf fst pq' = ?q" and q': "map_spmf snd pq' = ?q'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2518 | and *: "\<And>x x'. (x, x') \<in> set_spmf pq \<Longrightarrow> A x x'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2519 | and **: "\<And>y y'. (y, y') \<in> set_spmf pq' \<Longrightarrow> B y y'" by(auto elim!: rel_spmfE) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2520 | let ?f = "\<lambda>((x, x'), (y, y')). ((x, y), (x', y'))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2521 | let ?r = "1 / (weight_spmf p * weight_spmf q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2522 | let ?pq = "scale_spmf ?r (map_spmf ?f (pair_spmf pq pq'))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2523 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2524 |   { fix p :: "'x spmf" and q :: "'y spmf"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2525 | assume "weight_spmf q \<noteq> 0" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2526 | and "weight_spmf p \<noteq> 0" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2527 | and "1 / (weight_spmf p * weight_spmf q) \<le> weight_spmf p * weight_spmf q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2528 | hence "1 \<le> (weight_spmf p * weight_spmf q) * (weight_spmf p * weight_spmf q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2529 | by(simp add: pos_divide_le_eq order.strict_iff_order weight_spmf_nonneg) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2530 | moreover have "(weight_spmf p * weight_spmf q) * (weight_spmf p * weight_spmf q) \<le> (1 * 1) * (1 * 1)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2531 | by(intro mult_mono)(simp_all add: weight_spmf_nonneg weight_spmf_le_1) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2532 | ultimately have "(weight_spmf p * weight_spmf q) * (weight_spmf p * weight_spmf q) = 1" by simp | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2533 | hence *: "weight_spmf p * weight_spmf q = 1" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2534 | by(metis antisym_conv less_le mult_less_cancel_left1 weight_pair_spmf weight_spmf_le_1 weight_spmf_nonneg) | 
| 63540 | 2535 | hence **: "weight_spmf p = 1" by(metis antisym_conv mult_left_le weight_spmf_le_1 weight_spmf_nonneg) | 
| 2536 | moreover from * ** have "weight_spmf q = 1" by simp | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2537 | moreover note calculation } | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2538 | note full = this | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2539 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2540 | show ?lhs | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2541 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2542 | have [simp]: "fst \<circ> ?f = map_prod fst fst" by(simp add: fun_eq_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2543 | have "map_spmf fst ?pq = scale_spmf ?r (pair_spmf ?p ?q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2544 | by(simp add: pair_map_spmf[symmetric] p q map_scale_spmf spmf.map_comp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2545 | also have "\<dots> = pair_spmf p q" using full[of p q] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2546 | by(simp add: pair_scale_spmf1 pair_scale_spmf2 weight_spmf_le_1 weight_spmf_nonneg) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2547 | (auto simp: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2548 | finally show "map_spmf fst ?pq = \<dots>" . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2549 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2550 | have [simp]: "snd \<circ> ?f = map_prod snd snd" by(simp add: fun_eq_iff) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2551 | from \<open>?rhs\<close> have eq: "weight_spmf p * weight_spmf q = weight_spmf p' * weight_spmf q'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2552 | by(auto dest!: rel_spmf_weightD simp add: weight_spmf_le_1 weight_spmf_nonneg) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2553 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2554 | have "map_spmf snd ?pq = scale_spmf ?r (pair_spmf ?p' ?q')" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2555 | by(simp add: pair_map_spmf[symmetric] p' q' map_scale_spmf spmf.map_comp) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2556 | also have "\<dots> = pair_spmf p' q'" using full[of p' q'] eq | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2557 | by(simp add: pair_scale_spmf1 pair_scale_spmf2 weight_spmf_le_1 weight_spmf_nonneg) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2558 | (auto simp: scale_scale_spmf max_def min_def field_simps weight_spmf_nonneg weight_spmf_eq_0) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2559 | finally show "map_spmf snd ?pq = \<dots>" . | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2560 | qed(auto simp: set_scale_spmf split: if_split_asm dest: * ** ) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2561 | next | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2562 | assume ?lhs | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2563 | then obtain pq where pq: "map_spmf fst pq = pair_spmf p q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2564 | and pq': "map_spmf snd pq = pair_spmf p' q'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2565 | and *: "\<And>x y x' y'. ((x, y), (x', y')) \<in> set_spmf pq \<Longrightarrow> A x x' \<and> B y y'" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2566 | by(auto elim: rel_spmfE) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2567 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2568 | show ?A | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2569 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2570 | let ?f = "(\<lambda>((x, y), (x', y')). (x, x'))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2571 | let ?pq = "map_spmf ?f pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2572 | have [simp]: "fst \<circ> ?f = fst \<circ> fst" by(simp add: split_def o_def) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2573 | show "map_spmf fst ?pq = scale_spmf (weight_spmf q) p" using pq | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2574 | by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2575 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2576 | have [simp]: "snd \<circ> ?f = fst \<circ> snd" by(simp add: split_def o_def) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2577 | show "map_spmf snd ?pq = scale_spmf (weight_spmf q') p'" using pq' | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2578 | by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2579 | qed(auto dest: * ) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2580 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2581 | show ?B | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2582 | proof | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2583 | let ?f = "(\<lambda>((x, y), (x', y')). (y, y'))" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2584 | let ?pq = "map_spmf ?f pq" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2585 | have [simp]: "fst \<circ> ?f = snd \<circ> fst" by(simp add: split_def o_def) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2586 | show "map_spmf fst ?pq = scale_spmf (weight_spmf p) q" using pq | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2587 | by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2588 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2589 | have [simp]: "snd \<circ> ?f = snd \<circ> snd" by(simp add: split_def o_def) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63308diff
changeset | 2590 | show "map_spmf snd ?pq = scale_spmf (weight_spmf p') q'" using pq' | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2591 | by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2592 | qed(auto dest: * ) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2593 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2594 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2595 | lemma pair_pair_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2596 | "pair_spmf (pair_spmf p q) r = map_spmf (\<lambda>(x, (y, z)). ((x, y), z)) (pair_spmf p (pair_spmf q r))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2597 | by(simp add: pair_spmf_alt_def map_spmf_conv_bind_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2598 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2599 | lemma pair_commute_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2600 | "pair_spmf p q = map_spmf (\<lambda>(y, x). (x, y)) (pair_spmf q p)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2601 | unfolding pair_spmf_alt_def by(subst bind_commute_spmf)(simp add: map_spmf_conv_bind_spmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2602 | |
| 63308 | 2603 | subsection \<open>Assertions\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2604 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2605 | definition assert_spmf :: "bool \<Rightarrow> unit spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2606 | where "assert_spmf b = (if b then return_spmf () else return_pmf None)" | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2607 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2608 | lemma assert_spmf_simps [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2609 | "assert_spmf True = return_spmf ()" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2610 | "assert_spmf False = return_pmf None" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2611 | by(simp_all add: assert_spmf_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2612 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2613 | lemma in_set_assert_spmf [simp]: "x \<in> set_spmf (assert_spmf p) \<longleftrightarrow> p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2614 | by(cases p) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2615 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2616 | lemma set_spmf_assert_spmf_eq_empty [simp]: "set_spmf (assert_spmf b) = {} \<longleftrightarrow> \<not> b"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2617 | by auto | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2618 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2619 | lemma lossless_assert_spmf [iff]: "lossless_spmf (assert_spmf b) \<longleftrightarrow> b" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2620 | by(cases b) simp_all | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2621 | |
| 63308 | 2622 | subsection \<open>Try\<close> | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2623 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2624 | definition try_spmf :: "'a spmf \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf" ("TRY _ ELSE _" [0,60] 59)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2625 | where "try_spmf p q = bind_pmf p (\<lambda>x. case x of None \<Rightarrow> q | Some y \<Rightarrow> return_spmf y)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2626 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2627 | lemma try_spmf_lossless [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2628 | assumes "lossless_spmf p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2629 | shows "TRY p ELSE q = p" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2630 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2631 | have "TRY p ELSE q = bind_pmf p return_pmf" unfolding try_spmf_def using assms | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2632 | by(auto simp: lossless_iff_set_pmf_None split: option.split intro: bind_pmf_cong) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2633 | thus ?thesis by(simp add: bind_return_pmf') | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2634 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2635 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2636 | lemma try_spmf_return_spmf1: "TRY return_spmf x ELSE q = return_spmf x" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2637 | by simp | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2638 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2639 | lemma try_spmf_return_None [simp]: "TRY return_pmf None ELSE q = q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2640 | by(simp add: try_spmf_def bind_return_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2641 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2642 | lemma try_spmf_return_pmf_None2 [simp]: "TRY p ELSE return_pmf None = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2643 | by(simp add: try_spmf_def option.case_distrib[symmetric] bind_return_pmf' case_option_id) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2644 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2645 | lemma map_try_spmf: "map_spmf f (try_spmf p q) = try_spmf (map_spmf f p) (map_spmf f q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2646 | by(simp add: try_spmf_def map_bind_pmf bind_map_pmf option.case_distrib[where h="map_spmf f"] o_def cong del: option.case_cong_weak) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2647 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2648 | lemma try_spmf_bind_pmf: "TRY (bind_pmf p f) ELSE q = bind_pmf p (\<lambda>x. TRY (f x) ELSE q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2649 | by(simp add: try_spmf_def bind_assoc_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2650 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2651 | lemma try_spmf_bind_spmf_lossless: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2652 | "lossless_spmf p \<Longrightarrow> TRY (bind_spmf p f) ELSE q = bind_spmf p (\<lambda>x. TRY (f x) ELSE q)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2653 | by (metis (mono_tags, lifting) bind_spmf_of_pmf lossless_spmf_conv_spmf_of_pmf try_spmf_bind_pmf) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2654 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2655 | lemma try_spmf_bind_out: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2656 | "lossless_spmf p \<Longrightarrow> bind_spmf p (\<lambda>x. TRY (f x) ELSE q) = TRY (bind_spmf p f) ELSE q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2657 | by(simp add: try_spmf_bind_spmf_lossless) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2658 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2659 | lemma lossless_try_spmf [simp]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2660 | "lossless_spmf (TRY p ELSE q) \<longleftrightarrow> lossless_spmf p \<or> lossless_spmf q" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2661 | by(auto simp: try_spmf_def in_set_spmf lossless_iff_set_pmf_None split: option.splits) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2662 | |
| 63343 | 2663 | context includes lifting_syntax | 
| 2664 | begin | |
| 2665 | ||
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2666 | lemma try_spmf_parametric [transfer_rule]: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2667 | "(rel_spmf A ===> rel_spmf A ===> rel_spmf A) try_spmf try_spmf" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2668 | unfolding try_spmf_def[abs_def] by transfer_prover | 
| 63343 | 2669 | |
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2670 | end | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2671 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2672 | lemma try_spmf_cong: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2673 | "\<lbrakk> p = p'; \<not> lossless_spmf p' \<Longrightarrow> q = q' \<rbrakk> \<Longrightarrow> TRY p ELSE q = TRY p' ELSE q'" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2674 | unfolding try_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2675 | by(rule bind_pmf_cong)(auto split: option.split simp add: lossless_iff_set_pmf_None) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2676 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2677 | lemma rel_spmf_try_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2678 | "\<lbrakk> rel_spmf R p p'; \<not> lossless_spmf p' \<Longrightarrow> rel_spmf R q q' \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2679 | \<Longrightarrow> rel_spmf R (TRY p ELSE q) (TRY p' ELSE q')" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2680 | unfolding try_spmf_def | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2681 | apply(rule rel_pmf_bindI[where R="\<lambda>x y. rel_option R x y \<and> x \<in> set_pmf p \<and> y \<in> set_pmf p'"]) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2682 | apply (simp add: pmf.rel_mono_strong) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2683 | apply(auto split: option.split simp add: lossless_iff_set_pmf_None) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2684 | done | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2685 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2686 | lemma spmf_try_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2687 | "spmf (TRY p ELSE q) x = spmf p x + pmf p None * spmf q x" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2688 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2689 |   have "ennreal (spmf (TRY p ELSE q) x) = \<integral>\<^sup>+ y. ennreal (spmf q x) * indicator {None} y + indicator {Some x} y \<partial>measure_pmf p"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2690 | unfolding try_spmf_def ennreal_pmf_bind by(rule nn_integral_cong)(simp split: option.split split_indicator) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2691 |   also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (spmf q x) * indicator {None} y \<partial>measure_pmf p) + \<integral>\<^sup>+ y. indicator {Some x} y \<partial>measure_pmf p"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2692 | by(simp add: nn_integral_add) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2693 | also have "\<dots> = ennreal (spmf q x) * pmf p None + spmf p x" by(simp add: emeasure_pmf_single) | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2694 | finally show ?thesis by(simp flip: ennreal_plus ennreal_mult) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2695 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2696 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2697 | lemma try_scale_spmf_same [simp]: "lossless_spmf p \<Longrightarrow> TRY scale_spmf k p ELSE p = p" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2698 | by(rule spmf_eqI)(auto simp: spmf_try_spmf spmf_scale_spmf pmf_scale_spmf_None lossless_iff_pmf_None weight_spmf_conv_pmf_None min_def max_def field_simps) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2699 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2700 | lemma pmf_try_spmf_None [simp]: "pmf (TRY p ELSE q) None = pmf p None * pmf q None" (is "?lhs = ?rhs") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2701 | proof - | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2702 |   have "?lhs = \<integral> x. pmf q None * indicator {None} x \<partial>measure_pmf p"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63626diff
changeset | 2703 | unfolding try_spmf_def pmf_bind by(rule Bochner_Integration.integral_cong)(simp_all split: option.split) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2704 | also have "\<dots> = ?rhs" by(simp add: measure_pmf_single) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2705 | finally show ?thesis . | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2706 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2707 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2708 | lemma try_bind_spmf_lossless2: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2709 | "lossless_spmf q \<Longrightarrow> TRY (bind_spmf p f) ELSE q = TRY (p \<bind> (\<lambda>x. TRY (f x) ELSE q)) ELSE q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2710 | by(rule spmf_eqI)(simp add: spmf_try_spmf pmf_bind_spmf_None spmf_bind field_simps measure_spmf.integrable_const_bound[where B=1] pmf_le_1 lossless_iff_pmf_None) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2711 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2712 | lemma try_bind_spmf_lossless2': | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2713 | fixes f :: "'a \<Rightarrow> 'b spmf" shows | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2714 | "\<lbrakk> NO_MATCH (\<lambda>x :: 'a. try_spmf (g x :: 'b spmf) (h x)) f; lossless_spmf q \<rbrakk> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2715 | \<Longrightarrow> TRY (bind_spmf p f) ELSE q = TRY (p \<bind> (\<lambda>x :: 'a. TRY (f x) ELSE q)) ELSE q" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2716 | by(rule try_bind_spmf_lossless2) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2717 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2718 | lemma try_bind_assert_spmf: | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2719 | "TRY (assert_spmf b \<bind> f) ELSE q = (if b then TRY (f ()) ELSE q else q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2720 | by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2721 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2722 | subsection \<open>Miscellaneous\<close> | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2723 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2724 | lemma assumes "rel_spmf (\<lambda>x y. bad1 x = bad2 y \<and> (\<not> bad2 y \<longrightarrow> A x \<longleftrightarrow> B y)) p q" (is "rel_spmf ?A _ _") | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2725 |   shows fundamental_lemma_bad: "measure (measure_spmf p) {x. bad1 x} = measure (measure_spmf q) {y. bad2 y}" (is "?bad")
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2726 |   and fundamental_lemma: "\<bar>measure (measure_spmf p) {x. A x} - measure (measure_spmf q) {y. B y}\<bar> \<le>
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2727 |     measure (measure_spmf p) {x. bad1 x}" (is ?fundamental)
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2728 | proof - | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2729 | have good: "rel_fun ?A (=) (\<lambda>x. A x \<and> \<not> bad1 x) (\<lambda>y. B y \<and> \<not> bad2 y)" by(auto simp: rel_fun_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2730 |   from assms have 1: "measure (measure_spmf p) {x. A x \<and> \<not> bad1 x} = measure (measure_spmf q) {y. B y \<and> \<not> bad2 y}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2731 | by(rule measure_spmf_parametric[THEN rel_funD, THEN rel_funD])(rule Collect_parametric[THEN rel_funD, OF good]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2732 | |
| 67399 | 2733 | have bad: "rel_fun ?A (=) bad1 bad2" by(simp add: rel_fun_def) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2734 | show 2: ?bad using assms | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2735 | by(rule measure_spmf_parametric[THEN rel_funD, THEN rel_funD])(rule Collect_parametric[THEN rel_funD, OF bad]) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2736 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2737 | let ?\<mu>p = "measure (measure_spmf p)" and ?\<mu>q = "measure (measure_spmf q)" | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2738 |   have "{x. A x \<and> bad1 x} \<union> {x. A x \<and> \<not> bad1 x} = {x. A x}"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2739 |     and "{y. B y \<and> bad2 y} \<union> {y. B y \<and> \<not> bad2 y} = {y. B y}" by auto
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2740 |   then have "\<bar>?\<mu>p {x. A x} - ?\<mu>q {x. B x}\<bar> = \<bar>?\<mu>p ({x. A x \<and> bad1 x} \<union> {x. A x \<and> \<not> bad1 x}) - ?\<mu>q ({y. B y \<and> bad2 y} \<union> {y. B y \<and> \<not> bad2 y})\<bar>"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2741 | by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2742 |   also have "\<dots> = \<bar>?\<mu>p {x. A x \<and> bad1 x} + ?\<mu>p {x. A x \<and> \<not> bad1 x} - ?\<mu>q {y. B y \<and> bad2 y} - ?\<mu>q {y. B y \<and> \<not> bad2 y}\<bar>"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2743 | by(subst (1 2) measure_Union)(auto) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2744 |   also have "\<dots> = \<bar>?\<mu>p {x. A x \<and> bad1 x} - ?\<mu>q {y. B y \<and> bad2 y}\<bar>" using 1 by simp
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2745 |   also have "\<dots> \<le> max (?\<mu>p {x. A x \<and> bad1 x}) (?\<mu>q {y. B y \<and> bad2 y})"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
77434diff
changeset | 2746 | by(rule abs_leI)(auto simp: max_def not_le, simp_all only: add_increasing measure_nonneg mult_2) | 
| 63243 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2747 |   also have "\<dots> \<le> max (?\<mu>p {x. bad1 x}) (?\<mu>q {y. bad2 y})"
 | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2748 | by(rule max.mono; rule measure_spmf.finite_measure_mono; auto) | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2749 | also note 2[symmetric] | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2750 | finally show ?fundamental by simp | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2751 | qed | 
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2752 | |
| 
1bc6816fd525
add theory of discrete subprobability distributions
 Andreas Lochbihler parents: diff
changeset | 2753 | end |