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