| author | wenzelm |
| Mon, 05 Jan 2015 22:41:09 +0100 | |
| changeset 59292 | fef652c88263 |
| parent 59134 | a71f2e256ee2 |
| child 59325 | 922d31f5c3f5 |
| permissions | -rw-r--r-- |
| 58606 | 1 |
(* Title: HOL/Probability/Probability_Mass_Function.thy |
| 59023 | 2 |
Author: Johannes Hölzl, TU München |
3 |
Author: Andreas Lochbihler, ETH Zurich |
|
4 |
*) |
|
| 58606 | 5 |
|
| 59000 | 6 |
section \<open> Probability mass function \<close> |
7 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
8 |
theory Probability_Mass_Function |
| 59000 | 9 |
imports |
10 |
Giry_Monad |
|
| 59093 | 11 |
"~~/src/HOL/Number_Theory/Binomial" |
| 59000 | 12 |
"~~/src/HOL/Library/Multiset" |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
13 |
begin |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
14 |
|
| 59052 | 15 |
lemma bind_return'': "sets M = sets N \<Longrightarrow> M \<guillemotright>= return N = M" |
16 |
by (cases "space M = {}")
|
|
17 |
(simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return' |
|
18 |
cong: subprob_algebra_cong) |
|
19 |
||
20 |
||
21 |
lemma (in prob_space) distr_const[simp]: |
|
22 |
"c \<in> space N \<Longrightarrow> distr M N (\<lambda>x. c) = return N c" |
|
23 |
by (rule measure_eqI) (auto simp: emeasure_distr emeasure_space_1) |
|
24 |
||
25 |
lemma (in finite_measure) countable_support: |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
26 |
"countable {x. measure M {x} \<noteq> 0}"
|
| 59000 | 27 |
proof cases |
28 |
assume "measure M (space M) = 0" |
|
29 |
with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
|
|
30 |
by auto |
|
31 |
then show ?thesis |
|
32 |
by simp |
|
33 |
next |
|
34 |
let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
|
|
35 |
assume "?M \<noteq> 0" |
|
36 |
then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
|
|
37 |
using reals_Archimedean[of "?m x / ?M" for x] |
|
38 |
by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff) |
|
39 |
have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
|
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
40 |
proof (rule ccontr) |
| 59000 | 41 |
fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
42 |
then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X" |
| 58606 | 43 |
by (metis infinite_arbitrarily_large) |
| 59000 | 44 |
from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x" |
45 |
by auto |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
46 |
{ fix x assume "x \<in> X"
|
| 59000 | 47 |
from `?M \<noteq> 0` *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
48 |
then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
49 |
note singleton_sets = this |
| 59000 | 50 |
have "?M < (\<Sum>x\<in>X. ?M / Suc n)" |
51 |
using `?M \<noteq> 0` |
|
52 |
by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc field_simps less_le measure_nonneg) |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
53 |
also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
54 |
by (rule setsum_mono) fact |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
55 |
also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
56 |
using singleton_sets `finite X` |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
57 |
by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) |
| 59000 | 58 |
finally have "?M < measure M (\<Union>x\<in>X. {x})" .
|
59 |
moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
|
|
60 |
using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto |
|
61 |
ultimately show False by simp |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
62 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
63 |
show ?thesis |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
64 |
unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
65 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
66 |
|
| 59000 | 67 |
lemma (in finite_measure) AE_support_countable: |
68 |
assumes [simp]: "sets M = UNIV" |
|
69 |
shows "(AE x in M. measure M {x} \<noteq> 0) \<longleftrightarrow> (\<exists>S. countable S \<and> (AE x in M. x \<in> S))"
|
|
70 |
proof |
|
71 |
assume "\<exists>S. countable S \<and> (AE x in M. x \<in> S)" |
|
72 |
then obtain S where S[intro]: "countable S" and ae: "AE x in M. x \<in> S" |
|
73 |
by auto |
|
74 |
then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
|
|
75 |
(\<integral>\<^sup>+ x. emeasure M {x} * indicator {x\<in>S. emeasure M {x} \<noteq> 0} x \<partial>count_space UNIV)"
|
|
76 |
by (subst emeasure_UN_countable) |
|
77 |
(auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space) |
|
78 |
also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} * indicator S x \<partial>count_space UNIV)"
|
|
79 |
by (auto intro!: nn_integral_cong split: split_indicator) |
|
80 |
also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
|
|
81 |
by (subst emeasure_UN_countable) |
|
82 |
(auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space) |
|
83 |
also have "\<dots> = emeasure M (space M)" |
|
84 |
using ae by (intro emeasure_eq_AE) auto |
|
85 |
finally have "emeasure M {x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0} = emeasure M (space M)"
|
|
86 |
by (simp add: emeasure_single_in_space cong: rev_conj_cong) |
|
87 |
with finite_measure_compl[of "{x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0}"]
|
|
88 |
have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
|
|
89 |
by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure set_diff_eq cong: conj_cong) |
|
90 |
then show "AE x in M. measure M {x} \<noteq> 0"
|
|
91 |
by (auto simp: emeasure_eq_measure) |
|
92 |
qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
|
|
93 |
||
94 |
subsection {* PMF as measure *}
|
|
95 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
96 |
typedef 'a pmf = "{M :: 'a measure. prob_space M \<and> sets M = UNIV \<and> (AE x in M. measure M {x} \<noteq> 0)}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
97 |
morphisms measure_pmf Abs_pmf |
| 58606 | 98 |
by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
|
99 |
(auto intro!: prob_space_uniform_measure AE_uniform_measureI) |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
100 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
101 |
declare [[coercion measure_pmf]] |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
102 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
103 |
lemma prob_space_measure_pmf: "prob_space (measure_pmf p)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
104 |
using pmf.measure_pmf[of p] by auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
105 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
106 |
interpretation measure_pmf!: prob_space "measure_pmf M" for M |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
107 |
by (rule prob_space_measure_pmf) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
108 |
|
| 59000 | 109 |
interpretation measure_pmf!: subprob_space "measure_pmf M" for M |
110 |
by (rule prob_space_imp_subprob_space) unfold_locales |
|
111 |
||
| 59048 | 112 |
lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)" |
113 |
by unfold_locales |
|
114 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
115 |
locale pmf_as_measure |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
116 |
begin |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
117 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
118 |
setup_lifting type_definition_pmf |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
119 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
120 |
end |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
121 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
122 |
context |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
123 |
begin |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
124 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
125 |
interpretation pmf_as_measure . |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
126 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
127 |
lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
128 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
129 |
lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
130 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
131 |
lift_definition map_pmf :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf" is
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
132 |
"\<lambda>f M. distr M (count_space UNIV) f" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
133 |
proof safe |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
134 |
fix M and f :: "'a \<Rightarrow> 'b" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
135 |
let ?D = "distr M (count_space UNIV) f" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
136 |
assume "prob_space M" and [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
137 |
interpret prob_space M by fact |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
138 |
from ae have "AE x in M. measure M (f -` {f x}) \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
139 |
proof eventually_elim |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
140 |
fix x |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
141 |
have "measure M {x} \<le> measure M (f -` {f x})"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
142 |
by (intro finite_measure_mono) auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
143 |
then show "measure M {x} \<noteq> 0 \<Longrightarrow> measure M (f -` {f x}) \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
144 |
using measure_nonneg[of M "{x}"] by auto
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
145 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
146 |
then show "AE x in ?D. measure ?D {x} \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
147 |
by (simp add: AE_distr_iff measure_distr measurable_def) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
148 |
qed (auto simp: measurable_def prob_space.prob_space_distr) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
149 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
150 |
declare [[coercion set_pmf]] |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
151 |
|
| 59023 | 152 |
lemma countable_set_pmf [simp]: "countable (set_pmf p)" |
| 59000 | 153 |
by transfer (metis prob_space.finite_measure finite_measure.countable_support) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
154 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
155 |
lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
156 |
by transfer metis |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
157 |
|
| 59048 | 158 |
lemma sets_measure_pmf_count_space[measurable_cong]: |
159 |
"sets (measure_pmf M) = sets (count_space UNIV)" |
|
| 59000 | 160 |
by simp |
161 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
162 |
lemma space_measure_pmf[simp]: "space (measure_pmf p) = UNIV" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
163 |
using sets_eq_imp_space_eq[of "measure_pmf p" "count_space UNIV"] by simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
164 |
|
| 59048 | 165 |
lemma measure_pmf_in_subprob_algebra[measurable (raw)]: "measure_pmf x \<in> space (subprob_algebra (count_space UNIV))" |
166 |
by (simp add: space_subprob_algebra subprob_space_measure_pmf) |
|
167 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
168 |
lemma measurable_pmf_measure1[simp]: "measurable (M :: 'a pmf) N = UNIV \<rightarrow> space N" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
169 |
by (auto simp: measurable_def) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
170 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
171 |
lemma measurable_pmf_measure2[simp]: "measurable N (M :: 'a pmf) = measurable N (count_space UNIV)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
172 |
by (intro measurable_cong_sets) simp_all |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
173 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
174 |
lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
175 |
by transfer (simp add: less_le measure_nonneg) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
176 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
177 |
lemma pmf_nonneg: "0 \<le> pmf p x" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
178 |
by transfer (simp add: measure_nonneg) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
179 |
|
| 59000 | 180 |
lemma pmf_le_1: "pmf p x \<le> 1" |
181 |
by (simp add: pmf.rep_eq) |
|
182 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
183 |
lemma emeasure_pmf_single: |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
184 |
fixes M :: "'a pmf" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
185 |
shows "emeasure M {x} = pmf M x"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
186 |
by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
187 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
188 |
lemma AE_measure_pmf: "AE x in (M::'a pmf). x \<in> M" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
189 |
by transfer simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
190 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
191 |
lemma emeasure_pmf_single_eq_zero_iff: |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
192 |
fixes M :: "'a pmf" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
193 |
shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
194 |
by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
195 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
196 |
lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
197 |
proof - |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
198 |
{ fix y assume y: "y \<in> M" and P: "AE x in M. P x" "\<not> P y"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
199 |
with P have "AE x in M. x \<noteq> y" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
200 |
by auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
201 |
with y have False |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
202 |
by (simp add: emeasure_pmf_single_eq_zero_iff AE_iff_measurable[OF _ refl]) } |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
203 |
then show ?thesis |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
204 |
using AE_measure_pmf[of M] by auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
205 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
206 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
207 |
lemma set_pmf_not_empty: "set_pmf M \<noteq> {}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
208 |
using AE_measure_pmf[of M] by (intro notI) simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
209 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
210 |
lemma set_pmf_iff: "x \<in> set_pmf M \<longleftrightarrow> pmf M x \<noteq> 0" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
211 |
by transfer simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
212 |
|
| 59000 | 213 |
lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)" |
214 |
by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single) |
|
215 |
||
| 59023 | 216 |
lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S" |
217 |
using emeasure_measure_pmf_finite[of S M] |
|
218 |
by(simp add: measure_pmf.emeasure_eq_measure) |
|
219 |
||
| 59000 | 220 |
lemma nn_integral_measure_pmf_support: |
221 |
fixes f :: "'a \<Rightarrow> ereal" |
|
222 |
assumes f: "finite A" and nn: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> set_pmf M \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" |
|
223 |
shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>A. f x * pmf M x)" |
|
224 |
proof - |
|
225 |
have "(\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)" |
|
226 |
using nn by (intro nn_integral_cong_AE) (auto simp: AE_measure_pmf_iff split: split_indicator) |
|
227 |
also have "\<dots> = (\<Sum>x\<in>A. f x * emeasure M {x})"
|
|
228 |
using assms by (intro nn_integral_indicator_finite) auto |
|
229 |
finally show ?thesis |
|
230 |
by (simp add: emeasure_measure_pmf_finite) |
|
231 |
qed |
|
232 |
||
233 |
lemma nn_integral_measure_pmf_finite: |
|
234 |
fixes f :: "'a \<Rightarrow> ereal" |
|
235 |
assumes f: "finite (set_pmf M)" and nn: "\<And>x. x \<in> set_pmf M \<Longrightarrow> 0 \<le> f x" |
|
236 |
shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>set_pmf M. f x * pmf M x)" |
|
237 |
using assms by (intro nn_integral_measure_pmf_support) auto |
|
238 |
lemma integrable_measure_pmf_finite: |
|
239 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
|
|
240 |
shows "finite (set_pmf M) \<Longrightarrow> integrable M f" |
|
241 |
by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite) |
|
242 |
||
243 |
lemma integral_measure_pmf: |
|
244 |
assumes [simp]: "finite A" and "\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A" |
|
245 |
shows "(\<integral>x. f x \<partial>measure_pmf M) = (\<Sum>a\<in>A. f a * pmf M a)" |
|
246 |
proof - |
|
247 |
have "(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x * indicator A x \<partial>measure_pmf M)" |
|
248 |
using assms(2) by (intro integral_cong_AE) (auto split: split_indicator simp: AE_measure_pmf_iff) |
|
249 |
also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)" |
|
250 |
by (subst integral_indicator_finite_real) (auto simp: measure_def emeasure_measure_pmf_finite) |
|
251 |
finally show ?thesis . |
|
252 |
qed |
|
253 |
||
254 |
lemma integrable_pmf: "integrable (count_space X) (pmf M)" |
|
255 |
proof - |
|
256 |
have " (\<integral>\<^sup>+ x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+ x. pmf M x \<partial>count_space (M \<inter> X))" |
|
257 |
by (auto simp add: nn_integral_count_space_indicator set_pmf_iff intro!: nn_integral_cong split: split_indicator) |
|
258 |
then have "integrable (count_space X) (pmf M) = integrable (count_space (M \<inter> X)) (pmf M)" |
|
259 |
by (simp add: integrable_iff_bounded pmf_nonneg) |
|
260 |
then show ?thesis |
|
| 59023 | 261 |
by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def) |
| 59000 | 262 |
qed |
263 |
||
264 |
lemma integral_pmf: "(\<integral>x. pmf M x \<partial>count_space X) = measure M X" |
|
265 |
proof - |
|
266 |
have "(\<integral>x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+x. pmf M x \<partial>count_space X)" |
|
267 |
by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral) |
|
268 |
also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space (X \<inter> M))"
|
|
269 |
by (auto intro!: nn_integral_cong_AE split: split_indicator |
|
270 |
simp: pmf.rep_eq measure_pmf.emeasure_eq_measure nn_integral_count_space_indicator |
|
271 |
AE_count_space set_pmf_iff) |
|
272 |
also have "\<dots> = emeasure M (X \<inter> M)" |
|
273 |
by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf) |
|
274 |
also have "\<dots> = emeasure M X" |
|
275 |
by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff) |
|
276 |
finally show ?thesis |
|
277 |
by (simp add: measure_pmf.emeasure_eq_measure) |
|
278 |
qed |
|
279 |
||
280 |
lemma integral_pmf_restrict: |
|
281 |
"(f::'a \<Rightarrow> 'b::{banach, second_countable_topology}) \<in> borel_measurable (count_space UNIV) \<Longrightarrow>
|
|
282 |
(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x \<partial>restrict_space M M)" |
|
283 |
by (auto intro!: integral_cong_AE simp add: integral_restrict_space AE_measure_pmf_iff) |
|
284 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
285 |
lemma emeasure_pmf: "emeasure (M::'a pmf) M = 1" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
286 |
proof - |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
287 |
have "emeasure (M::'a pmf) M = emeasure (M::'a pmf) (space M)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
288 |
by (intro emeasure_eq_AE) (simp_all add: AE_measure_pmf) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
289 |
then show ?thesis |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
290 |
using measure_pmf.emeasure_space_1 by simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
291 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
292 |
|
| 59023 | 293 |
lemma in_null_sets_measure_pmfI: |
294 |
"A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
|
|
295 |
using emeasure_eq_0_AE[where ?P="\<lambda>x. x \<in> A" and M="measure_pmf p"] |
|
296 |
by(auto simp add: null_sets_def AE_measure_pmf_iff) |
|
297 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
298 |
lemma map_pmf_id[simp]: "map_pmf id = id" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
299 |
by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
300 |
|
| 59053 | 301 |
lemma map_pmf_ident[simp]: "map_pmf (\<lambda>x. x) = (\<lambda>x. x)" |
302 |
using map_pmf_id unfolding id_def . |
|
303 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
304 |
lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
305 |
by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
306 |
|
| 59000 | 307 |
lemma map_pmf_comp: "map_pmf f (map_pmf g M) = map_pmf (\<lambda>x. f (g x)) M" |
308 |
using map_pmf_compose[of f g] by (simp add: comp_def) |
|
309 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
310 |
lemma map_pmf_cong: |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
311 |
assumes "p = q" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
312 |
shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g q" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
313 |
unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] map_pmf.rep_eq |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
314 |
by (auto simp add: emeasure_distr AE_measure_pmf_iff intro!: emeasure_eq_AE measure_eqI) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
315 |
|
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
316 |
lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
317 |
unfolding map_pmf.rep_eq by (subst emeasure_distr) auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
318 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
319 |
lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
320 |
unfolding map_pmf.rep_eq by (intro nn_integral_distr) auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
321 |
|
| 59023 | 322 |
lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
|
323 |
proof(transfer fixing: f x) |
|
324 |
fix p :: "'b measure" |
|
325 |
presume "prob_space p" |
|
326 |
then interpret prob_space p . |
|
327 |
presume "sets p = UNIV" |
|
328 |
then show "ereal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
|
|
329 |
by(simp add: measure_distr measurable_def emeasure_eq_measure) |
|
330 |
qed simp_all |
|
331 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
332 |
lemma pmf_set_map: |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
333 |
fixes f :: "'a \<Rightarrow> 'b" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
334 |
shows "set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
335 |
proof (rule, transfer, clarsimp simp add: measure_distr measurable_def) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
336 |
fix f :: "'a \<Rightarrow> 'b" and M :: "'a measure" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
337 |
assume "prob_space M" and ae: "AE x in M. measure M {x} \<noteq> 0" and [simp]: "sets M = UNIV"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
338 |
interpret prob_space M by fact |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
339 |
show "{x. measure M (f -` {x}) \<noteq> 0} = f ` {x. measure M {x} \<noteq> 0}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
340 |
proof safe |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
341 |
fix x assume "measure M (f -` {x}) \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
342 |
moreover have "measure M (f -` {x}) = measure M {y. f y = x \<and> measure M {y} \<noteq> 0}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
343 |
using ae by (intro finite_measure_eq_AE) auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
344 |
ultimately have "{y. f y = x \<and> measure M {y} \<noteq> 0} \<noteq> {}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
345 |
by (metis measure_empty) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
346 |
then show "x \<in> f ` {x. measure M {x} \<noteq> 0}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
347 |
by auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
348 |
next |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
349 |
fix x assume "measure M {x} \<noteq> 0"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
350 |
then have "0 < measure M {x}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
351 |
using measure_nonneg[of M "{x}"] by auto
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
352 |
also have "measure M {x} \<le> measure M (f -` {f x})"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
353 |
by (intro finite_measure_mono) auto |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
354 |
finally show "measure M (f -` {f x}) = 0 \<Longrightarrow> False"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
355 |
by simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
356 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
357 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
358 |
|
| 59000 | 359 |
lemma set_map_pmf: "set_pmf (map_pmf f M) = f`set_pmf M" |
360 |
using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff) |
|
361 |
||
| 59023 | 362 |
lemma nn_integral_pmf: "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = emeasure (measure_pmf p) A" |
363 |
proof - |
|
364 |
have "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = (\<integral>\<^sup>+ x. pmf p x \<partial>count_space (A \<inter> set_pmf p))" |
|
365 |
by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong) |
|
366 |
also have "\<dots> = emeasure (measure_pmf p) (\<Union>x\<in>A \<inter> set_pmf p. {x})"
|
|
367 |
by(subst emeasure_UN_countable)(auto simp add: emeasure_pmf_single disjoint_family_on_def) |
|
368 |
also have "\<dots> = emeasure (measure_pmf p) ((\<Union>x\<in>A \<inter> set_pmf p. {x}) \<union> {x. x \<in> A \<and> x \<notin> set_pmf p})"
|
|
369 |
by(rule emeasure_Un_null_set[symmetric])(auto intro: in_null_sets_measure_pmfI) |
|
370 |
also have "\<dots> = emeasure (measure_pmf p) A" |
|
371 |
by(auto intro: arg_cong2[where f=emeasure]) |
|
372 |
finally show ?thesis . |
|
373 |
qed |
|
374 |
||
| 59000 | 375 |
subsection {* PMFs as function *}
|
376 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
377 |
context |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
378 |
fixes f :: "'a \<Rightarrow> real" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
379 |
assumes nonneg: "\<And>x. 0 \<le> f x" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
380 |
assumes prob: "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
381 |
begin |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
382 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
383 |
lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ereal \<circ> f)" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
384 |
proof (intro conjI) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
385 |
have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
386 |
by (simp split: split_indicator) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
387 |
show "AE x in density (count_space UNIV) (ereal \<circ> f). |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
388 |
measure (density (count_space UNIV) (ereal \<circ> f)) {x} \<noteq> 0"
|
|
59092
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
59053
diff
changeset
|
389 |
by (simp add: AE_density nonneg measure_def emeasure_density max_def) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
390 |
show "prob_space (density (count_space UNIV) (ereal \<circ> f))" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
391 |
by default (simp add: emeasure_density prob) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
392 |
qed simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
393 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
394 |
lemma pmf_embed_pmf: "pmf embed_pmf x = f x" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
395 |
proof transfer |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
396 |
have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
397 |
by (simp split: split_indicator) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
398 |
fix x show "measure (density (count_space UNIV) (ereal \<circ> f)) {x} = f x"
|
|
59092
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
59053
diff
changeset
|
399 |
by transfer (simp add: measure_def emeasure_density nonneg max_def) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
400 |
qed |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
401 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
402 |
end |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
403 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
404 |
lemma embed_pmf_transfer: |
| 58730 | 405 |
"rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ereal \<circ> f)) embed_pmf" |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
406 |
by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
407 |
|
| 59000 | 408 |
lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)" |
409 |
proof (transfer, elim conjE) |
|
410 |
fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
|
|
411 |
assume "prob_space M" then interpret prob_space M . |
|
412 |
show "M = density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))"
|
|
413 |
proof (rule measure_eqI) |
|
414 |
fix A :: "'a set" |
|
415 |
have "(\<integral>\<^sup>+ x. ereal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
|
|
416 |
(\<integral>\<^sup>+ x. emeasure M {x} * indicator (A \<inter> {x. measure M {x} \<noteq> 0}) x \<partial>count_space UNIV)"
|
|
417 |
by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator) |
|
418 |
also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} \<partial>count_space (A \<inter> {x. measure M {x} \<noteq> 0}))"
|
|
419 |
by (subst nn_integral_restrict_space[symmetric]) (auto simp: restrict_count_space) |
|
420 |
also have "\<dots> = emeasure M (\<Union>x\<in>(A \<inter> {x. measure M {x} \<noteq> 0}). {x})"
|
|
421 |
by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support) |
|
422 |
(auto simp: disjoint_family_on_def) |
|
423 |
also have "\<dots> = emeasure M A" |
|
424 |
using ae by (intro emeasure_eq_AE) auto |
|
425 |
finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))) A"
|
|
426 |
using emeasure_space_1 by (simp add: emeasure_density) |
|
427 |
qed simp |
|
428 |
qed |
|
429 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
430 |
lemma td_pmf_embed_pmf: |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
431 |
"type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1}"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
432 |
unfolding type_definition_def |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
433 |
proof safe |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
434 |
fix p :: "'a pmf" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
435 |
have "(\<integral>\<^sup>+ x. 1 \<partial>measure_pmf p) = 1" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
436 |
using measure_pmf.emeasure_space_1[of p] by simp |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
437 |
then show *: "(\<integral>\<^sup>+ x. ereal (pmf p x) \<partial>count_space UNIV) = 1" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
438 |
by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
439 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
440 |
show "embed_pmf (pmf p) = p" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
441 |
by (intro measure_pmf_inject[THEN iffD1]) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
442 |
(simp add: * embed_pmf.rep_eq pmf_nonneg measure_pmf_eq_density[of p] comp_def) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
443 |
next |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
444 |
fix f :: "'a \<Rightarrow> real" assume "\<forall>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
445 |
then show "pmf (embed_pmf f) = f" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
446 |
by (auto intro!: pmf_embed_pmf) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
447 |
qed (rule pmf_nonneg) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
448 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
449 |
end |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
450 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
451 |
locale pmf_as_function |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
452 |
begin |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
453 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
454 |
setup_lifting td_pmf_embed_pmf |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
455 |
|
| 58730 | 456 |
lemma set_pmf_transfer[transfer_rule]: |
457 |
assumes "bi_total A" |
|
458 |
shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
|
|
459 |
using `bi_total A` |
|
460 |
by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff) |
|
461 |
metis+ |
|
462 |
||
| 59000 | 463 |
end |
464 |
||
465 |
context |
|
466 |
begin |
|
467 |
||
468 |
interpretation pmf_as_function . |
|
469 |
||
470 |
lemma pmf_eqI: "(\<And>i. pmf M i = pmf N i) \<Longrightarrow> M = N" |
|
471 |
by transfer auto |
|
472 |
||
473 |
lemma pmf_eq_iff: "M = N \<longleftrightarrow> (\<forall>i. pmf M i = pmf N i)" |
|
474 |
by (auto intro: pmf_eqI) |
|
475 |
||
476 |
end |
|
477 |
||
478 |
context |
|
479 |
begin |
|
480 |
||
481 |
interpretation pmf_as_function . |
|
482 |
||
| 59093 | 483 |
subsubsection \<open> Bernoulli Distribution \<close> |
484 |
||
| 59000 | 485 |
lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is |
486 |
"\<lambda>p b. ((\<lambda>p. if b then p else 1 - p) \<circ> min 1 \<circ> max 0) p" |
|
487 |
by (auto simp: nn_integral_count_space_finite[where A="{False, True}"] UNIV_bool
|
|
488 |
split: split_max split_min) |
|
489 |
||
490 |
lemma pmf_bernoulli_True[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) True = p" |
|
491 |
by transfer simp |
|
492 |
||
493 |
lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p" |
|
494 |
by transfer simp |
|
495 |
||
496 |
lemma set_pmf_bernoulli: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV" |
|
497 |
by (auto simp add: set_pmf_iff UNIV_bool) |
|
498 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
499 |
lemma nn_integral_bernoulli_pmf[simp]: |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
500 |
assumes [simp]: "0 \<le> p" "p \<le> 1" "\<And>x. 0 \<le> f x" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
501 |
shows "(\<integral>\<^sup>+x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
502 |
by (subst nn_integral_measure_pmf_support[of UNIV]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
503 |
(auto simp: UNIV_bool field_simps) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
504 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
505 |
lemma integral_bernoulli_pmf[simp]: |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
506 |
assumes [simp]: "0 \<le> p" "p \<le> 1" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
507 |
shows "(\<integral>x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
508 |
by (subst integral_measure_pmf[of UNIV]) (auto simp: UNIV_bool) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
509 |
|
| 59093 | 510 |
subsubsection \<open> Geometric Distribution \<close> |
511 |
||
| 59000 | 512 |
lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. 1 / 2^Suc n" |
513 |
proof |
|
514 |
note geometric_sums[of "1 / 2"] |
|
515 |
note sums_mult[OF this, of "1 / 2"] |
|
516 |
from sums_suminf_ereal[OF this] |
|
517 |
show "(\<integral>\<^sup>+ x. ereal (1 / 2 ^ Suc x) \<partial>count_space UNIV) = 1" |
|
518 |
by (simp add: nn_integral_count_space_nat field_simps) |
|
519 |
qed simp |
|
520 |
||
521 |
lemma pmf_geometric[simp]: "pmf geometric_pmf n = 1 / 2^Suc n" |
|
522 |
by transfer rule |
|
523 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
524 |
lemma set_pmf_geometric[simp]: "set_pmf geometric_pmf = UNIV" |
| 59000 | 525 |
by (auto simp: set_pmf_iff) |
526 |
||
| 59093 | 527 |
subsubsection \<open> Uniform Multiset Distribution \<close> |
528 |
||
| 59000 | 529 |
context |
530 |
fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
|
|
531 |
begin |
|
532 |
||
533 |
lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M" |
|
534 |
proof |
|
535 |
show "(\<integral>\<^sup>+ x. ereal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1" |
|
536 |
using M_not_empty |
|
537 |
by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size |
|
538 |
setsum_divide_distrib[symmetric]) |
|
539 |
(auto simp: size_multiset_overloaded_eq intro!: setsum.cong) |
|
540 |
qed simp |
|
541 |
||
542 |
lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M" |
|
543 |
by transfer rule |
|
544 |
||
545 |
lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_of M" |
|
546 |
by (auto simp: set_pmf_iff) |
|
547 |
||
548 |
end |
|
549 |
||
| 59093 | 550 |
subsubsection \<open> Uniform Distribution \<close> |
551 |
||
| 59000 | 552 |
context |
553 |
fixes S :: "'a set" assumes S_not_empty: "S \<noteq> {}" and S_finite: "finite S"
|
|
554 |
begin |
|
555 |
||
556 |
lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S" |
|
557 |
proof |
|
558 |
show "(\<integral>\<^sup>+ x. ereal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1" |
|
559 |
using S_not_empty S_finite by (subst nn_integral_count_space'[of S]) auto |
|
560 |
qed simp |
|
561 |
||
562 |
lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S" |
|
563 |
by transfer rule |
|
564 |
||
565 |
lemma set_pmf_of_set[simp]: "set_pmf pmf_of_set = S" |
|
566 |
using S_finite S_not_empty by (auto simp: set_pmf_iff) |
|
567 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
568 |
lemma emeasure_pmf_of_set[simp]: "emeasure pmf_of_set S = 1" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
569 |
by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
570 |
|
| 59000 | 571 |
end |
572 |
||
| 59093 | 573 |
subsubsection \<open> Poisson Distribution \<close> |
574 |
||
575 |
context |
|
576 |
fixes rate :: real assumes rate_pos: "0 < rate" |
|
577 |
begin |
|
578 |
||
579 |
lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)" |
|
580 |
proof |
|
581 |
(* Proof by Manuel Eberl *) |
|
582 |
||
583 |
have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp |
|
584 |
by (simp add: field_simps field_divide_inverse[symmetric]) |
|
585 |
have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x * exp (-rate) \<partial>count_space UNIV) = |
|
586 |
exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)" |
|
587 |
by (simp add: field_simps nn_integral_cmult[symmetric]) |
|
588 |
also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)" |
|
589 |
by (simp_all add: nn_integral_count_space_nat suminf_ereal summable suminf_ereal_finite) |
|
590 |
also have "... = exp rate" unfolding exp_def |
|
591 |
by (simp add: field_simps field_divide_inverse[symmetric] transfer_int_nat_factorial) |
|
592 |
also have "ereal (exp (-rate)) * ereal (exp rate) = 1" |
|
593 |
by (simp add: mult_exp_exp) |
|
594 |
finally show "(\<integral>\<^sup>+ x. ereal (rate ^ x / real (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" . |
|
595 |
qed (simp add: rate_pos[THEN less_imp_le]) |
|
596 |
||
597 |
lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)" |
|
598 |
by transfer rule |
|
599 |
||
600 |
lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV" |
|
601 |
using rate_pos by (auto simp: set_pmf_iff) |
|
602 |
||
| 59000 | 603 |
end |
604 |
||
| 59093 | 605 |
subsubsection \<open> Binomial Distribution \<close> |
606 |
||
607 |
context |
|
608 |
fixes n :: nat and p :: real assumes p_nonneg: "0 \<le> p" and p_le_1: "p \<le> 1" |
|
609 |
begin |
|
610 |
||
611 |
lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)" |
|
612 |
proof |
|
613 |
have "(\<integral>\<^sup>+k. ereal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) = |
|
614 |
ereal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))" |
|
615 |
using p_le_1 p_nonneg by (subst nn_integral_count_space') auto |
|
616 |
also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n" |
|
617 |
by (subst binomial_ring) (simp add: atLeast0AtMost real_of_nat_def) |
|
618 |
finally show "(\<integral>\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1" |
|
619 |
by simp |
|
620 |
qed (insert p_nonneg p_le_1, simp) |
|
621 |
||
622 |
lemma pmf_binomial[simp]: "pmf binomial_pmf k = (n choose k) * p^k * (1 - p)^(n - k)" |
|
623 |
by transfer rule |
|
624 |
||
625 |
lemma set_pmf_binomial_eq: "set_pmf binomial_pmf = (if p = 0 then {0} else if p = 1 then {n} else {.. n})"
|
|
626 |
using p_nonneg p_le_1 unfolding set_eq_iff set_pmf_iff pmf_binomial by (auto simp: set_pmf_iff) |
|
627 |
||
628 |
end |
|
629 |
||
630 |
end |
|
631 |
||
632 |
lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
|
|
633 |
by (simp add: set_pmf_binomial_eq) |
|
634 |
||
635 |
lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
|
|
636 |
by (simp add: set_pmf_binomial_eq) |
|
637 |
||
638 |
lemma set_pmf_binomial[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (binomial_pmf n p) = {..n}"
|
|
639 |
by (simp add: set_pmf_binomial_eq) |
|
640 |
||
641 |
subsection \<open> Monad Interpretation \<close> |
|
| 59000 | 642 |
|
643 |
lemma measurable_measure_pmf[measurable]: |
|
644 |
"(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))" |
|
645 |
by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales |
|
646 |
||
647 |
lemma bind_pmf_cong: |
|
648 |
assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)" |
|
649 |
assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i" |
|
650 |
shows "bind (measure_pmf x) A = bind (measure_pmf x) B" |
|
651 |
proof (rule measure_eqI) |
|
652 |
show "sets (measure_pmf x \<guillemotright>= A) = sets (measure_pmf x \<guillemotright>= B)" |
|
| 59048 | 653 |
using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra) |
| 59000 | 654 |
next |
655 |
fix X assume "X \<in> sets (measure_pmf x \<guillemotright>= A)" |
|
656 |
then have X: "X \<in> sets N" |
|
| 59048 | 657 |
using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra) |
| 59000 | 658 |
show "emeasure (measure_pmf x \<guillemotright>= A) X = emeasure (measure_pmf x \<guillemotright>= B) X" |
659 |
using assms |
|
660 |
by (subst (1 2) emeasure_bind[where N=N, OF _ _ X]) |
|
661 |
(auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff) |
|
662 |
qed |
|
663 |
||
664 |
context |
|
665 |
begin |
|
666 |
||
667 |
interpretation pmf_as_measure . |
|
668 |
||
669 |
lift_definition join_pmf :: "'a pmf pmf \<Rightarrow> 'a pmf" is "\<lambda>M. measure_pmf M \<guillemotright>= measure_pmf" |
|
670 |
proof (intro conjI) |
|
671 |
fix M :: "'a pmf pmf" |
|
672 |
||
673 |
interpret bind: prob_space "measure_pmf M \<guillemotright>= measure_pmf" |
|
| 59048 | 674 |
apply (intro measure_pmf.prob_space_bind[where S="count_space UNIV"] AE_I2) |
675 |
apply (auto intro!: subprob_space_measure_pmf simp: space_subprob_algebra) |
|
| 59000 | 676 |
apply unfold_locales |
677 |
done |
|
678 |
show "prob_space (measure_pmf M \<guillemotright>= measure_pmf)" |
|
679 |
by intro_locales |
|
680 |
show "sets (measure_pmf M \<guillemotright>= measure_pmf) = UNIV" |
|
| 59048 | 681 |
by (subst sets_bind) auto |
| 59000 | 682 |
have "AE x in measure_pmf M \<guillemotright>= measure_pmf. emeasure (measure_pmf M \<guillemotright>= measure_pmf) {x} \<noteq> 0"
|
| 59048 | 683 |
by (auto simp: AE_bind[where B="count_space UNIV"] measure_pmf_in_subprob_algebra |
684 |
emeasure_bind[where N="count_space UNIV"] AE_measure_pmf_iff nn_integral_0_iff_AE |
|
685 |
measure_pmf.emeasure_eq_measure measure_le_0_iff set_pmf_iff pmf.rep_eq) |
|
| 59000 | 686 |
then show "AE x in measure_pmf M \<guillemotright>= measure_pmf. measure (measure_pmf M \<guillemotright>= measure_pmf) {x} \<noteq> 0"
|
687 |
unfolding bind.emeasure_eq_measure by simp |
|
688 |
qed |
|
689 |
||
690 |
lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)" |
|
691 |
proof (transfer fixing: N i) |
|
692 |
have N: "subprob_space (measure_pmf N)" |
|
693 |
by (rule prob_space_imp_subprob_space) intro_locales |
|
694 |
show "measure (measure_pmf N \<guillemotright>= measure_pmf) {i} = integral\<^sup>L (measure_pmf N) (\<lambda>M. measure M {i})"
|
|
695 |
using measurable_measure_pmf[of "\<lambda>x. x"] |
|
696 |
by (intro subprob_space.measure_bind[where N="count_space UNIV", OF N]) auto |
|
697 |
qed (auto simp: Transfer.Rel_def rel_fun_def cr_pmf_def) |
|
698 |
||
| 59024 | 699 |
lemma set_pmf_join_pmf: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)" |
700 |
apply(simp add: set_eq_iff set_pmf_iff pmf_join) |
|
701 |
apply(subst integral_nonneg_eq_0_iff_AE) |
|
702 |
apply(auto simp add: pmf_le_1 pmf_nonneg AE_measure_pmf_iff intro!: measure_pmf.integrable_const_bound[where B=1]) |
|
703 |
done |
|
704 |
||
| 59000 | 705 |
lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)" |
706 |
by (auto intro!: prob_space_return simp: AE_return measure_return) |
|
707 |
||
708 |
lemma join_return_pmf: "join_pmf (return_pmf M) = M" |
|
709 |
by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq) |
|
710 |
||
711 |
lemma map_return_pmf: "map_pmf f (return_pmf x) = return_pmf (f x)" |
|
712 |
by transfer (simp add: distr_return) |
|
713 |
||
| 59052 | 714 |
lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c" |
715 |
by transfer (auto simp: prob_space.distr_const) |
|
716 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
717 |
lemma set_return_pmf: "set_pmf (return_pmf x) = {x}"
|
| 59000 | 718 |
by transfer (auto simp add: measure_return split: split_indicator) |
719 |
||
720 |
lemma pmf_return: "pmf (return_pmf x) y = indicator {y} x"
|
|
721 |
by transfer (simp add: measure_return) |
|
722 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
723 |
lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
724 |
unfolding return_pmf.rep_eq by (intro nn_integral_return) auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
725 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
726 |
lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
727 |
unfolding return_pmf.rep_eq by (intro emeasure_return) auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
728 |
|
| 59000 | 729 |
end |
730 |
||
731 |
definition "bind_pmf M f = join_pmf (map_pmf f M)" |
|
732 |
||
733 |
lemma (in pmf_as_measure) bind_transfer[transfer_rule]: |
|
734 |
"rel_fun pmf_as_measure.cr_pmf (rel_fun (rel_fun op = pmf_as_measure.cr_pmf) pmf_as_measure.cr_pmf) op \<guillemotright>= bind_pmf" |
|
735 |
proof (auto simp: pmf_as_measure.cr_pmf_def rel_fun_def bind_pmf_def join_pmf.rep_eq map_pmf.rep_eq) |
|
736 |
fix M f and g :: "'a \<Rightarrow> 'b pmf" assume "\<forall>x. f x = measure_pmf (g x)" |
|
737 |
then have f: "f = (\<lambda>x. measure_pmf (g x))" |
|
738 |
by auto |
|
739 |
show "measure_pmf M \<guillemotright>= f = distr (measure_pmf M) (count_space UNIV) g \<guillemotright>= measure_pmf" |
|
740 |
unfolding f by (subst bind_distr[OF _ measurable_measure_pmf]) auto |
|
741 |
qed |
|
742 |
||
743 |
lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)" |
|
744 |
by (auto intro!: integral_distr simp: bind_pmf_def pmf_join map_pmf.rep_eq) |
|
745 |
||
746 |
lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x" |
|
747 |
unfolding bind_pmf_def map_return_pmf join_return_pmf .. |
|
748 |
||
| 59052 | 749 |
lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id" |
750 |
by (simp add: bind_pmf_def) |
|
751 |
||
752 |
lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c" |
|
753 |
unfolding bind_pmf_def map_pmf_const join_return_pmf .. |
|
754 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
755 |
lemma set_bind_pmf: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
756 |
apply (simp add: set_eq_iff set_pmf_iff pmf_bind) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
757 |
apply (subst integral_nonneg_eq_0_iff_AE) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
758 |
apply (auto simp: pmf_nonneg pmf_le_1 AE_measure_pmf_iff |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
759 |
intro!: measure_pmf.integrable_const_bound[where B=1]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
760 |
done |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
761 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
762 |
lemma measurable_pair_restrict_pmf2: |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
763 |
assumes "countable A" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
764 |
assumes "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
765 |
shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
766 |
apply (subst measurable_cong_sets) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
767 |
apply (rule sets_pair_measure_cong sets_restrict_space_cong sets_measure_pmf_count_space refl)+ |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
768 |
apply (simp_all add: restrict_count_space) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
769 |
apply (subst split_eta[symmetric]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
770 |
unfolding measurable_split_conv |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
771 |
apply (rule measurable_compose_countable'[OF _ measurable_snd `countable A`]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
772 |
apply (rule measurable_compose[OF measurable_fst]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
773 |
apply fact |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
774 |
done |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
775 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
776 |
lemma measurable_pair_restrict_pmf1: |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
777 |
assumes "countable A" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
778 |
assumes "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
779 |
shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
780 |
apply (subst measurable_cong_sets) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
781 |
apply (rule sets_pair_measure_cong sets_restrict_space_cong sets_measure_pmf_count_space refl)+ |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
782 |
apply (simp_all add: restrict_count_space) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
783 |
apply (subst split_eta[symmetric]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
784 |
unfolding measurable_split_conv |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
785 |
apply (rule measurable_compose_countable'[OF _ measurable_fst `countable A`]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
786 |
apply (rule measurable_compose[OF measurable_snd]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
787 |
apply fact |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
788 |
done |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
789 |
|
| 59000 | 790 |
lemma bind_commute_pmf: "bind_pmf A (\<lambda>x. bind_pmf B (C x)) = bind_pmf B (\<lambda>y. bind_pmf A (\<lambda>x. C x y))" |
791 |
unfolding pmf_eq_iff pmf_bind |
|
792 |
proof |
|
793 |
fix i |
|
794 |
interpret B: prob_space "restrict_space B B" |
|
795 |
by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE) |
|
796 |
(auto simp: AE_measure_pmf_iff) |
|
797 |
interpret A: prob_space "restrict_space A A" |
|
798 |
by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE) |
|
799 |
(auto simp: AE_measure_pmf_iff) |
|
800 |
||
801 |
interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B" |
|
802 |
by unfold_locales |
|
803 |
||
804 |
have "(\<integral> x. \<integral> y. pmf (C x y) i \<partial>B \<partial>A) = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>A)" |
|
805 |
by (rule integral_cong) (auto intro!: integral_pmf_restrict) |
|
806 |
also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)" |
|
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
807 |
by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2 |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
808 |
countable_set_pmf borel_measurable_count_space) |
| 59000 | 809 |
also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)" |
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
810 |
by (rule AB.Fubini_integral[symmetric]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
811 |
(auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2 |
| 59023 | 812 |
simp: pmf_nonneg pmf_le_1 measurable_restrict_space1) |
| 59000 | 813 |
also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)" |
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
814 |
by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2 |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
815 |
countable_set_pmf borel_measurable_count_space) |
| 59000 | 816 |
also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)" |
817 |
by (rule integral_cong) (auto intro!: integral_pmf_restrict[symmetric]) |
|
818 |
finally show "(\<integral> x. \<integral> y. pmf (C x y) i \<partial>B \<partial>A) = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)" . |
|
819 |
qed |
|
820 |
||
821 |
||
822 |
context |
|
823 |
begin |
|
824 |
||
825 |
interpretation pmf_as_measure . |
|
826 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
827 |
lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<guillemotright>= (\<lambda>x. measure_pmf (f x)))" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
828 |
by transfer simp |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
829 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
830 |
lemma nn_integral_bind_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>bind_pmf M N) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
831 |
using measurable_measure_pmf[of N] |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
832 |
unfolding measure_pmf_bind |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
833 |
apply (subst (1 3) nn_integral_max_0[symmetric]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
834 |
apply (intro nn_integral_bind[where B="count_space UNIV"]) |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
835 |
apply auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
836 |
done |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
837 |
|
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
838 |
lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
839 |
using measurable_measure_pmf[of N] |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
840 |
unfolding measure_pmf_bind |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
841 |
by (subst emeasure_bind[where N="count_space UNIV"]) auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
842 |
|
| 59000 | 843 |
lemma bind_return_pmf': "bind_pmf N return_pmf = N" |
844 |
proof (transfer, clarify) |
|
845 |
fix N :: "'a measure" assume "sets N = UNIV" then show "N \<guillemotright>= return (count_space UNIV) = N" |
|
846 |
by (subst return_sets_cong[where N=N]) (simp_all add: bind_return') |
|
847 |
qed |
|
848 |
||
849 |
lemma bind_return_pmf'': "bind_pmf N (\<lambda>x. return_pmf (f x)) = map_pmf f N" |
|
850 |
proof (transfer, clarify) |
|
851 |
fix N :: "'b measure" and f :: "'b \<Rightarrow> 'a" assume "prob_space N" "sets N = UNIV" |
|
852 |
then show "N \<guillemotright>= (\<lambda>x. return (count_space UNIV) (f x)) = distr N (count_space UNIV) f" |
|
853 |
by (subst bind_return_distr[symmetric]) |
|
854 |
(auto simp: prob_space.not_empty measurable_def comp_def) |
|
855 |
qed |
|
856 |
||
857 |
lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)" |
|
858 |
by transfer |
|
859 |
(auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"] |
|
860 |
simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space) |
|
861 |
||
862 |
end |
|
863 |
||
| 59052 | 864 |
lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)" |
865 |
unfolding bind_pmf_def[symmetric] |
|
866 |
unfolding bind_return_pmf''[symmetric] join_eq_bind_pmf bind_assoc_pmf |
|
867 |
by (simp add: bind_return_pmf'') |
|
868 |
||
| 59000 | 869 |
definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))" |
870 |
||
871 |
lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b" |
|
872 |
unfolding pair_pmf_def pmf_bind pmf_return |
|
873 |
apply (subst integral_measure_pmf[where A="{b}"])
|
|
874 |
apply (auto simp: indicator_eq_0_iff) |
|
875 |
apply (subst integral_measure_pmf[where A="{a}"])
|
|
876 |
apply (auto simp: indicator_eq_0_iff setsum_nonneg_eq_0_iff pmf_nonneg) |
|
877 |
done |
|
878 |
||
|
59002
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
879 |
lemma set_pair_pmf: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B" |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
880 |
unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto |
|
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents:
59000
diff
changeset
|
881 |
|
| 59048 | 882 |
lemma measure_pmf_in_subprob_space[measurable (raw)]: |
883 |
"measure_pmf M \<in> space (subprob_algebra (count_space UNIV))" |
|
884 |
by (simp add: space_subprob_algebra) intro_locales |
|
885 |
||
| 59134 | 886 |
lemma nn_integral_pair_pmf': "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)" |
887 |
proof - |
|
888 |
have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. max 0 (f x) * indicator (A \<times> B) x \<partial>pair_pmf A B)" |
|
889 |
by (subst nn_integral_max_0[symmetric]) |
|
890 |
(auto simp: AE_measure_pmf_iff set_pair_pmf intro!: nn_integral_cong_AE) |
|
891 |
also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)" |
|
892 |
by (simp add: pair_pmf_def) |
|
893 |
also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) \<partial>B \<partial>A)" |
|
894 |
by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff) |
|
895 |
finally show ?thesis |
|
896 |
unfolding nn_integral_max_0 . |
|
897 |
qed |
|
898 |
||
899 |
lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)" |
|
900 |
proof (safe intro!: pmf_eqI) |
|
901 |
fix a :: "'a" and b :: "'b" |
|
902 |
have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ereal)"
|
|
903 |
by (auto split: split_indicator) |
|
904 |
||
905 |
have "ereal (pmf (pair_pmf (map_pmf f A) B) (a, b)) = |
|
906 |
ereal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))" |
|
907 |
unfolding pmf_pair ereal_pmf_map |
|
908 |
by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg |
|
909 |
emeasure_map_pmf[symmetric] del: emeasure_map_pmf) |
|
910 |
then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)" |
|
911 |
by simp |
|
912 |
qed |
|
913 |
||
914 |
lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)" |
|
915 |
proof (safe intro!: pmf_eqI) |
|
916 |
fix a :: "'a" and b :: "'b" |
|
917 |
have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ereal)"
|
|
918 |
by (auto split: split_indicator) |
|
919 |
||
920 |
have "ereal (pmf (pair_pmf A (map_pmf f B)) (a, b)) = |
|
921 |
ereal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))" |
|
922 |
unfolding pmf_pair ereal_pmf_map |
|
923 |
by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg |
|
924 |
emeasure_map_pmf[symmetric] del: emeasure_map_pmf) |
|
925 |
then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)" |
|
926 |
by simp |
|
927 |
qed |
|
928 |
||
929 |
lemma map_pair: "map_pmf (\<lambda>(a, b). (f a, g b)) (pair_pmf A B) = pair_pmf (map_pmf f A) (map_pmf g B)" |
|
930 |
by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta') |
|
931 |
||
| 59000 | 932 |
lemma bind_pair_pmf: |
933 |
assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)" |
|
934 |
shows "measure_pmf (pair_pmf A B) \<guillemotright>= M = (measure_pmf A \<guillemotright>= (\<lambda>x. measure_pmf B \<guillemotright>= (\<lambda>y. M (x, y))))" |
|
935 |
(is "?L = ?R") |
|
936 |
proof (rule measure_eqI) |
|
937 |
have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)" |
|
938 |
using M[THEN measurable_space] by (simp_all add: space_pair_measure) |
|
939 |
||
| 59048 | 940 |
note measurable_bind[where N="count_space UNIV", measurable] |
941 |
note measure_pmf_in_subprob_space[simp] |
|
942 |
||
| 59000 | 943 |
have sets_eq_N: "sets ?L = N" |
| 59048 | 944 |
by (subst sets_bind[OF sets_kernel[OF M']]) auto |
| 59000 | 945 |
show "sets ?L = sets ?R" |
| 59048 | 946 |
using measurable_space[OF M] |
947 |
by (simp add: sets_eq_N space_pair_measure space_subprob_algebra) |
|
| 59000 | 948 |
fix X assume "X \<in> sets ?L" |
949 |
then have X[measurable]: "X \<in> sets N" |
|
950 |
unfolding sets_eq_N . |
|
951 |
then show "emeasure ?L X = emeasure ?R X" |
|
952 |
apply (simp add: emeasure_bind[OF _ M' X]) |
|
| 59048 | 953 |
apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A] |
954 |
nn_integral_measure_pmf_finite set_return_pmf emeasure_nonneg pmf_return one_ereal_def[symmetric]) |
|
955 |
apply (subst emeasure_bind[OF _ _ X]) |
|
| 59000 | 956 |
apply measurable |
957 |
apply (subst emeasure_bind[OF _ _ X]) |
|
958 |
apply measurable |
|
959 |
done |
|
960 |
qed |
|
961 |
||
| 59052 | 962 |
lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A" |
963 |
unfolding bind_pmf_def[symmetric] bind_return_pmf' .. |
|
964 |
||
965 |
lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A" |
|
966 |
by (simp add: pair_pmf_def bind_return_pmf''[symmetric] bind_assoc_pmf bind_return_pmf bind_return_pmf') |
|
967 |
||
968 |
lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B" |
|
969 |
by (simp add: pair_pmf_def bind_return_pmf''[symmetric] bind_assoc_pmf bind_return_pmf bind_return_pmf') |
|
970 |
||
| 59053 | 971 |
lemma nn_integral_pmf': |
972 |
"inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)" |
|
973 |
by (subst nn_integral_bij_count_space[where g=f and B="f`A"]) |
|
974 |
(auto simp: bij_betw_def nn_integral_pmf) |
|
975 |
||
976 |
lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0" |
|
977 |
using pmf_nonneg[of M p] by simp |
|
978 |
||
979 |
lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0" |
|
980 |
using pmf_nonneg[of M p] by simp_all |
|
981 |
||
982 |
lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M" |
|
983 |
unfolding set_pmf_iff by simp |
|
984 |
||
985 |
lemma pmf_map_inj: "inj_on f (set_pmf M) \<Longrightarrow> x \<in> set_pmf M \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x" |
|
986 |
by (auto simp: pmf.rep_eq map_pmf.rep_eq measure_distr AE_measure_pmf_iff inj_onD |
|
987 |
intro!: measure_pmf.finite_measure_eq_AE) |
|
988 |
||
| 59023 | 989 |
inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
|
990 |
for R p q |
|
991 |
where |
|
992 |
"\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y; |
|
993 |
map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk> |
|
994 |
\<Longrightarrow> rel_pmf R p q" |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
995 |
|
| 59023 | 996 |
bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
997 |
proof - |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
998 |
show "map_pmf id = id" by (rule map_pmf_id) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
999 |
show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1000 |
show "\<And>f g::'a \<Rightarrow> 'b. \<And>p. (\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g p" |
| 59023 | 1001 |
by (intro map_pmf_cong refl) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1002 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1003 |
show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1004 |
by (rule pmf_set_map) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1005 |
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1006 |
{ fix p :: "'s pmf"
|
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1007 |
have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1008 |
by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"]) |
| 59053 | 1009 |
(auto intro: countable_set_pmf) |
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1010 |
also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq" |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1011 |
by (metis Field_natLeq card_of_least natLeq_Well_order) |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1012 |
finally show "(card_of (set_pmf p), natLeq) \<in> ordLeq" . } |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1013 |
|
| 59023 | 1014 |
show "\<And>R. rel_pmf R = |
1015 |
(BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf fst))\<inverse>\<inverse> OO
|
|
1016 |
BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf snd)"
|
|
1017 |
by (auto simp add: fun_eq_iff BNF_Def.Grp_def OO_def rel_pmf.simps) |
|
1018 |
||
1019 |
{ fix p :: "'a pmf" and f :: "'a \<Rightarrow> 'b" and g x
|
|
1020 |
assume p: "\<And>z. z \<in> set_pmf p \<Longrightarrow> f z = g z" |
|
1021 |
and x: "x \<in> set_pmf p" |
|
1022 |
thus "f x = g x" by simp } |
|
1023 |
||
1024 |
fix R :: "'a => 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool" |
|
1025 |
{ fix p q r
|
|
1026 |
assume pq: "rel_pmf R p q" |
|
1027 |
and qr:"rel_pmf S q r" |
|
1028 |
from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" |
|
1029 |
and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto |
|
1030 |
from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z" |
|
1031 |
and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto |
|
1032 |
||
| 59053 | 1033 |
note pmf_nonneg[intro, simp] |
| 59023 | 1034 |
|
| 59053 | 1035 |
def A \<equiv> "\<lambda>y. {x. (x, y) \<in> set_pmf pq}"
|
1036 |
then have "\<And>y. A y \<subseteq> set_pmf p" by (auto simp add: p set_map_pmf intro: rev_image_eqI) |
|
1037 |
then have [simp]: "\<And>y. countable (A y)" by (rule countable_subset) simp |
|
1038 |
have A: "\<And>x y. (x, y) \<in> set_pmf pq \<longleftrightarrow> x \<in> A y" |
|
1039 |
by (simp add: A_def) |
|
| 59023 | 1040 |
|
| 59053 | 1041 |
let ?P = "\<lambda>y. to_nat_on (A y)" |
1042 |
def pp \<equiv> "map_pmf (\<lambda>(x, y). (y, ?P y x)) pq" |
|
1043 |
let ?pp = "\<lambda>y x. pmf pp (y, x)" |
|
1044 |
{ fix x y have "x \<in> A y \<Longrightarrow> pmf pp (y, ?P y x) = pmf pq (x, y)"
|
|
1045 |
unfolding pp_def |
|
1046 |
by (intro pmf_map_inj[of "\<lambda>(x, y). (y, ?P y x)" pq "(x, y)", simplified]) |
|
1047 |
(auto simp: inj_on_def A) } |
|
1048 |
note pmf_pp = this |
|
| 59023 | 1049 |
|
| 59053 | 1050 |
def B \<equiv> "\<lambda>y. {z. (y, z) \<in> set_pmf qr}"
|
1051 |
then have "\<And>y. B y \<subseteq> set_pmf r" by (auto simp add: r set_map_pmf intro: rev_image_eqI) |
|
1052 |
then have [simp]: "\<And>y. countable (B y)" by (rule countable_subset) simp |
|
1053 |
have B: "\<And>y z. (y, z) \<in> set_pmf qr \<longleftrightarrow> z \<in> B y" |
|
1054 |
by (simp add: B_def) |
|
| 59023 | 1055 |
|
| 59053 | 1056 |
let ?R = "\<lambda>y. to_nat_on (B y)" |
1057 |
def rr \<equiv> "map_pmf (\<lambda>(y, z). (y, ?R y z)) qr" |
|
1058 |
let ?rr = "\<lambda>y z. pmf rr (y, z)" |
|
1059 |
{ fix y z have "z \<in> B y \<Longrightarrow> pmf rr (y, ?R y z) = pmf qr (y, z)"
|
|
1060 |
unfolding rr_def |
|
1061 |
by (intro pmf_map_inj[of "\<lambda>(y, z). (y, ?R y z)" qr "(y, z)", simplified]) |
|
1062 |
(auto simp: inj_on_def B) } |
|
1063 |
note pmf_rr = this |
|
| 59023 | 1064 |
|
| 59053 | 1065 |
have eq: "\<And>y. (\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV)" |
| 59023 | 1066 |
proof - |
1067 |
fix y |
|
| 59053 | 1068 |
have "(\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = pmf q y" |
1069 |
by (simp add: nn_integral_pmf' inj_on_def pp_def q) |
|
1070 |
(auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
|
1071 |
also have "\<dots> = (\<integral>\<^sup>+ x. ?rr y x \<partial>count_space UNIV)" |
|
1072 |
by (simp add: nn_integral_pmf' inj_on_def rr_def q') |
|
1073 |
(auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
|
| 59023 | 1074 |
finally show "?thesis y" . |
1075 |
qed |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1076 |
|
| 59023 | 1077 |
def assign_aux \<equiv> "\<lambda>y remainder start weight z. |
1078 |
if z < start then 0 |
|
1079 |
else if z = start then min weight remainder |
|
| 59053 | 1080 |
else if remainder + setsum (?rr y) {Suc start ..<z} < weight then min (weight - remainder - setsum (?rr y) {Suc start..<z}) (?rr y z) else 0"
|
| 59023 | 1081 |
hence assign_aux_alt_def: "\<And>y remainder start weight z. assign_aux y remainder start weight z = |
1082 |
(if z < start then 0 |
|
1083 |
else if z = start then min weight remainder |
|
| 59053 | 1084 |
else if remainder + setsum (?rr y) {Suc start ..<z} < weight then min (weight - remainder - setsum (?rr y) {Suc start..<z}) (?rr y z) else 0)"
|
| 59023 | 1085 |
by simp |
1086 |
{ fix y and remainder :: real and start and weight :: real
|
|
1087 |
assume weight_nonneg: "0 \<le> weight" |
|
1088 |
let ?assign_aux = "assign_aux y remainder start weight" |
|
1089 |
{ fix z
|
|
1090 |
have "setsum ?assign_aux {..<z} =
|
|
| 59053 | 1091 |
(if z \<le> start then 0 else if remainder + setsum (?rr y) {Suc start..<z} < weight then remainder + setsum (?rr y) {Suc start..<z} else weight)"
|
| 59023 | 1092 |
proof(induction z) |
1093 |
case (Suc z) show ?case |
|
| 59053 | 1094 |
by (auto simp add: Suc.IH assign_aux_alt_def[where z=z] not_less) |
1095 |
(metis add.commute add.left_commute add_increasing pmf_nonneg) |
|
| 59023 | 1096 |
qed(auto simp add: assign_aux_def) } |
1097 |
note setsum_start_assign_aux = this |
|
1098 |
moreover {
|
|
1099 |
assume remainder_nonneg: "0 \<le> remainder" |
|
1100 |
have [simp]: "\<And>z. 0 \<le> ?assign_aux z" |
|
1101 |
by(simp add: assign_aux_def weight_nonneg remainder_nonneg) |
|
| 59053 | 1102 |
moreover have "\<And>z. \<lbrakk> ?rr y z = 0; remainder \<le> ?rr y start \<rbrakk> \<Longrightarrow> ?assign_aux z = 0" |
| 59023 | 1103 |
using remainder_nonneg weight_nonneg |
1104 |
by(auto simp add: assign_aux_def min_def) |
|
1105 |
moreover have "(\<integral>\<^sup>+ z. ?assign_aux z \<partial>count_space UNIV) = |
|
| 59053 | 1106 |
min weight (\<integral>\<^sup>+ z. (if z < start then 0 else if z = start then remainder else ?rr y z) \<partial>count_space UNIV)" |
| 59023 | 1107 |
(is "?lhs = ?rhs" is "_ = min _ (\<integral>\<^sup>+ y. ?f y \<partial>_)") |
1108 |
proof - |
|
1109 |
have "?lhs = (SUP n. \<Sum>z<n. ereal (?assign_aux z))" |
|
1110 |
by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
|
1111 |
also have "\<dots> = (SUP n. min weight (\<Sum>z<n. ?f z))" |
|
1112 |
proof(rule arg_cong2[where f=SUPREMUM] ext refl)+ |
|
1113 |
fix n |
|
1114 |
have "(\<Sum>z<n. ereal (?assign_aux z)) = min weight ((if n > start then remainder else 0) + setsum ?f {Suc start..<n})"
|
|
1115 |
using weight_nonneg remainder_nonneg by(simp add: setsum_start_assign_aux min_def) |
|
1116 |
also have "\<dots> = min weight (setsum ?f {start..<n})"
|
|
1117 |
by(simp add: setsum_head_upt_Suc) |
|
1118 |
also have "\<dots> = min weight (setsum ?f {..<n})"
|
|
1119 |
by(intro arg_cong2[where f=min] setsum.mono_neutral_left) auto |
|
1120 |
finally show "(\<Sum>z<n. ereal (?assign_aux z)) = \<dots>" . |
|
1121 |
qed |
|
1122 |
also have "\<dots> = min weight (SUP n. setsum ?f {..<n})"
|
|
1123 |
unfolding inf_min[symmetric] by(subst inf_SUP) simp |
|
1124 |
also have "\<dots> = ?rhs" |
|
1125 |
by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP remainder_nonneg) |
|
1126 |
finally show ?thesis . |
|
1127 |
qed |
|
1128 |
moreover note calculation } |
|
1129 |
moreover note calculation } |
|
1130 |
note setsum_start_assign_aux = this(1) |
|
1131 |
and assign_aux_nonneg [simp] = this(2) |
|
1132 |
and assign_aux_eq_0_outside = this(3) |
|
1133 |
and nn_integral_assign_aux = this(4) |
|
1134 |
{ fix y and remainder :: real and start target
|
|
| 59053 | 1135 |
have "setsum (?rr y) {Suc start..<target} \<ge> 0" by (simp add: setsum_nonneg)
|
| 59023 | 1136 |
moreover assume "0 \<le> remainder" |
1137 |
ultimately have "assign_aux y remainder start 0 target = 0" |
|
1138 |
by(auto simp add: assign_aux_def min_def) } |
|
1139 |
note assign_aux_weight_0 [simp] = this |
|
1140 |
||
| 59053 | 1141 |
def find_start \<equiv> "\<lambda>y weight. if \<exists>n. weight \<le> setsum (?rr y) {..n} then Some (LEAST n. weight \<le> setsum (?rr y) {..n}) else None"
|
| 59023 | 1142 |
have find_start_eq_Some_above: |
| 59053 | 1143 |
"\<And>y weight n. find_start y weight = Some n \<Longrightarrow> weight \<le> setsum (?rr y) {..n}"
|
| 59023 | 1144 |
by(drule sym)(auto simp add: find_start_def split: split_if_asm intro: LeastI) |
1145 |
{ fix y weight n
|
|
1146 |
assume find_start: "find_start y weight = Some n" |
|
1147 |
and weight: "0 \<le> weight" |
|
| 59053 | 1148 |
have "setsum (?rr y) {..n} \<le> ?rr y n + weight"
|
| 59023 | 1149 |
proof(rule ccontr) |
1150 |
assume "\<not> ?thesis" |
|
| 59053 | 1151 |
hence "?rr y n + weight < setsum (?rr y) {..n}" by simp
|
| 59023 | 1152 |
moreover with weight obtain n' where "n = Suc n'" by(cases n) auto |
| 59053 | 1153 |
ultimately have "weight \<le> setsum (?rr y) {..n'}" by simp
|
1154 |
hence "(LEAST n. weight \<le> setsum (?rr y) {..n}) \<le> n'" by(rule Least_le)
|
|
1155 |
moreover from find_start have "n = (LEAST n. weight \<le> setsum (?rr y) {..n})"
|
|
| 59023 | 1156 |
by(auto simp add: find_start_def split: split_if_asm) |
1157 |
ultimately show False using \<open>n = Suc n'\<close> by auto |
|
1158 |
qed } |
|
1159 |
note find_start_eq_Some_least = this |
|
1160 |
have find_start_0 [simp]: "\<And>y. find_start y 0 = Some 0" |
|
1161 |
by(auto simp add: find_start_def intro!: exI[where x=0]) |
|
1162 |
{ fix y and weight :: real
|
|
| 59053 | 1163 |
assume "weight < \<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV" |
1164 |
also have "(\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV) = (SUP n. \<Sum>z<n. ereal (?rr y z))" |
|
| 59023 | 1165 |
by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
| 59053 | 1166 |
finally obtain n where "weight < (\<Sum>z<n. ?rr y z)" by(auto simp add: less_SUP_iff) |
| 59023 | 1167 |
hence "weight \<in> dom (find_start y)" |
| 59053 | 1168 |
by(auto simp add: find_start_def)(meson atMost_iff finite_atMost lessThan_iff less_imp_le order_trans pmf_nonneg setsum_mono3 subsetI) } |
| 59023 | 1169 |
note in_dom_find_startI = this |
1170 |
{ fix y and w w' :: real and m
|
|
| 59053 | 1171 |
let ?m' = "LEAST m. w' \<le> setsum (?rr y) {..m}"
|
| 59023 | 1172 |
assume "w' \<le> w" |
1173 |
also assume "find_start y w = Some m" |
|
| 59053 | 1174 |
hence "w \<le> setsum (?rr y) {..m}" by(rule find_start_eq_Some_above)
|
| 59023 | 1175 |
finally have "find_start y w' = Some ?m'" by(auto simp add: find_start_def) |
| 59053 | 1176 |
moreover from \<open>w' \<le> setsum (?rr y) {..m}\<close> have "?m' \<le> m" by(rule Least_le)
|
| 59023 | 1177 |
ultimately have "\<exists>m'. find_start y w' = Some m' \<and> m' \<le> m" by blast } |
1178 |
note find_start_mono = this[rotated] |
|
1179 |
||
| 59053 | 1180 |
def assign \<equiv> "\<lambda>y x z. let used = setsum (?pp y) {..<x}
|
| 59023 | 1181 |
in case find_start y used of None \<Rightarrow> 0 |
| 59053 | 1182 |
| Some start \<Rightarrow> assign_aux y (setsum (?rr y) {..start} - used) start (?pp y x) z"
|
| 59023 | 1183 |
hence assign_alt_def: "\<And>y x z. assign y x z = |
| 59053 | 1184 |
(let used = setsum (?pp y) {..<x}
|
| 59023 | 1185 |
in case find_start y used of None \<Rightarrow> 0 |
| 59053 | 1186 |
| Some start \<Rightarrow> assign_aux y (setsum (?rr y) {..start} - used) start (?pp y x) z)"
|
| 59023 | 1187 |
by simp |
1188 |
have assign_nonneg [simp]: "\<And>y x z. 0 \<le> assign y x z" |
|
| 59053 | 1189 |
by(simp add: assign_def diff_le_iff find_start_eq_Some_above Let_def split: option.split) |
1190 |
have assign_eq_0_outside: "\<And>y x z. \<lbrakk> ?pp y x = 0 \<or> ?rr y z = 0 \<rbrakk> \<Longrightarrow> assign y x z = 0" |
|
1191 |
by(auto simp add: assign_def assign_aux_eq_0_outside diff_le_iff find_start_eq_Some_above find_start_eq_Some_least setsum_nonneg Let_def split: option.split) |
|
| 59023 | 1192 |
|
1193 |
{ fix y x z
|
|
1194 |
have "(\<Sum>n<Suc x. assign y n z) = |
|
| 59053 | 1195 |
(case find_start y (setsum (?pp y) {..<x}) of None \<Rightarrow> ?rr y z
|
1196 |
| Some m \<Rightarrow> if z < m then ?rr y z |
|
1197 |
else min (?rr y z) (max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z})))"
|
|
| 59023 | 1198 |
(is "?lhs x = ?rhs x") |
1199 |
proof(induction x) |
|
1200 |
case 0 thus ?case |
|
1201 |
by(auto simp add: assign_def assign_aux_def setsum_head_upt_Suc atLeast0LessThan[symmetric] not_less field_simps max_def) |
|
1202 |
next |
|
1203 |
case (Suc x) |
|
1204 |
have "?lhs (Suc x) = ?lhs x + assign y (Suc x) z" by simp |
|
1205 |
also have "?lhs x = ?rhs x" by(rule Suc.IH) |
|
1206 |
also have "?rhs x + assign y (Suc x) z = ?rhs (Suc x)" |
|
| 59053 | 1207 |
proof(cases "find_start y (setsum (?pp y) {..<Suc x})")
|
| 59023 | 1208 |
case None |
1209 |
thus ?thesis |
|
1210 |
by(auto split: option.split simp add: assign_def min_def max_def diff_le_iff setsum_nonneg not_le field_simps) |
|
1211 |
(metis add.commute add_increasing find_start_def lessThan_Suc_atMost less_imp_le option.distinct(1) setsum_lessThan_Suc)+ |
|
| 59053 | 1212 |
next |
| 59023 | 1213 |
case (Some m) |
| 59053 | 1214 |
have [simp]: "setsum (?rr y) {..m} = ?rr y m + setsum (?rr y) {..<m}"
|
| 59023 | 1215 |
by(simp add: ivl_disj_un(2)[symmetric]) |
| 59053 | 1216 |
from Some obtain m' where m': "find_start y (setsum (?pp y) {..<x}) = Some m'" "m' \<le> m"
|
1217 |
by(auto dest: find_start_mono[where w'2="setsum (?pp y) {..<x}"])
|
|
| 59023 | 1218 |
moreover {
|
1219 |
assume "z < m" |
|
| 59053 | 1220 |
then have "setsum (?rr y) {..z} \<le> setsum (?rr y) {..<m}"
|
| 59023 | 1221 |
by(auto intro: setsum_mono3) |
| 59053 | 1222 |
also have "\<dots> \<le> setsum (?pp y) {..<Suc x}" using find_start_eq_Some_least[OF Some]
|
| 59023 | 1223 |
by(simp add: ivl_disj_un(2)[symmetric] setsum_nonneg) |
| 59053 | 1224 |
finally have "?rr y z \<le> max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z})"
|
1225 |
by(auto simp add: ivl_disj_un(2)[symmetric] max_def diff_le_iff simp del: pmf_le_0_iff) |
|
| 59023 | 1226 |
} moreover {
|
1227 |
assume "m \<le> z" |
|
| 59053 | 1228 |
have "setsum (?pp y) {..<Suc x} \<le> setsum (?rr y) {..m}"
|
| 59023 | 1229 |
using Some by(rule find_start_eq_Some_above) |
| 59053 | 1230 |
also have "\<dots> \<le> setsum (?rr y) {..<Suc z}" using \<open>m \<le> z\<close> by(intro setsum_mono3) auto
|
1231 |
finally have "max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z}) \<le> ?rr y z" by simp
|
|
1232 |
moreover have "z \<noteq> m \<Longrightarrow> setsum (?rr y) {..m} + setsum (?rr y) {Suc m..<z} = setsum (?rr y) {..<z}"
|
|
| 59023 | 1233 |
using \<open>m \<le> z\<close> |
1234 |
by(subst ivl_disj_un(8)[where l="Suc m", symmetric]) |
|
1235 |
(simp_all add: setsum_Un ivl_disj_un(2)[symmetric] setsum.neutral) |
|
1236 |
moreover note calculation |
|
1237 |
} moreover {
|
|
1238 |
assume "m < z" |
|
| 59053 | 1239 |
have "setsum (?pp y) {..<Suc x} \<le> setsum (?rr y) {..m}"
|
| 59023 | 1240 |
using Some by(rule find_start_eq_Some_above) |
| 59053 | 1241 |
also have "\<dots> \<le> setsum (?rr y) {..<z}" using \<open>m < z\<close> by(intro setsum_mono3) auto
|
1242 |
finally have "max 0 (setsum (?pp y) {..<Suc x} - setsum (?rr y) {..<z}) = 0" by simp }
|
|
1243 |
moreover have "setsum (?pp y) {..<Suc x} \<ge> setsum (?rr y) {..<m}"
|
|
| 59023 | 1244 |
using find_start_eq_Some_least[OF Some] |
1245 |
by(simp add: setsum_nonneg ivl_disj_un(2)[symmetric]) |
|
| 59053 | 1246 |
moreover hence "setsum (?pp y) {..<Suc (Suc x)} \<ge> setsum (?rr y) {..<m}"
|
| 59023 | 1247 |
by(fastforce intro: order_trans) |
1248 |
ultimately show ?thesis using Some |
|
1249 |
by(auto simp add: assign_def assign_aux_def Let_def field_simps max_def) |
|
1250 |
qed |
|
1251 |
finally show ?case . |
|
1252 |
qed } |
|
1253 |
note setsum_assign = this |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1254 |
|
| 59053 | 1255 |
have nn_integral_assign1: "\<And>y z. (\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = ?rr y z" |
| 59023 | 1256 |
proof - |
1257 |
fix y z |
|
1258 |
have "(\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = (SUP n. ereal (\<Sum>x<n. assign y x z))" |
|
1259 |
by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
|
| 59053 | 1260 |
also have "\<dots> = ?rr y z" |
| 59023 | 1261 |
proof(rule antisym) |
| 59053 | 1262 |
show "(SUP n. ereal (\<Sum>x<n. assign y x z)) \<le> ?rr y z" |
| 59023 | 1263 |
proof(rule SUP_least) |
1264 |
fix n |
|
| 59053 | 1265 |
show "ereal (\<Sum>x<n. (assign y x z)) \<le> ?rr y z" |
| 59023 | 1266 |
using setsum_assign[of y z "n - 1"] |
1267 |
by(cases n)(simp_all split: option.split) |
|
1268 |
qed |
|
| 59053 | 1269 |
show "?rr y z \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
1270 |
proof(cases "setsum (?rr y) {..z} < \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV")
|
|
| 59023 | 1271 |
case True |
| 59053 | 1272 |
then obtain n where "setsum (?rr y) {..z} < setsum (?pp y) {..<n}"
|
| 59023 | 1273 |
by(auto simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP less_SUP_iff) |
| 59053 | 1274 |
moreover have "\<And>k. k < z \<Longrightarrow> setsum (?rr y) {..k} \<le> setsum (?rr y) {..<z}"
|
| 59023 | 1275 |
by(auto intro: setsum_mono3) |
| 59053 | 1276 |
ultimately have "?rr y z \<le> (\<Sum>x<Suc n. assign y x z)" |
| 59023 | 1277 |
by(subst setsum_assign)(auto split: option.split dest!: find_start_eq_Some_above simp add: ivl_disj_un(2)[symmetric] add.commute add_increasing le_diff_eq le_max_iff_disj) |
1278 |
also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
|
1279 |
by(rule SUP_upper) simp |
|
1280 |
finally show ?thesis by simp |
|
1281 |
next |
|
1282 |
case False |
|
| 59053 | 1283 |
have "setsum (?rr y) {..z} = \<integral>\<^sup>+ z. ?rr y z \<partial>count_space {..z}"
|
| 59023 | 1284 |
by(simp add: nn_integral_count_space_finite max_def) |
| 59053 | 1285 |
also have "\<dots> \<le> \<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV" |
| 59023 | 1286 |
by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
| 59053 | 1287 |
also have "\<dots> = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" by(simp add: eq) |
1288 |
finally have *: "setsum (?rr y) {..z} = \<dots>" using False by simp
|
|
1289 |
also have "\<dots> = (SUP n. ereal (\<Sum>x<n. ?pp y x))" |
|
| 59023 | 1290 |
by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
| 59053 | 1291 |
also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z)) + setsum (?rr y) {..<z}"
|
| 59023 | 1292 |
proof(rule SUP_least) |
1293 |
fix n |
|
| 59053 | 1294 |
have "setsum (?pp y) {..<n} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<n}"
|
| 59023 | 1295 |
by(simp add: nn_integral_count_space_finite max_def) |
| 59053 | 1296 |
also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
| 59023 | 1297 |
by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
| 59053 | 1298 |
also have "\<dots> = setsum (?rr y) {..z}" using * by simp
|
1299 |
finally obtain k where k: "find_start y (setsum (?pp y) {..<n}) = Some k"
|
|
| 59023 | 1300 |
by(fastforce simp add: find_start_def) |
| 59053 | 1301 |
with \<open>ereal (setsum (?pp y) {..<n}) \<le> setsum (?rr y) {..z}\<close>
|
| 59023 | 1302 |
have "k \<le> z" by(auto simp add: find_start_def split: split_if_asm intro: Least_le) |
| 59053 | 1303 |
then have "setsum (?pp y) {..<n} - setsum (?rr y) {..<z} \<le> ereal (\<Sum>x<Suc n. assign y x z)"
|
1304 |
using \<open>ereal (setsum (?pp y) {..<n}) \<le> setsum (?rr y) {..z}\<close>
|
|
1305 |
apply (subst setsum_assign) |
|
1306 |
apply (auto simp add: field_simps max_def k ivl_disj_un(2)[symmetric]) |
|
1307 |
apply (meson add_increasing le_cases pmf_nonneg) |
|
1308 |
done |
|
| 59023 | 1309 |
also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
1310 |
by(rule SUP_upper) simp |
|
| 59053 | 1311 |
finally show "ereal (\<Sum>x<n. ?pp y x) \<le> \<dots> + setsum (?rr y) {..<z}"
|
| 59023 | 1312 |
by(simp add: ereal_minus(1)[symmetric] ereal_minus_le del: ereal_minus(1)) |
1313 |
qed |
|
1314 |
finally show ?thesis |
|
1315 |
by(simp add: ivl_disj_un(2)[symmetric] plus_ereal.simps(1)[symmetric] ereal_add_le_add_iff2 del: plus_ereal.simps(1)) |
|
1316 |
qed |
|
1317 |
qed |
|
1318 |
finally show "?thesis y z" . |
|
1319 |
qed |
|
1320 |
||
1321 |
{ fix y x
|
|
| 59053 | 1322 |
have "(\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = ?pp y x" |
1323 |
proof(cases "setsum (?pp y) {..<x} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV")
|
|
| 59023 | 1324 |
case False |
| 59053 | 1325 |
let ?used = "setsum (?pp y) {..<x}"
|
1326 |
have "?used = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<x}"
|
|
| 59023 | 1327 |
by(simp add: nn_integral_count_space_finite max_def) |
| 59053 | 1328 |
also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
| 59023 | 1329 |
by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_mono) |
1330 |
finally have "?used < \<dots>" using False by auto |
|
1331 |
also note eq finally have "?used \<in> dom (find_start y)" by(rule in_dom_find_startI) |
|
1332 |
then obtain k where k: "find_start y ?used = Some k" by auto |
|
| 59053 | 1333 |
let ?f = "\<lambda>z. if z < k then 0 else if z = k then setsum (?rr y) {..k} - ?used else ?rr y z"
|
1334 |
let ?g = "\<lambda>x'. if x' < x then 0 else ?pp y x'" |
|
1335 |
have "?pp y x = ?g x" by simp |
|
| 59023 | 1336 |
also have "?g x \<le> \<integral>\<^sup>+ x'. ?g x' \<partial>count_space UNIV" by(rule nn_integral_ge_point) simp |
1337 |
also {
|
|
| 59053 | 1338 |
have "?used = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<x}"
|
| 59023 | 1339 |
by(simp add: nn_integral_count_space_finite max_def) |
| 59053 | 1340 |
also have "\<dots> = \<integral>\<^sup>+ x'. (if x' < x then ?pp y x' else 0) \<partial>count_space UNIV" |
1341 |
by(simp add: nn_integral_count_space_indicator indicator_def if_distrib zero_ereal_def cong del: if_cong) |
|
1342 |
also have "(\<integral>\<^sup>+ x'. ?g x' \<partial>count_space UNIV) + \<dots> = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
| 59023 | 1343 |
by(subst nn_integral_add[symmetric])(auto intro: nn_integral_cong) |
1344 |
also note calculation } |
|
| 59053 | 1345 |
ultimately have "ereal (?pp y x) + ?used \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
| 59023 | 1346 |
by (metis (no_types, lifting) ereal_add_mono order_refl) |
1347 |
also note eq |
|
| 59053 | 1348 |
also have "(\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV) + (\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space UNIV)"
|
| 59023 | 1349 |
using k by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong simp add: ivl_disj_un(2)[symmetric] setsum_nonneg dest: find_start_eq_Some_least find_start_eq_Some_above) |
| 59053 | 1350 |
also have "(\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space UNIV) =
|
1351 |
(\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space {..k})"
|
|
| 59023 | 1352 |
by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_cong) |
1353 |
also have "\<dots> = ?used" |
|
1354 |
using k by(auto simp add: nn_integral_count_space_finite max_def ivl_disj_un(2)[symmetric] diff_le_iff setsum_nonneg dest: find_start_eq_Some_least) |
|
| 59053 | 1355 |
finally have "?pp y x \<le> (\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV)" |
| 59023 | 1356 |
by(cases "\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV") simp_all |
1357 |
then show ?thesis using k |
|
1358 |
by(simp add: assign_def nn_integral_assign_aux diff_le_iff find_start_eq_Some_above min_def) |
|
1359 |
next |
|
1360 |
case True |
|
| 59053 | 1361 |
have "setsum (?pp y) {..x} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..x}"
|
| 59023 | 1362 |
by(simp add: nn_integral_count_space_finite max_def) |
| 59053 | 1363 |
also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
| 59023 | 1364 |
by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
| 59053 | 1365 |
also have "\<dots> = setsum (?pp y) {..<x}" by(simp add: True)
|
1366 |
finally have "?pp y x = 0" by(simp add: ivl_disj_un(2)[symmetric] eq_iff del: pmf_le_0_iff) |
|
| 59023 | 1367 |
thus ?thesis |
| 59053 | 1368 |
by(cases "find_start y (setsum (?pp y) {..<x})")(simp_all add: assign_def diff_le_iff find_start_eq_Some_above)
|
| 59023 | 1369 |
qed } |
1370 |
note nn_integral_assign2 = this |
|
1371 |
||
| 59053 | 1372 |
def a \<equiv> "embed_pmf (\<lambda>(y, x, z). assign y x z)" |
| 59023 | 1373 |
{ fix y x z
|
| 59053 | 1374 |
have "assign y x z = pmf a (y, x, z)" |
1375 |
unfolding a_def |
|
1376 |
proof (subst pmf_embed_pmf) |
|
1377 |
have "(\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>count_space UNIV) = |
|
1378 |
(\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>(count_space ((\<lambda>((y, x), z). (y, x, z)) ` (pp \<times> UNIV))))" |
|
1379 |
by (force simp add: nn_integral_count_space_indicator pmf_eq_0_set_pmf split: split_indicator |
|
1380 |
intro!: nn_integral_cong assign_eq_0_outside) |
|
1381 |
also have "\<dots> = (\<integral>\<^sup>+ x. ereal ((\<lambda>((y, x), z). assign y x z) x) \<partial>(count_space (pp \<times> UNIV)))" |
|
1382 |
by (subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric]) |
|
1383 |
(auto simp: inj_on_def intro!: nn_integral_cong) |
|
1384 |
also have "\<dots> = (\<integral>\<^sup>+ y. \<integral>\<^sup>+z. ereal ((\<lambda>((y, x), z). assign y x z) (y, z)) \<partial>count_space UNIV \<partial>count_space pp)" |
|
1385 |
by (subst sigma_finite_measure.nn_integral_fst) |
|
1386 |
(auto simp: pair_measure_countable sigma_finite_measure_count_space_countable) |
|
1387 |
also have "\<dots> = (\<integral>\<^sup>+ z. ?pp (fst z) (snd z) \<partial>count_space pp)" |
|
1388 |
by (subst nn_integral_assign2[symmetric]) (auto intro!: nn_integral_cong) |
|
1389 |
finally show "(\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>count_space UNIV) = 1" |
|
1390 |
by (simp add: nn_integral_pmf emeasure_pmf) |
|
1391 |
qed auto } |
|
1392 |
note a = this |
|
| 59023 | 1393 |
|
| 59053 | 1394 |
def pr \<equiv> "map_pmf (\<lambda>(y, x, z). (from_nat_into (A y) x, from_nat_into (B y) z)) a" |
| 59023 | 1395 |
|
| 59053 | 1396 |
have "rel_pmf (R OO S) p r" |
| 59023 | 1397 |
proof |
| 59053 | 1398 |
have pp_eq: "pp = map_pmf (\<lambda>(y, x, z). (y, x)) a" |
1399 |
proof (rule pmf_eqI) |
|
1400 |
fix i |
|
1401 |
show "pmf pp i = pmf (map_pmf (\<lambda>(y, x, z). (y, x)) a) i" |
|
1402 |
using nn_integral_assign2[of "fst i" "snd i", symmetric] |
|
1403 |
by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map |
|
1404 |
simp del: ereal.inject intro!: arg_cong2[where f=emeasure]) |
|
1405 |
qed |
|
1406 |
moreover have pq_eq: "pq = map_pmf (\<lambda>(y, x). (from_nat_into (A y) x, y)) pp" |
|
1407 |
by (simp add: pp_def map_pmf_comp split_beta A[symmetric] cong: map_pmf_cong) |
|
1408 |
ultimately show "map_pmf fst pr = p" |
|
1409 |
unfolding p pr_def by (simp add: map_pmf_comp split_beta) |
|
1410 |
||
1411 |
have rr_eq: "rr = map_pmf (\<lambda>(y, x, z). (y, z)) a" |
|
1412 |
proof (rule pmf_eqI) |
|
1413 |
fix i show "pmf rr i = pmf (map_pmf (\<lambda>(y, x, z). (y, z)) a) i" |
|
1414 |
using nn_integral_assign1[of "fst i" "snd i", symmetric] |
|
1415 |
by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map |
|
1416 |
simp del: ereal.inject intro!: arg_cong2[where f=emeasure]) |
|
1417 |
qed |
|
1418 |
moreover have qr_eq: "qr = map_pmf (\<lambda>(y, z). (y, from_nat_into (B y) z)) rr" |
|
1419 |
by (simp add: rr_def map_pmf_comp split_beta B[symmetric] cong: map_pmf_cong) |
|
1420 |
ultimately show "map_pmf snd pr = r" |
|
1421 |
unfolding r pr_def by (simp add: map_pmf_comp split_beta) |
|
1422 |
||
1423 |
fix x z assume "(x, z) \<in> set_pmf pr" |
|
1424 |
then have "\<exists>y. (x, y) \<in> set_pmf pq \<and> (y, z) \<in> set_pmf qr" |
|
1425 |
by (force simp add: pp_eq pq_eq rr_eq qr_eq set_map_pmf pr_def image_image) |
|
1426 |
with pq qr show "(R OO S) x z" |
|
1427 |
by blast |
|
1428 |
qed } |
|
| 59023 | 1429 |
then show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)" |
1430 |
by(auto simp add: le_fun_def) |
|
1431 |
qed (fact natLeq_card_order natLeq_cinfinite)+ |
|
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1432 |
|
| 59134 | 1433 |
lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)" |
1434 |
proof safe |
|
1435 |
fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M" |
|
1436 |
then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b" |
|
1437 |
and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq" |
|
1438 |
by (force elim: rel_pmf.cases) |
|
1439 |
moreover have "set_pmf (return_pmf x) = {x}"
|
|
1440 |
by (simp add: set_return_pmf) |
|
1441 |
with `a \<in> M` have "(x, a) \<in> pq" |
|
1442 |
by (force simp: eq set_map_pmf) |
|
1443 |
with * show "R x a" |
|
1444 |
by auto |
|
1445 |
qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"] |
|
1446 |
simp: map_fst_pair_pmf map_snd_pair_pmf set_pair_pmf set_return_pmf) |
|
1447 |
||
1448 |
lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)" |
|
1449 |
by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1) |
|
1450 |
||
1451 |
lemma rel_pmf_rel_prod: |
|
1452 |
"rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B') \<longleftrightarrow> rel_pmf R A B \<and> rel_pmf S A' B'" |
|
1453 |
proof safe |
|
1454 |
assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')" |
|
1455 |
then obtain pq where pq: "\<And>a b c d. ((a, c), (b, d)) \<in> set_pmf pq \<Longrightarrow> R a b \<and> S c d" |
|
1456 |
and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'" |
|
1457 |
by (force elim: rel_pmf.cases) |
|
1458 |
show "rel_pmf R A B" |
|
1459 |
proof (rule rel_pmf.intros) |
|
1460 |
let ?f = "\<lambda>(a, b). (fst a, fst b)" |
|
1461 |
have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd" |
|
1462 |
by auto |
|
1463 |
||
1464 |
show "map_pmf fst (map_pmf ?f pq) = A" |
|
1465 |
by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf) |
|
1466 |
show "map_pmf snd (map_pmf ?f pq) = B" |
|
1467 |
by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf) |
|
1468 |
||
1469 |
fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)" |
|
1470 |
then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq" |
|
1471 |
by (auto simp: set_map_pmf) |
|
1472 |
from pq[OF this] show "R a b" .. |
|
1473 |
qed |
|
1474 |
show "rel_pmf S A' B'" |
|
1475 |
proof (rule rel_pmf.intros) |
|
1476 |
let ?f = "\<lambda>(a, b). (snd a, snd b)" |
|
1477 |
have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd" |
|
1478 |
by auto |
|
1479 |
||
1480 |
show "map_pmf fst (map_pmf ?f pq) = A'" |
|
1481 |
by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf) |
|
1482 |
show "map_pmf snd (map_pmf ?f pq) = B'" |
|
1483 |
by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf) |
|
1484 |
||
1485 |
fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)" |
|
1486 |
then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq" |
|
1487 |
by (auto simp: set_map_pmf) |
|
1488 |
from pq[OF this] show "S c d" .. |
|
1489 |
qed |
|
1490 |
next |
|
1491 |
assume "rel_pmf R A B" "rel_pmf S A' B'" |
|
1492 |
then obtain Rpq Spq |
|
1493 |
where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b" |
|
1494 |
"map_pmf fst Rpq = A" "map_pmf snd Rpq = B" |
|
1495 |
and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b" |
|
1496 |
"map_pmf fst Spq = A'" "map_pmf snd Spq = B'" |
|
1497 |
by (force elim: rel_pmf.cases) |
|
1498 |
||
1499 |
let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))" |
|
1500 |
let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)" |
|
1501 |
have [simp]: "(\<lambda>x. fst (?f x)) = (\<lambda>(a, b). (fst a, fst b))" "(\<lambda>x. snd (?f x)) = (\<lambda>(a, b). (snd a, snd b))" |
|
1502 |
by auto |
|
1503 |
||
1504 |
show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')" |
|
1505 |
by (rule rel_pmf.intros[where pq="?pq"]) |
|
1506 |
(auto simp: map_snd_pair_pmf map_fst_pair_pmf set_pair_pmf set_map_pmf map_pmf_comp Rpq Spq |
|
1507 |
map_pair) |
|
1508 |
qed |
|
1509 |
||
|
58587
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1510 |
end |
|
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff
changeset
|
1511 |