1(*  Title:      HOL/Probability/Probability_Mass_Function.thy
2    Author:     Johannes H��lzl, TU M��nchen
3    Author:     Andreas Lochbihler, ETH Zurich
4*)
5
6section \<open> Probability mass function \<close>
7
8theory Probability_Mass_Function
9imports
10  Giry_Monad
11  "HOL-Library.Multiset"
12begin
13
14lemma AE_emeasure_singleton:
15  assumes x: "emeasure M {x} \<noteq> 0" and ae: "AE x in M. P x" shows "P x"
16proof -
17  from x have x_M: "{x} \<in> sets M"
18    by (auto intro: emeasure_notin_sets)
19  from ae obtain N where N: "{x\<in>space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
20    by (auto elim: AE_E)
21  { assume "\<not> P x"
22    with x_M[THEN sets.sets_into_space] N have "emeasure M {x} \<le> emeasure M N"
23      by (intro emeasure_mono) auto
24    with x N have False
25      by (auto simp:) }
26  then show "P x" by auto
27qed
28
29lemma AE_measure_singleton: "measure M {x} \<noteq> 0 \<Longrightarrow> AE x in M. P x \<Longrightarrow> P x"
30  by (metis AE_emeasure_singleton measure_def emeasure_empty measure_empty)
31
32lemma (in finite_measure) AE_support_countable:
33  assumes [simp]: "sets M = UNIV"
34  shows "(AE x in M. measure M {x} \<noteq> 0) \<longleftrightarrow> (\<exists>S. countable S \<and> (AE x in M. x \<in> S))"
35proof
36  assume "\<exists>S. countable S \<and> (AE x in M. x \<in> S)"
37  then obtain S where S[intro]: "countable S" and ae: "AE x in M. x \<in> S"
38    by auto
39  then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
40    (\<integral>\<^sup>+ x. emeasure M {x} * indicator {x\<in>S. emeasure M {x} \<noteq> 0} x \<partial>count_space UNIV)"
41    by (subst emeasure_UN_countable)
42       (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space)
43  also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} * indicator S x \<partial>count_space UNIV)"
44    by (auto intro!: nn_integral_cong split: split_indicator)
45  also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
46    by (subst emeasure_UN_countable)
47       (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space)
48  also have "\<dots> = emeasure M (space M)"
49    using ae by (intro emeasure_eq_AE) auto
50  finally have "emeasure M {x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0} = emeasure M (space M)"
51    by (simp add: emeasure_single_in_space cong: rev_conj_cong)
52  with finite_measure_compl[of "{x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0}"]
53  have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
54    by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure measure_nonneg set_diff_eq cong: conj_cong)
55  then show "AE x in M. measure M {x} \<noteq> 0"
56    by (auto simp: emeasure_eq_measure)
57qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
58
59subsection \<open> PMF as measure \<close>
60
61typedef 'a pmf = "{M :: 'a measure. prob_space M \<and> sets M = UNIV \<and> (AE x in M. measure M {x} \<noteq> 0)}"
62  morphisms measure_pmf Abs_pmf
63  by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
64     (auto intro!: prob_space_uniform_measure AE_uniform_measureI)
65
66declare [[coercion measure_pmf]]
67
68lemma prob_space_measure_pmf: "prob_space (measure_pmf p)"
69  using pmf.measure_pmf[of p] by auto
70
71interpretation measure_pmf: prob_space "measure_pmf M" for M
72  by (rule prob_space_measure_pmf)
73
74interpretation measure_pmf: subprob_space "measure_pmf M" for M
75  by (rule prob_space_imp_subprob_space) unfold_locales
76
77lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)"
78  by unfold_locales
79
80locale pmf_as_measure
81begin
82
83setup_lifting type_definition_pmf
84
85end
86
87context
88begin
89
90interpretation pmf_as_measure .
91
92lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV"
93  by transfer blast
94
95lemma sets_measure_pmf_count_space[measurable_cong]:
96  "sets (measure_pmf M) = sets (count_space UNIV)"
97  by simp
98
99lemma space_measure_pmf[simp]: "space (measure_pmf p) = UNIV"
100  using sets_eq_imp_space_eq[of "measure_pmf p" "count_space UNIV"] by simp
101
102lemma measure_pmf_UNIV [simp]: "measure (measure_pmf p) UNIV = 1"
103using measure_pmf.prob_space[of p] by simp
104
105lemma measure_pmf_in_subprob_algebra[measurable (raw)]: "measure_pmf x \<in> space (subprob_algebra (count_space UNIV))"
106  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
107
108lemma measurable_pmf_measure1[simp]: "measurable (M :: 'a pmf) N = UNIV \<rightarrow> space N"
109  by (auto simp: measurable_def)
110
111lemma measurable_pmf_measure2[simp]: "measurable N (M :: 'a pmf) = measurable N (count_space UNIV)"
112  by (intro measurable_cong_sets) simp_all
113
114lemma measurable_pair_restrict_pmf2:
115  assumes "countable A"
116  assumes [measurable]: "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L"
117  shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L" (is "f \<in> measurable ?M _")
118proof -
119  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
120    by (simp add: restrict_count_space)
121
122  show ?thesis
123    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (fst b, a)" and g=snd and I=A,
124                                            unfolded prod.collapse] assms)
125        measurable
126qed
127
128lemma measurable_pair_restrict_pmf1:
129  assumes "countable A"
130  assumes [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L"
131  shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L"
132proof -
133  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
134    by (simp add: restrict_count_space)
135
136  show ?thesis
137    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A,
138                                            unfolded prod.collapse] assms)
139        measurable
140qed
141
142lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
143
144lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
145declare [[coercion set_pmf]]
146
147lemma AE_measure_pmf: "AE x in (M::'a pmf). x \<in> M"
148  by transfer simp
149
150lemma emeasure_pmf_single_eq_zero_iff:
151  fixes M :: "'a pmf"
152  shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
153  unfolding set_pmf.rep_eq by (simp add: measure_pmf.emeasure_eq_measure)
154
155lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"
156  using AE_measure_singleton[of M] AE_measure_pmf[of M]
157  by (auto simp: set_pmf.rep_eq)
158
159lemma AE_pmfI: "(\<And>y. y \<in> set_pmf M \<Longrightarrow> P y) \<Longrightarrow> almost_everywhere (measure_pmf M) P"
160by(simp add: AE_measure_pmf_iff)
161
162lemma countable_set_pmf [simp]: "countable (set_pmf p)"
163  by transfer (metis prob_space.finite_measure finite_measure.countable_support)
164
165lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"
166  by transfer (simp add: less_le)
167
168lemma pmf_nonneg[simp]: "0 \<le> pmf p x"
169  by transfer simp
170
171lemma pmf_not_neg [simp]: "\<not>pmf p x < 0"
172  by (simp add: not_less pmf_nonneg)
173
174lemma pmf_pos [simp]: "pmf p x \<noteq> 0 \<Longrightarrow> pmf p x > 0"
175  using pmf_nonneg[of p x] by linarith
176
177lemma pmf_le_1: "pmf p x \<le> 1"
178  by (simp add: pmf.rep_eq)
179
180lemma set_pmf_not_empty: "set_pmf M \<noteq> {}"
181  using AE_measure_pmf[of M] by (intro notI) simp
182
183lemma set_pmf_iff: "x \<in> set_pmf M \<longleftrightarrow> pmf M x \<noteq> 0"
184  by transfer simp
185
186lemma pmf_positive_iff: "0 < pmf p x \<longleftrightarrow> x \<in> set_pmf p"
187  unfolding less_le by (simp add: set_pmf_iff)
188
189lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
190  by (auto simp: set_pmf_iff)
191
192lemma set_pmf_eq': "set_pmf p = {x. pmf p x > 0}"
193proof safe
194  fix x assume "x \<in> set_pmf p"
195  hence "pmf p x \<noteq> 0" by (auto simp: set_pmf_eq)
196  with pmf_nonneg[of p x] show "pmf p x > 0" by simp
197qed (auto simp: set_pmf_eq)
198
199lemma emeasure_pmf_single:
200  fixes M :: "'a pmf"
201  shows "emeasure M {x} = pmf M x"
202  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
203
204lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
205  using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure pmf_nonneg measure_nonneg)
206
207lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"
208  by (subst emeasure_eq_sum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)
209
210lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = sum (pmf M) S"
211  using emeasure_measure_pmf_finite[of S M]
212  by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg sum_nonneg pmf_nonneg)
213
214lemma sum_pmf_eq_1:
215  assumes "finite A" "set_pmf p \<subseteq> A"
216  shows   "(\<Sum>x\<in>A. pmf p x) = 1"
217proof -
218  have "(\<Sum>x\<in>A. pmf p x) = measure_pmf.prob p A"
219    by (simp add: measure_measure_pmf_finite assms)
220  also from assms have "\<dots> = 1"
221    by (subst measure_pmf.prob_eq_1) (auto simp: AE_measure_pmf_iff)
222  finally show ?thesis .
223qed
224
225lemma nn_integral_measure_pmf_support:
226  fixes f :: "'a \<Rightarrow> ennreal"
227  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"
228  shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>A. f x * pmf M x)"
229proof -
230  have "(\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
231    using nn by (intro nn_integral_cong_AE) (auto simp: AE_measure_pmf_iff split: split_indicator)
232  also have "\<dots> = (\<Sum>x\<in>A. f x * emeasure M {x})"
233    using assms by (intro nn_integral_indicator_finite) auto
234  finally show ?thesis
235    by (simp add: emeasure_measure_pmf_finite)
236qed
237
238lemma nn_integral_measure_pmf_finite:
239  fixes f :: "'a \<Rightarrow> ennreal"
240  assumes f: "finite (set_pmf M)" and nn: "\<And>x. x \<in> set_pmf M \<Longrightarrow> 0 \<le> f x"
241  shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>set_pmf M. f x * pmf M x)"
242  using assms by (intro nn_integral_measure_pmf_support) auto
243
244lemma integrable_measure_pmf_finite:
245  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
246  shows "finite (set_pmf M) \<Longrightarrow> integrable M f"
247  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite ennreal_mult_less_top)
248
249lemma integral_measure_pmf_real:
250  assumes [simp]: "finite A" and "\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A"
251  shows "(\<integral>x. f x \<partial>measure_pmf M) = (\<Sum>a\<in>A. f a * pmf M a)"
252proof -
253  have "(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x * indicator A x \<partial>measure_pmf M)"
254    using assms(2) by (intro integral_cong_AE) (auto split: split_indicator simp: AE_measure_pmf_iff)
255  also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)"
256    by (subst integral_indicator_finite_real)
257       (auto simp: measure_def emeasure_measure_pmf_finite pmf_nonneg)
258  finally show ?thesis .
259qed
260
261lemma integrable_pmf: "integrable (count_space X) (pmf M)"
262proof -
263  have " (\<integral>\<^sup>+ x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+ x. pmf M x \<partial>count_space (M \<inter> X))"
264    by (auto simp add: nn_integral_count_space_indicator set_pmf_iff intro!: nn_integral_cong split: split_indicator)
265  then have "integrable (count_space X) (pmf M) = integrable (count_space (M \<inter> X)) (pmf M)"
266    by (simp add: integrable_iff_bounded pmf_nonneg)
267  then show ?thesis
268    by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def)
269qed
270
271lemma integral_pmf: "(\<integral>x. pmf M x \<partial>count_space X) = measure M X"
272proof -
273  have "(\<integral>x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+x. pmf M x \<partial>count_space X)"
274    by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral)
275  also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space (X \<inter> M))"
276    by (auto intro!: nn_integral_cong_AE split: split_indicator
277             simp: pmf.rep_eq measure_pmf.emeasure_eq_measure nn_integral_count_space_indicator
278                   AE_count_space set_pmf_iff)
279  also have "\<dots> = emeasure M (X \<inter> M)"
280    by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf)
281  also have "\<dots> = emeasure M X"
282    by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff)
283  finally show ?thesis
284    by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg integral_nonneg pmf_nonneg)
285qed
286
287lemma integral_pmf_restrict:
288  "(f::'a \<Rightarrow> 'b::{banach, second_countable_topology}) \<in> borel_measurable (count_space UNIV) \<Longrightarrow>
289    (\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x \<partial>restrict_space M M)"
290  by (auto intro!: integral_cong_AE simp add: integral_restrict_space AE_measure_pmf_iff)
291
292lemma emeasure_pmf: "emeasure (M::'a pmf) M = 1"
293proof -
294  have "emeasure (M::'a pmf) M = emeasure (M::'a pmf) (space M)"
295    by (intro emeasure_eq_AE) (simp_all add: AE_measure_pmf)
296  then show ?thesis
297    using measure_pmf.emeasure_space_1 by simp
298qed
299
300lemma emeasure_pmf_UNIV [simp]: "emeasure (measure_pmf M) UNIV = 1"
301using measure_pmf.emeasure_space_1[of M] by simp
302
303lemma in_null_sets_measure_pmfI:
304  "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
305using emeasure_eq_0_AE[where ?P="\<lambda>x. x \<in> A" and M="measure_pmf p"]
306by(auto simp add: null_sets_def AE_measure_pmf_iff)
307
308lemma measure_subprob: "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
309  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
310
311subsection \<open> Monad Interpretation \<close>
312
313lemma measurable_measure_pmf[measurable]:
314  "(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))"
315  by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales
316
317lemma bind_measure_pmf_cong:
318  assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)"
319  assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i"
320  shows "bind (measure_pmf x) A = bind (measure_pmf x) B"
321proof (rule measure_eqI)
322  show "sets (measure_pmf x \<bind> A) = sets (measure_pmf x \<bind> B)"
323    using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra)
324next
325  fix X assume "X \<in> sets (measure_pmf x \<bind> A)"
326  then have X: "X \<in> sets N"
327    using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra)
328  show "emeasure (measure_pmf x \<bind> A) X = emeasure (measure_pmf x \<bind> B) X"
329    using assms
330    by (subst (1 2) emeasure_bind[where N=N, OF _ _ X])
331       (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
332qed
333
334lift_definition bind_pmf :: "'a pmf \<Rightarrow> ('a \<Rightarrow> 'b pmf ) \<Rightarrow> 'b pmf" is bind
335proof (clarify, intro conjI)
336  fix f :: "'a measure" and g :: "'a \<Rightarrow> 'b measure"
337  assume "prob_space f"
338  then interpret f: prob_space f .
339  assume "sets f = UNIV" and ae_f: "AE x in f. measure f {x} \<noteq> 0"
340  then have s_f[simp]: "sets f = sets (count_space UNIV)"
341    by simp
342  assume g: "\<And>x. prob_space (g x) \<and> sets (g x) = UNIV \<and> (AE y in g x. measure (g x) {y} \<noteq> 0)"
343  then have g: "\<And>x. prob_space (g x)" and s_g[simp]: "\<And>x. sets (g x) = sets (count_space UNIV)"
344    and ae_g: "\<And>x. AE y in g x. measure (g x) {y} \<noteq> 0"
345    by auto
346
347  have [measurable]: "g \<in> measurable f (subprob_algebra (count_space UNIV))"
348    by (auto simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space g)
349
350  show "prob_space (f \<bind> g)"
351    using g by (intro f.prob_space_bind[where S="count_space UNIV"]) auto
352  then interpret fg: prob_space "f \<bind> g" .
353  show [simp]: "sets (f \<bind> g) = UNIV"
354    using sets_eq_imp_space_eq[OF s_f]
355    by (subst sets_bind[where N="count_space UNIV"]) auto
356  show "AE x in f \<bind> g. measure (f \<bind> g) {x} \<noteq> 0"
357    apply (simp add: fg.prob_eq_0 AE_bind[where B="count_space UNIV"])
358    using ae_f
359    apply eventually_elim
360    using ae_g
361    apply eventually_elim
362    apply (auto dest: AE_measure_singleton)
363    done
364qed
365
366adhoc_overloading Monad_Syntax.bind bind_pmf
367
368lemma ennreal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"
369  unfolding pmf.rep_eq bind_pmf.rep_eq
370  by (auto simp: measure_pmf.measure_bind[where N="count_space UNIV"] measure_subprob measure_nonneg
371           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
372
373lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"
374  using ennreal_pmf_bind[of N f i]
375  by (subst (asm) nn_integral_eq_integral)
376     (auto simp: pmf_nonneg pmf_le_1 pmf_nonneg integral_nonneg
377           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
378
379lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"
380  by transfer (simp add: bind_const' prob_space_imp_subprob_space)
381
382lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
383proof -
384  have "set_pmf (bind_pmf M N) = {x. ennreal (pmf (bind_pmf M N) x) \<noteq> 0}"
385    by (simp add: set_pmf_eq pmf_nonneg)
386  also have "\<dots> = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
387    unfolding ennreal_pmf_bind
388    by (subst nn_integral_0_iff_AE) (auto simp: AE_measure_pmf_iff pmf_nonneg set_pmf_eq)
389  finally show ?thesis .
390qed
391
392lemma bind_pmf_cong [fundef_cong]:
393  assumes "p = q"
394  shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
395  unfolding \<open>p = q\<close>[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq
396  by (auto simp: AE_measure_pmf_iff Pi_iff space_subprob_algebra subprob_space_measure_pmf
397                 sets_bind[where N="count_space UNIV"] emeasure_bind[where N="count_space UNIV"]
398           intro!: nn_integral_cong_AE measure_eqI)
399
400lemma bind_pmf_cong_simp:
401  "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q =simp=> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
402  by (simp add: simp_implies_def cong: bind_pmf_cong)
403
404lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<bind> (\<lambda>x. measure_pmf (f x)))"
405  by transfer simp
406
407lemma 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)"
408  using measurable_measure_pmf[of N]
409  unfolding measure_pmf_bind
410  apply (intro nn_integral_bind[where B="count_space UNIV"])
411  apply auto
412  done
413
414lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)"
415  using measurable_measure_pmf[of N]
416  unfolding measure_pmf_bind
417  by (subst emeasure_bind[where N="count_space UNIV"]) auto
418
419lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)"
420  by (auto intro!: prob_space_return simp: AE_return measure_return)
421
422lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x"
423  by transfer
424     (auto intro!: prob_space_imp_subprob_space bind_return[where N="count_space UNIV"]
425           simp: space_subprob_algebra)
426
427lemma set_return_pmf[simp]: "set_pmf (return_pmf x) = {x}"
428  by transfer (auto simp add: measure_return split: split_indicator)
429
430lemma bind_return_pmf': "bind_pmf N return_pmf = N"
431proof (transfer, clarify)
432  fix N :: "'a measure" assume "sets N = UNIV" then show "N \<bind> return (count_space UNIV) = N"
433    by (subst return_sets_cong[where N=N]) (simp_all add: bind_return')
434qed
435
436lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)"
437  by transfer
438     (auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"]
439           simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space)
440
441definition "map_pmf f M = bind_pmf M (\<lambda>x. return_pmf (f x))"
442
443lemma map_bind_pmf: "map_pmf f (bind_pmf M g) = bind_pmf M (\<lambda>x. map_pmf f (g x))"
444  by (simp add: map_pmf_def bind_assoc_pmf)
445
446lemma bind_map_pmf: "bind_pmf (map_pmf f M) g = bind_pmf M (\<lambda>x. g (f x))"
447  by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)
448
449lemma map_pmf_transfer[transfer_rule]:
450  "rel_fun (=) (rel_fun cr_pmf cr_pmf) (\<lambda>f M. distr M (count_space UNIV) f) map_pmf"
451proof -
452  have "rel_fun (=) (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf)
453     (\<lambda>f M. M \<bind> (return (count_space UNIV) o f)) map_pmf"
454    unfolding map_pmf_def[abs_def] comp_def by transfer_prover
455  then show ?thesis
456    by (force simp: rel_fun_def cr_pmf_def bind_return_distr)
457qed
458
459lemma map_pmf_rep_eq:
460  "measure_pmf (map_pmf f M) = distr (measure_pmf M) (count_space UNIV) f"
461  unfolding map_pmf_def bind_pmf.rep_eq comp_def return_pmf.rep_eq
462  using bind_return_distr[of M f "count_space UNIV"] by (simp add: comp_def)
463
464lemma map_pmf_id[simp]: "map_pmf id = id"
465  by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI)
466
467lemma map_pmf_ident[simp]: "map_pmf (\<lambda>x. x) = (\<lambda>x. x)"
468  using map_pmf_id unfolding id_def .
469
470lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g"
471  by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def)
472
473lemma map_pmf_comp: "map_pmf f (map_pmf g M) = map_pmf (\<lambda>x. f (g x)) M"
474  using map_pmf_compose[of f g] by (simp add: comp_def)
475
476lemma map_pmf_cong: "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g q"
477  unfolding map_pmf_def by (rule bind_pmf_cong) auto
478
479lemma pmf_set_map: "set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"
480  by (auto simp add: comp_def fun_eq_iff map_pmf_def)
481
482lemma set_map_pmf[simp]: "set_pmf (map_pmf f M) = f`set_pmf M"
483  using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff)
484
485lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)"
486  unfolding map_pmf_rep_eq by (subst emeasure_distr) auto
487
488lemma measure_map_pmf[simp]: "measure (map_pmf f M) X = measure M (f -` X)"
489using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
490
491lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"
492  unfolding map_pmf_rep_eq by (intro nn_integral_distr) auto
493
494lemma ennreal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
495proof (transfer fixing: f x)
496  fix p :: "'b measure"
497  presume "prob_space p"
498  then interpret prob_space p .
499  presume "sets p = UNIV"
500  then show "ennreal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
501    by(simp add: measure_distr measurable_def emeasure_eq_measure)
502qed simp_all
503
504lemma pmf_map: "pmf (map_pmf f p) x = measure p (f -` {x})"
505proof (transfer fixing: f x)
506  fix p :: "'b measure"
507  presume "prob_space p"
508  then interpret prob_space p .
509  presume "sets p = UNIV"
510  then show "measure (distr p (count_space UNIV) f) {x} = measure p (f -` {x})"
511    by(simp add: measure_distr measurable_def emeasure_eq_measure)
512qed simp_all
513
514lemma nn_integral_pmf: "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = emeasure (measure_pmf p) A"
515proof -
516  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))"
517    by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong)
518  also have "\<dots> = emeasure (measure_pmf p) (\<Union>x\<in>A \<inter> set_pmf p. {x})"
519    by(subst emeasure_UN_countable)(auto simp add: emeasure_pmf_single disjoint_family_on_def)
520  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})"
521    by(rule emeasure_Un_null_set[symmetric])(auto intro: in_null_sets_measure_pmfI)
522  also have "\<dots> = emeasure (measure_pmf p) A"
523    by(auto intro: arg_cong2[where f=emeasure])
524  finally show ?thesis .
525qed
526
527lemma integral_map_pmf[simp]:
528  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
529  shows "integral\<^sup>L (map_pmf g p) f = integral\<^sup>L p (\<lambda>x. f (g x))"
530  by (simp add: integral_distr map_pmf_rep_eq)
531
532lemma pmf_abs_summable [intro]: "pmf p abs_summable_on A"
533  by (rule abs_summable_on_subset[OF _ subset_UNIV])
534     (auto simp:  abs_summable_on_def integrable_iff_bounded nn_integral_pmf)
535
536lemma measure_pmf_conv_infsetsum: "measure (measure_pmf p) A = infsetsum (pmf p) A"
537  unfolding infsetsum_def by (simp add: integral_eq_nn_integral nn_integral_pmf measure_def)
538
539lemma infsetsum_pmf_eq_1:
540  assumes "set_pmf p \<subseteq> A"
541  shows   "infsetsum (pmf p) A = 1"
542proof -
543  have "infsetsum (pmf p) A = lebesgue_integral (count_space UNIV) (pmf p)"
544    using assms unfolding infsetsum_altdef set_lebesgue_integral_def
545    by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def set_pmf_eq)
546  also have "\<dots> = 1"
547    by (subst integral_eq_nn_integral) (auto simp: nn_integral_pmf)
548  finally show ?thesis .
549qed
550
551lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)"
552  by transfer (simp add: distr_return)
553
554lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c"
555  by transfer (auto simp: prob_space.distr_const)
556
557lemma pmf_return [simp]: "pmf (return_pmf x) y = indicator {y} x"
558  by transfer (simp add: measure_return)
559
560lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x"
561  unfolding return_pmf.rep_eq by (intro nn_integral_return) auto
562
563lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x"
564  unfolding return_pmf.rep_eq by (intro emeasure_return) auto
565
566lemma measure_return_pmf [simp]: "measure_pmf.prob (return_pmf x) A = indicator A x"
567proof -
568  have "ennreal (measure_pmf.prob (return_pmf x) A) =
569          emeasure (measure_pmf (return_pmf x)) A"
570    by (simp add: measure_pmf.emeasure_eq_measure)
571  also have "\<dots> = ennreal (indicator A x)" by (simp add: ennreal_indicator)
572  finally show ?thesis by simp
573qed
574
575lemma return_pmf_inj[simp]: "return_pmf x = return_pmf y \<longleftrightarrow> x = y"
576  by (metis insertI1 set_return_pmf singletonD)
577
578lemma map_pmf_eq_return_pmf_iff:
579  "map_pmf f p = return_pmf x \<longleftrightarrow> (\<forall>y \<in> set_pmf p. f y = x)"
580proof
581  assume "map_pmf f p = return_pmf x"
582  then have "set_pmf (map_pmf f p) = set_pmf (return_pmf x)" by simp
583  then show "\<forall>y \<in> set_pmf p. f y = x" by auto
584next
585  assume "\<forall>y \<in> set_pmf p. f y = x"
586  then show "map_pmf f p = return_pmf x"
587    unfolding map_pmf_const[symmetric, of _ p] by (intro map_pmf_cong) auto
588qed
589
590definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))"
591
592lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b"
593  unfolding pair_pmf_def pmf_bind pmf_return
594  apply (subst integral_measure_pmf_real[where A="{b}"])
595  apply (auto simp: indicator_eq_0_iff)
596  apply (subst integral_measure_pmf_real[where A="{a}"])
597  apply (auto simp: indicator_eq_0_iff sum_nonneg_eq_0_iff pmf_nonneg)
598  done
599
600lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"
601  unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto
602
603lemma measure_pmf_in_subprob_space[measurable (raw)]:
604  "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
605  by (simp add: space_subprob_algebra) intro_locales
606
607lemma 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)"
608proof -
609  have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. f x * indicator (A \<times> B) x \<partial>pair_pmf A B)"
610    by (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)
611  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
612    by (simp add: pair_pmf_def)
613  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
614    by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
615  finally show ?thesis .
616qed
617
618lemma bind_pair_pmf:
619  assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)"
620  shows "measure_pmf (pair_pmf A B) \<bind> M = (measure_pmf A \<bind> (\<lambda>x. measure_pmf B \<bind> (\<lambda>y. M (x, y))))"
621    (is "?L = ?R")
622proof (rule measure_eqI)
623  have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)"
624    using M[THEN measurable_space] by (simp_all add: space_pair_measure)
625
626  note measurable_bind[where N="count_space UNIV", measurable]
627  note measure_pmf_in_subprob_space[simp]
628
629  have sets_eq_N: "sets ?L = N"
630    by (subst sets_bind[OF sets_kernel[OF M']]) auto
631  show "sets ?L = sets ?R"
632    using measurable_space[OF M]
633    by (simp add: sets_eq_N space_pair_measure space_subprob_algebra)
634  fix X assume "X \<in> sets ?L"
635  then have X[measurable]: "X \<in> sets N"
636    unfolding sets_eq_N .
637  then show "emeasure ?L X = emeasure ?R X"
638    apply (simp add: emeasure_bind[OF _ M' X])
639    apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A]
640                     nn_integral_measure_pmf_finite)
641    apply (subst emeasure_bind[OF _ _ X])
642    apply measurable
643    apply (subst emeasure_bind[OF _ _ X])
644    apply measurable
645    done
646qed
647
648lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A"
649  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
650
651lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B"
652  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
653
654lemma nn_integral_pmf':
655  "inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)"
656  by (subst nn_integral_bij_count_space[where g=f and B="f`A"])
657     (auto simp: bij_betw_def nn_integral_pmf)
658
659lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"
660  using pmf_nonneg[of M p] by arith
661
662lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"
663  using pmf_nonneg[of M p] by arith+
664
665lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"
666  unfolding set_pmf_iff by simp
667
668lemma 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"
669  by (auto simp: pmf.rep_eq map_pmf_rep_eq measure_distr AE_measure_pmf_iff inj_onD
670           intro!: measure_pmf.finite_measure_eq_AE)
671
672lemma pmf_map_inj': "inj f \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"
673apply(cases "x \<in> set_pmf M")
674 apply(simp add: pmf_map_inj[OF subset_inj_on])
675apply(simp add: pmf_eq_0_set_pmf[symmetric])
676apply(auto simp add: pmf_eq_0_set_pmf dest: injD)
677done
678
679lemma pmf_map_outside: "x \<notin> f ` set_pmf M \<Longrightarrow> pmf (map_pmf f M) x = 0"
680  unfolding pmf_eq_0_set_pmf by simp
681
682lemma measurable_set_pmf[measurable]: "Measurable.pred (count_space UNIV) (\<lambda>x. x \<in> set_pmf M)"
683  by simp
684
685
686subsection \<open> PMFs as function \<close>
687
688context
689  fixes f :: "'a \<Rightarrow> real"
690  assumes nonneg: "\<And>x. 0 \<le> f x"
691  assumes prob: "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"
692begin
693
694lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ennreal \<circ> f)"
695proof (intro conjI)
696  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
697    by (simp split: split_indicator)
698  show "AE x in density (count_space UNIV) (ennreal \<circ> f).
699    measure (density (count_space UNIV) (ennreal \<circ> f)) {x} \<noteq> 0"
700    by (simp add: AE_density nonneg measure_def emeasure_density max_def)
701  show "prob_space (density (count_space UNIV) (ennreal \<circ> f))"
702    by standard (simp add: emeasure_density prob)
703qed simp
704
705lemma pmf_embed_pmf: "pmf embed_pmf x = f x"
706proof transfer
707  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
708    by (simp split: split_indicator)
709  fix x show "measure (density (count_space UNIV) (ennreal \<circ> f)) {x} = f x"
710    by transfer (simp add: measure_def emeasure_density nonneg max_def)
711qed
712
713lemma set_embed_pmf: "set_pmf embed_pmf = {x. f x \<noteq> 0}"
714by(auto simp add: set_pmf_eq pmf_embed_pmf)
715
716end
717
718lemma embed_pmf_transfer:
719  "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ennreal \<circ> f)) embed_pmf"
720  by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer)
721
722lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)"
723proof (transfer, elim conjE)
724  fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
725  assume "prob_space M" then interpret prob_space M .
726  show "M = density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))"
727  proof (rule measure_eqI)
728    fix A :: "'a set"
729    have "(\<integral>\<^sup>+ x. ennreal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
730      (\<integral>\<^sup>+ x. emeasure M {x} * indicator (A \<inter> {x. measure M {x} \<noteq> 0}) x \<partial>count_space UNIV)"
731      by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator)
732    also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} \<partial>count_space (A \<inter> {x. measure M {x} \<noteq> 0}))"
733      by (subst nn_integral_restrict_space[symmetric]) (auto simp: restrict_count_space)
734    also have "\<dots> = emeasure M (\<Union>x\<in>(A \<inter> {x. measure M {x} \<noteq> 0}). {x})"
735      by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support)
736         (auto simp: disjoint_family_on_def)
737    also have "\<dots> = emeasure M A"
738      using ae by (intro emeasure_eq_AE) auto
739    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))) A"
740      using emeasure_space_1 by (simp add: emeasure_density)
741  qed simp
742qed
743
744lemma td_pmf_embed_pmf:
745  "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1}"
746  unfolding type_definition_def
747proof safe
748  fix p :: "'a pmf"
749  have "(\<integral>\<^sup>+ x. 1 \<partial>measure_pmf p) = 1"
750    using measure_pmf.emeasure_space_1[of p] by simp
751  then show *: "(\<integral>\<^sup>+ x. ennreal (pmf p x) \<partial>count_space UNIV) = 1"
752    by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const)
753
754  show "embed_pmf (pmf p) = p"
755    by (intro measure_pmf_inject[THEN iffD1])
756       (simp add: * embed_pmf.rep_eq pmf_nonneg measure_pmf_eq_density[of p] comp_def)
757next
758  fix f :: "'a \<Rightarrow> real" assume "\<forall>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"
759  then show "pmf (embed_pmf f) = f"
760    by (auto intro!: pmf_embed_pmf)
761qed (rule pmf_nonneg)
762
763end
764
765lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ennreal (pmf p x) * f x \<partial>count_space UNIV"
766by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg)
767
768lemma integral_measure_pmf:
769  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
770  assumes A: "finite A"
771  shows "(\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A) \<Longrightarrow> (LINT x|M. f x) = (\<Sum>a\<in>A. pmf M a *\<^sub>R f a)"
772  unfolding measure_pmf_eq_density
773  apply (simp add: integral_density)
774  apply (subst lebesgue_integral_count_space_finite_support)
775  apply (auto intro!: finite_subset[OF _ \<open>finite A\<close>] sum.mono_neutral_left simp: pmf_eq_0_set_pmf)
776  done
777
778lemma expectation_return_pmf [simp]:
779  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
780  shows "measure_pmf.expectation (return_pmf x) f = f x"
781  by (subst integral_measure_pmf[of "{x}"]) simp_all
782
783lemma pmf_expectation_bind:
784  fixes p :: "'a pmf" and f :: "'a \<Rightarrow> 'b pmf"
785    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
786  assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))" "set_pmf p \<subseteq> A"
787  shows "measure_pmf.expectation (p \<bind> f) h =
788           (\<Sum>a\<in>A. pmf p a *\<^sub>R measure_pmf.expectation (f a) h)"
789proof -
790  have "measure_pmf.expectation (p \<bind> f) h = (\<Sum>a\<in>(\<Union>x\<in>A. set_pmf (f x)). pmf (p \<bind> f) a *\<^sub>R h a)"
791    using assms by (intro integral_measure_pmf) auto
792  also have "\<dots> = (\<Sum>x\<in>(\<Union>x\<in>A. set_pmf (f x)). (\<Sum>a\<in>A. (pmf p a * pmf (f a) x) *\<^sub>R h x))"
793  proof (intro sum.cong refl, goal_cases)
794    case (1 x)
795    thus ?case
796      by (subst pmf_bind, subst integral_measure_pmf[of A])
797         (insert assms, auto simp: scaleR_sum_left)
798  qed
799  also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R (\<Sum>i\<in>(\<Union>x\<in>A. set_pmf (f x)). pmf (f j) i *\<^sub>R h i))"
800    by (subst sum.swap) (simp add: scaleR_sum_right)
801  also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R measure_pmf.expectation (f j) h)"
802  proof (intro sum.cong refl, goal_cases)
803    case (1 x)
804    thus ?case
805      by (subst integral_measure_pmf[of "(\<Union>x\<in>A. set_pmf (f x))"])
806         (insert assms, auto simp: scaleR_sum_left)
807  qed
808  finally show ?thesis .
809qed
810
811lemma continuous_on_LINT_pmf: \<comment> \<open>This is dominated convergence!?\<close>
812  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::{banach, second_countable_topology}"
813  assumes f: "\<And>i. i \<in> set_pmf M \<Longrightarrow> continuous_on A (f i)"
814    and bnd: "\<And>a i. a \<in> A \<Longrightarrow> i \<in> set_pmf M \<Longrightarrow> norm (f i a) \<le> B"
815  shows "continuous_on A (\<lambda>a. LINT i|M. f i a)"
816proof cases
817  assume "finite M" with f show ?thesis
818    using integral_measure_pmf[OF \<open>finite M\<close>]
819    by (subst integral_measure_pmf[OF \<open>finite M\<close>])
820       (auto intro!: continuous_on_sum continuous_on_scaleR continuous_on_const)
821next
822  assume "infinite M"
823  let ?f = "\<lambda>i x. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) x"
824
825  show ?thesis
826  proof (rule uniform_limit_theorem)
827    show "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>a. \<Sum>i<n. ?f i a)"
828      by (intro always_eventually allI continuous_on_sum continuous_on_scaleR continuous_on_const f
829                from_nat_into set_pmf_not_empty)
830    show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. LINT i|M. f i a) sequentially"
831    proof (subst uniform_limit_cong[where g="\<lambda>n a. \<Sum>i<n. ?f i a"])
832      fix a assume "a \<in> A"
833      have 1: "(LINT i|M. f i a) = (LINT i|map_pmf (to_nat_on M) M. f (from_nat_into M i) a)"
834        by (auto intro!: integral_cong_AE AE_pmfI)
835      have 2: "\<dots> = (LINT i|count_space UNIV. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) a)"
836        by (simp add: measure_pmf_eq_density integral_density)
837      have "(\<lambda>n. ?f n a) sums (LINT i|M. f i a)"
838        unfolding 1 2
839      proof (intro sums_integral_count_space_nat)
840        have A: "integrable M (\<lambda>i. f i a)"
841          using \<open>a\<in>A\<close> by (auto intro!: measure_pmf.integrable_const_bound AE_pmfI bnd)
842        have "integrable (map_pmf (to_nat_on M) M) (\<lambda>i. f (from_nat_into M i) a)"
843          by (auto simp add: map_pmf_rep_eq integrable_distr_eq intro!: AE_pmfI integrable_cong_AE_imp[OF A])
844        then show "integrable (count_space UNIV) (\<lambda>n. ?f n a)"
845          by (simp add: measure_pmf_eq_density integrable_density)
846      qed
847      then show "(LINT i|M. f i a) = (\<Sum> n. ?f n a)"
848        by (simp add: sums_unique)
849    next
850      show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. (\<Sum> n. ?f n a)) sequentially"
851      proof (rule Weierstrass_m_test)
852        fix n a assume "a\<in>A"
853        then show "norm (?f n a) \<le> pmf (map_pmf (to_nat_on M) M) n * B"
854          using bnd by (auto intro!: mult_mono simp: from_nat_into set_pmf_not_empty)
855      next
856        have "integrable (map_pmf (to_nat_on M) M) (\<lambda>n. B)"
857          by auto
858        then show "summable (\<lambda>n. pmf (map_pmf (to_nat_on (set_pmf M)) M) n * B)"
859          by (fastforce simp add: measure_pmf_eq_density integrable_density integrable_count_space_nat_iff summable_mult2)
860      qed
861    qed simp
862  qed simp
863qed
864
865lemma continuous_on_LBINT:
866  fixes f :: "real \<Rightarrow> real"
867  assumes f: "\<And>b. a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f"
868  shows "continuous_on UNIV (\<lambda>b. LBINT x:{a..b}. f x)"
869proof (subst set_borel_integral_eq_integral)
870  { fix b :: real assume "a \<le> b"
871    from f[OF this] have "continuous_on {a..b} (\<lambda>b. integral {a..b} f)"
872      by (intro indefinite_integral_continuous_1 set_borel_integral_eq_integral) }
873  note * = this
874
875  have "continuous_on (\<Union>b\<in>{a..}. {a <..< b}) (\<lambda>b. integral {a..b} f)"
876  proof (intro continuous_on_open_UN)
877    show "b \<in> {a..} \<Longrightarrow> continuous_on {a<..<b} (\<lambda>b. integral {a..b} f)" for b
878      using *[of b] by (rule continuous_on_subset) auto
879  qed simp
880  also have "(\<Union>b\<in>{a..}. {a <..< b}) = {a <..}"
881    by (auto simp: lt_ex gt_ex less_imp_le) (simp add: Bex_def less_imp_le gt_ex cong: rev_conj_cong)
882  finally have "continuous_on {a+1 ..} (\<lambda>b. integral {a..b} f)"
883    by (rule continuous_on_subset) auto
884  moreover have "continuous_on {a..a+1} (\<lambda>b. integral {a..b} f)"
885    by (rule *) simp
886  moreover
887  have "x \<le> a \<Longrightarrow> {a..x} = (if a = x then {a} else {})" for x
888    by auto
889  then have "continuous_on {..a} (\<lambda>b. integral {a..b} f)"
890    by (subst continuous_on_cong[OF refl, where g="\<lambda>x. 0"]) (auto intro!: continuous_on_const)
891  ultimately have "continuous_on ({..a} \<union> {a..a+1} \<union> {a+1 ..}) (\<lambda>b. integral {a..b} f)"
892    by (intro continuous_on_closed_Un) auto
893  also have "{..a} \<union> {a..a+1} \<union> {a+1 ..} = UNIV"
894    by auto
895  finally show "continuous_on UNIV (\<lambda>b. integral {a..b} f)"
896    by auto
897next
898  show "set_integrable lborel {a..b} f" for b
899    using f by (cases "a \<le> b") auto
900qed
901
902locale pmf_as_function
903begin
904
905setup_lifting td_pmf_embed_pmf
906
907lemma set_pmf_transfer[transfer_rule]:
908  assumes "bi_total A"
909  shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
910  using \<open>bi_total A\<close>
911  by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff)
912     metis+
913
914end
915
916context
917begin
918
919interpretation pmf_as_function .
920
921lemma pmf_eqI: "(\<And>i. pmf M i = pmf N i) \<Longrightarrow> M = N"
922  by transfer auto
923
924lemma pmf_eq_iff: "M = N \<longleftrightarrow> (\<forall>i. pmf M i = pmf N i)"
925  by (auto intro: pmf_eqI)
926
927lemma pmf_neq_exists_less:
928  assumes "M \<noteq> N"
929  shows   "\<exists>x. pmf M x < pmf N x"
930proof (rule ccontr)
931  assume "\<not>(\<exists>x. pmf M x < pmf N x)"
932  hence ge: "pmf M x \<ge> pmf N x" for x by (auto simp: not_less)
933  from assms obtain x where "pmf M x \<noteq> pmf N x" by (auto simp: pmf_eq_iff)
934  with ge[of x] have gt: "pmf M x > pmf N x" by simp
935  have "1 = measure (measure_pmf M) UNIV" by simp
936  also have "\<dots> = measure (measure_pmf N) {x} + measure (measure_pmf N) (UNIV - {x})"
937    by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all
938  also from gt have "measure (measure_pmf N) {x} < measure (measure_pmf M) {x}"
939    by (simp add: measure_pmf_single)
940  also have "measure (measure_pmf N) (UNIV - {x}) \<le> measure (measure_pmf M) (UNIV - {x})"
941    by (subst (1 2) integral_pmf [symmetric])
942       (intro integral_mono integrable_pmf, simp_all add: ge)
943  also have "measure (measure_pmf M) {x} + \<dots> = 1"
944    by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all
945  finally show False by simp_all
946qed
947
948lemma 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))"
949  unfolding pmf_eq_iff pmf_bind
950proof
951  fix i
952  interpret B: prob_space "restrict_space B B"
953    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
954       (auto simp: AE_measure_pmf_iff)
955  interpret A: prob_space "restrict_space A A"
956    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
957       (auto simp: AE_measure_pmf_iff)
958
959  interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B"
960    by unfold_locales
961
962  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)"
963    by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict)
964  also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)"
965    by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
966              countable_set_pmf borel_measurable_count_space)
967  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)"
968    by (rule AB.Fubini_integral[symmetric])
969       (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2
970             simp: pmf_nonneg pmf_le_1 measurable_restrict_space1)
971  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)"
972    by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
973              countable_set_pmf borel_measurable_count_space)
974  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)"
975    by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict[symmetric])
976  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)" .
977qed
978
979lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"
980proof (safe intro!: pmf_eqI)
981  fix a :: "'a" and b :: "'b"
982  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ennreal)"
983    by (auto split: split_indicator)
984
985  have "ennreal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
986         ennreal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
987    unfolding pmf_pair ennreal_pmf_map
988    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg
989                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
990  then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"
991    by (simp add: pmf_nonneg)
992qed
993
994lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"
995proof (safe intro!: pmf_eqI)
996  fix a :: "'a" and b :: "'b"
997  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ennreal)"
998    by (auto split: split_indicator)
999
1000  have "ennreal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
1001         ennreal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
1002    unfolding pmf_pair ennreal_pmf_map
1003    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg
1004                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
1005  then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"
1006    by (simp add: pmf_nonneg)
1007qed
1008
1009lemma 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)"
1010  by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta')
1011
1012end
1013
1014lemma pair_return_pmf1: "pair_pmf (return_pmf x) y = map_pmf (Pair x) y"
1015by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)
1016
1017lemma pair_return_pmf2: "pair_pmf x (return_pmf y) = map_pmf (\<lambda>x. (x, y)) x"
1018by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)
1019
1020lemma pair_pair_pmf: "pair_pmf (pair_pmf u v) w = map_pmf (\<lambda>(x, (y, z)). ((x, y), z)) (pair_pmf u (pair_pmf v w))"
1021by(simp add: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf)
1022
1023lemma pair_commute_pmf: "pair_pmf x y = map_pmf (\<lambda>(x, y). (y, x)) (pair_pmf y x)"
1024unfolding pair_pmf_def by(subst bind_commute_pmf)(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)
1025
1026lemma set_pmf_subset_singleton: "set_pmf p \<subseteq> {x} \<longleftrightarrow> p = return_pmf x"
1027proof(intro iffI pmf_eqI)
1028  fix i
1029  assume x: "set_pmf p \<subseteq> {x}"
1030  hence *: "set_pmf p = {x}" using set_pmf_not_empty[of p] by auto
1031  have "ennreal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
1032  also have "\<dots> = \<integral>\<^sup>+ i. 1 \<partial>p" by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
1033  also have "\<dots> = 1" by simp
1034  finally show "pmf p i = pmf (return_pmf x) i" using x
1035    by(auto split: split_indicator simp add: pmf_eq_0_set_pmf)
1036qed auto
1037
1038lemma bind_eq_return_pmf:
1039  "bind_pmf p f = return_pmf x \<longleftrightarrow> (\<forall>y\<in>set_pmf p. f y = return_pmf x)"
1040  (is "?lhs \<longleftrightarrow> ?rhs")
1041proof(intro iffI strip)
1042  fix y
1043  assume y: "y \<in> set_pmf p"
1044  assume "?lhs"
1045  hence "set_pmf (bind_pmf p f) = {x}" by simp
1046  hence "(\<Union>y\<in>set_pmf p. set_pmf (f y)) = {x}" by simp
1047  hence "set_pmf (f y) \<subseteq> {x}" using y by auto
1048  thus "f y = return_pmf x" by(simp add: set_pmf_subset_singleton)
1049next
1050  assume *: ?rhs
1051  show ?lhs
1052  proof(rule pmf_eqI)
1053    fix i
1054    have "ennreal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ennreal (pmf (f y) i) \<partial>p"
1055      by (simp add: ennreal_pmf_bind)
1056    also have "\<dots> = \<integral>\<^sup>+ y. ennreal (pmf (return_pmf x) i) \<partial>p"
1057      by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
1058    also have "\<dots> = ennreal (pmf (return_pmf x) i)"
1059      by simp
1060    finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i"
1061      by (simp add: pmf_nonneg)
1062  qed
1063qed
1064
1065lemma pmf_False_conv_True: "pmf p False = 1 - pmf p True"
1066proof -
1067  have "pmf p False + pmf p True = measure p {False} + measure p {True}"
1068    by(simp add: measure_pmf_single)
1069  also have "\<dots> = measure p ({False} \<union> {True})"
1070    by(subst measure_pmf.finite_measure_Union) simp_all
1071  also have "{False} \<union> {True} = space p" by auto
1072  finally show ?thesis by simp
1073qed
1074
1075lemma pmf_True_conv_False: "pmf p True = 1 - pmf p False"
1076by(simp add: pmf_False_conv_True)
1077
1078subsection \<open> Conditional Probabilities \<close>
1079
1080lemma measure_pmf_zero_iff: "measure (measure_pmf p) s = 0 \<longleftrightarrow> set_pmf p \<inter> s = {}"
1081  by (subst measure_pmf.prob_eq_0) (auto simp: AE_measure_pmf_iff)
1082
1083context
1084  fixes p :: "'a pmf" and s :: "'a set"
1085  assumes not_empty: "set_pmf p \<inter> s \<noteq> {}"
1086begin
1087
1088interpretation pmf_as_measure .
1089
1090lemma emeasure_measure_pmf_not_zero: "emeasure (measure_pmf p) s \<noteq> 0"
1091proof
1092  assume "emeasure (measure_pmf p) s = 0"
1093  then have "AE x in measure_pmf p. x \<notin> s"
1094    by (rule AE_I[rotated]) auto
1095  with not_empty show False
1096    by (auto simp: AE_measure_pmf_iff)
1097qed
1098
1099lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0"
1100  using emeasure_measure_pmf_not_zero by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
1101
1102lift_definition cond_pmf :: "'a pmf" is
1103  "uniform_measure (measure_pmf p) s"
1104proof (intro conjI)
1105  show "prob_space (uniform_measure (measure_pmf p) s)"
1106    by (intro prob_space_uniform_measure) (auto simp: emeasure_measure_pmf_not_zero)
1107  show "AE x in uniform_measure (measure_pmf p) s. measure (uniform_measure (measure_pmf p) s) {x} \<noteq> 0"
1108    by (simp add: emeasure_measure_pmf_not_zero measure_measure_pmf_not_zero AE_uniform_measure
1109                  AE_measure_pmf_iff set_pmf.rep_eq less_top[symmetric])
1110qed simp
1111
1112lemma pmf_cond: "pmf cond_pmf x = (if x \<in> s then pmf p x / measure p s else 0)"
1113  by transfer (simp add: emeasure_measure_pmf_not_zero pmf.rep_eq)
1114
1115lemma set_cond_pmf[simp]: "set_pmf cond_pmf = set_pmf p \<inter> s"
1116  by (auto simp add: set_pmf_iff pmf_cond measure_measure_pmf_not_zero split: if_split_asm)
1117
1118end
1119
1120lemma measure_pmf_posI: "x \<in> set_pmf p \<Longrightarrow> x \<in> A \<Longrightarrow> measure_pmf.prob p A > 0"
1121  using measure_measure_pmf_not_zero[of p A] by (subst zero_less_measure_iff) blast
1122
1123lemma cond_map_pmf:
1124  assumes "set_pmf p \<inter> f -` s \<noteq> {}"
1125  shows "cond_pmf (map_pmf f p) s = map_pmf f (cond_pmf p (f -` s))"
1126proof -
1127  have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
1128    using assms by auto
1129  { fix x
1130    have "ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =
1131      emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
1132      unfolding ennreal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)
1133    also have "f -` s \<inter> f -` {x} = (if x \<in> s then f -` {x} else {})"
1134      by auto
1135    also have "emeasure p (if x \<in> s then f -` {x} else {}) / emeasure p (f -` s) =
1136      ennreal (pmf (cond_pmf (map_pmf f p) s) x)"
1137      using measure_measure_pmf_not_zero[OF *]
1138      by (simp add: pmf_cond[OF *] ennreal_pmf_map measure_pmf.emeasure_eq_measure
1139                    divide_ennreal pmf_nonneg measure_nonneg zero_less_measure_iff pmf_map)
1140    finally have "ennreal (pmf (cond_pmf (map_pmf f p) s) x) = ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"
1141      by simp }
1142  then show ?thesis
1143    by (intro pmf_eqI) (simp add: pmf_nonneg)
1144qed
1145
1146lemma bind_cond_pmf_cancel:
1147  assumes [simp]: "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
1148  assumes [simp]: "\<And>y. y \<in> set_pmf q \<Longrightarrow> set_pmf p \<inter> {x. R x y} \<noteq> {}"
1149  assumes [simp]: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow> measure q {y. R x y} = measure p {x. R x y}"
1150  shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q"
1151proof (rule pmf_eqI)
1152  fix i
1153  have "ennreal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
1154    (\<integral>\<^sup>+x. ennreal (pmf q i / measure p {x. R x i}) * ennreal (indicator {x. R x i} x) \<partial>p)"
1155    by (auto simp add: ennreal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf pmf_nonneg measure_nonneg
1156             intro!: nn_integral_cong_AE)
1157  also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
1158    by (simp add: pmf_nonneg measure_nonneg zero_ennreal_def[symmetric] ennreal_indicator
1159                  nn_integral_cmult measure_pmf.emeasure_eq_measure ennreal_mult[symmetric])
1160  also have "\<dots> = pmf q i"
1161    by (cases "pmf q i = 0")
1162       (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero pmf_nonneg)
1163  finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
1164    by (simp add: pmf_nonneg)
1165qed
1166
1167subsection \<open> Relator \<close>
1168
1169inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
1170for R p q
1171where
1172  "\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y;
1173     map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk>
1174  \<Longrightarrow> rel_pmf R p q"
1175
1176lemma rel_pmfI:
1177  assumes R: "rel_set R (set_pmf p) (set_pmf q)"
1178  assumes eq: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow>
1179    measure p {x. R x y} = measure q {y. R x y}"
1180  shows "rel_pmf R p q"
1181proof
1182  let ?pq = "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q {y. R x y}) (\<lambda>y. return_pmf (x, y)))"
1183  have "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
1184    using R by (auto simp: rel_set_def)
1185  then show "\<And>x y. (x, y) \<in> set_pmf ?pq \<Longrightarrow> R x y"
1186    by auto
1187  show "map_pmf fst ?pq = p"
1188    by (simp add: map_bind_pmf bind_return_pmf')
1189
1190  show "map_pmf snd ?pq = q"
1191    using R eq
1192    apply (simp add: bind_cond_pmf_cancel map_bind_pmf bind_return_pmf')
1193    apply (rule bind_cond_pmf_cancel)
1194    apply (auto simp: rel_set_def)
1195    done
1196qed
1197
1198lemma rel_pmf_imp_rel_set: "rel_pmf R p q \<Longrightarrow> rel_set R (set_pmf p) (set_pmf q)"
1199  by (force simp add: rel_pmf.simps rel_set_def)
1200
1201lemma rel_pmfD_measure:
1202  assumes rel_R: "rel_pmf R p q" and R: "\<And>a b. R a b \<Longrightarrow> R a y \<longleftrightarrow> R x b"
1203  assumes "x \<in> set_pmf p" "y \<in> set_pmf q"
1204  shows "measure p {x. R x y} = measure q {y. R x y}"
1205proof -
1206  from rel_R obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
1207    and eq: "p = map_pmf fst pq" "q = map_pmf snd pq"
1208    by (auto elim: rel_pmf.cases)
1209  have "measure p {x. R x y} = measure pq {x. R (fst x) y}"
1210    by (simp add: eq map_pmf_rep_eq measure_distr)
1211  also have "\<dots> = measure pq {y. R x (snd y)}"
1212    by (intro measure_pmf.finite_measure_eq_AE)
1213       (auto simp: AE_measure_pmf_iff R dest!: pq)
1214  also have "\<dots> = measure q {y. R x y}"
1215    by (simp add: eq map_pmf_rep_eq measure_distr)
1216  finally show "measure p {x. R x y} = measure q {y. R x y}" .
1217qed
1218
1219lemma rel_pmf_measureD:
1220  assumes "rel_pmf R p q"
1221  shows "measure (measure_pmf p) A \<le> measure (measure_pmf q) {y. \<exists>x\<in>A. R x y}" (is "?lhs \<le> ?rhs")
1222using assms
1223proof cases
1224  fix pq
1225  assume R: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
1226    and p[symmetric]: "map_pmf fst pq = p"
1227    and q[symmetric]: "map_pmf snd pq = q"
1228  have "?lhs = measure (measure_pmf pq) (fst -` A)" by(simp add: p)
1229  also have "\<dots> \<le> measure (measure_pmf pq) {y. \<exists>x\<in>A. R x (snd y)}"
1230    by(rule measure_pmf.finite_measure_mono_AE)(auto 4 3 simp add: AE_measure_pmf_iff dest: R)
1231  also have "\<dots> = ?rhs" by(simp add: q)
1232  finally show ?thesis .
1233qed
1234
1235lemma rel_pmf_iff_measure:
1236  assumes "symp R" "transp R"
1237  shows "rel_pmf R p q \<longleftrightarrow>
1238    rel_set R (set_pmf p) (set_pmf q) \<and>
1239    (\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y})"
1240  by (safe intro!: rel_pmf_imp_rel_set rel_pmfI)
1241     (auto intro!: rel_pmfD_measure dest: sympD[OF \<open>symp R\<close>] transpD[OF \<open>transp R\<close>])
1242
1243lemma quotient_rel_set_disjoint:
1244  "equivp R \<Longrightarrow> C \<in> UNIV // {(x, y). R x y} \<Longrightarrow> rel_set R A B \<Longrightarrow> A \<inter> C = {} \<longleftrightarrow> B \<inter> C = {}"
1245  using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C]
1246  by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE)
1247     (blast dest: equivp_symp)+
1248
1249lemma quotientD: "equiv X R \<Longrightarrow> A \<in> X // R \<Longrightarrow> x \<in> A \<Longrightarrow> A = R `` {x}"
1250  by (metis Image_singleton_iff equiv_class_eq_iff quotientE)
1251
1252lemma rel_pmf_iff_equivp:
1253  assumes "equivp R"
1254  shows "rel_pmf R p q \<longleftrightarrow> (\<forall>C\<in>UNIV // {(x, y). R x y}. measure p C = measure q C)"
1255    (is "_ \<longleftrightarrow>   (\<forall>C\<in>_//?R. _)")
1256proof (subst rel_pmf_iff_measure, safe)
1257  show "symp R" "transp R"
1258    using assms by (auto simp: equivp_reflp_symp_transp)
1259next
1260  fix C assume C: "C \<in> UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)"
1261  assume eq: "\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y}"
1262
1263  show "measure p C = measure q C"
1264  proof (cases "p \<inter> C = {}")
1265    case True
1266    then have "q \<inter> C = {}"
1267      using quotient_rel_set_disjoint[OF assms C R] by simp
1268    with True show ?thesis
1269      unfolding measure_pmf_zero_iff[symmetric] by simp
1270  next
1271    case False
1272    then have "q \<inter> C \<noteq> {}"
1273      using quotient_rel_set_disjoint[OF assms C R] by simp
1274    with False obtain x y where in_set: "x \<in> set_pmf p" "y \<in> set_pmf q" and in_C: "x \<in> C" "y \<in> C"
1275      by auto
1276    then have "R x y"
1277      using in_quotient_imp_in_rel[of UNIV ?R C x y] C assms
1278      by (simp add: equivp_equiv)
1279    with in_set eq have "measure p {x. R x y} = measure q {y. R x y}"
1280      by auto
1281    moreover have "{y. R x y} = C"
1282      using assms \<open>x \<in> C\<close> C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)
1283    moreover have "{x. R x y} = C"
1284      using assms \<open>y \<in> C\<close> C quotientD[of UNIV "?R" C y] sympD[of R]
1285      by (auto simp add: equivp_equiv elim: equivpE)
1286    ultimately show ?thesis
1287      by auto
1288  qed
1289next
1290  assume eq: "\<forall>C\<in>UNIV // ?R. measure p C = measure q C"
1291  show "rel_set R (set_pmf p) (set_pmf q)"
1292    unfolding rel_set_def
1293  proof safe
1294    fix x assume x: "x \<in> set_pmf p"
1295    have "{y. R x y} \<in> UNIV // ?R"
1296      by (auto simp: quotient_def)
1297    with eq have *: "measure q {y. R x y} = measure p {y. R x y}"
1298      by auto
1299    have "measure q {y. R x y} \<noteq> 0"
1300      using x assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
1301    then show "\<exists>y\<in>set_pmf q. R x y"
1302      unfolding measure_pmf_zero_iff by auto
1303  next
1304    fix y assume y: "y \<in> set_pmf q"
1305    have "{x. R x y} \<in> UNIV // ?R"
1306      using assms by (auto simp: quotient_def dest: equivp_symp)
1307    with eq have *: "measure p {x. R x y} = measure q {x. R x y}"
1308      by auto
1309    have "measure p {x. R x y} \<noteq> 0"
1310      using y assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
1311    then show "\<exists>x\<in>set_pmf p. R x y"
1312      unfolding measure_pmf_zero_iff by auto
1313  qed
1314
1315  fix x y assume "x \<in> set_pmf p" "y \<in> set_pmf q" "R x y"
1316  have "{y. R x y} \<in> UNIV // ?R" "{x. R x y} = {y. R x y}"
1317    using assms \<open>R x y\<close> by (auto simp: quotient_def dest: equivp_symp equivp_transp)
1318  with eq show "measure p {x. R x y} = measure q {y. R x y}"
1319    by auto
1320qed
1321
1322bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf
1323proof -
1324  show "map_pmf id = id" by (rule map_pmf_id)
1325  show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose)
1326  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"
1327    by (intro map_pmf_cong refl)
1328
1329  show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"
1330    by (rule pmf_set_map)
1331
1332  show "(card_of (set_pmf p), natLeq) \<in> ordLeq" for p :: "'s pmf"
1333  proof -
1334    have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq"
1335      by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"])
1336         (auto intro: countable_set_pmf)
1337    also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq"
1338      by (metis Field_natLeq card_of_least natLeq_Well_order)
1339    finally show ?thesis .
1340  qed
1341
1342  show "\<And>R. rel_pmf R = (\<lambda>x y. \<exists>z. set_pmf z \<subseteq> {(x, y). R x y} \<and>
1343    map_pmf fst z = x \<and> map_pmf snd z = y)"
1344     by (auto simp add: fun_eq_iff rel_pmf.simps)
1345
1346  show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)"
1347    for R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
1348  proof -
1349    { fix p q r
1350      assume pq: "rel_pmf R p q"
1351        and qr:"rel_pmf S q r"
1352      from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
1353        and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto
1354      from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z"
1355        and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto
1356
1357      define pr where "pr =
1358        bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy})
1359          (\<lambda>yz. return_pmf (fst xy, snd yz)))"
1360      have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
1361        by (force simp: q')
1362
1363      have "rel_pmf (R OO S) p r"
1364      proof (rule rel_pmf.intros)
1365        fix x z assume "(x, z) \<in> pr"
1366        then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr"
1367          by (auto simp: q pr_welldefined pr_def split_beta)
1368        with pq qr show "(R OO S) x z"
1369          by blast
1370      next
1371        have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {yz. fst yz = y}))"
1372          by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_pmf_comp)
1373        then show "map_pmf snd pr = r"
1374          unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) (auto simp: eq_commute)
1375      qed (simp add: pr_def map_bind_pmf split_beta map_pmf_def[symmetric] p map_pmf_comp)
1376    }
1377    then show ?thesis
1378      by(auto simp add: le_fun_def)
1379  qed
1380qed (fact natLeq_card_order natLeq_cinfinite)+
1381
1382lemma map_pmf_idI: "(\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_pmf f p = p"
1383by(simp cong: pmf.map_cong)
1384
1385lemma rel_pmf_conj[simp]:
1386  "rel_pmf (\<lambda>x y. P \<and> Q x y) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
1387  "rel_pmf (\<lambda>x y. Q x y \<and> P) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
1388  using set_pmf_not_empty by (fastforce simp: pmf.in_rel subset_eq)+
1389
1390lemma rel_pmf_top[simp]: "rel_pmf top = top"
1391  by (auto simp: pmf.in_rel[abs_def] fun_eq_iff map_fst_pair_pmf map_snd_pair_pmf
1392           intro: exI[of _ "pair_pmf x y" for x y])
1393
1394lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)"
1395proof safe
1396  fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M"
1397  then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b"
1398    and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq"
1399    by (force elim: rel_pmf.cases)
1400  moreover have "set_pmf (return_pmf x) = {x}"
1401    by simp
1402  with \<open>a \<in> M\<close> have "(x, a) \<in> pq"
1403    by (force simp: eq)
1404  with * show "R x a"
1405    by auto
1406qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"]
1407          simp: map_fst_pair_pmf map_snd_pair_pmf)
1408
1409lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)"
1410  by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1)
1411
1412lemma rel_return_pmf[simp]: "rel_pmf R (return_pmf x1) (return_pmf x2) = R x1 x2"
1413  unfolding rel_pmf_return_pmf2 set_return_pmf by simp
1414
1415lemma rel_pmf_False[simp]: "rel_pmf (\<lambda>x y. False) x y = False"
1416  unfolding pmf.in_rel fun_eq_iff using set_pmf_not_empty by fastforce
1417
1418lemma rel_pmf_rel_prod:
1419  "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'"
1420proof safe
1421  assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
1422  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"
1423    and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'"
1424    by (force elim: rel_pmf.cases)
1425  show "rel_pmf R A B"
1426  proof (rule rel_pmf.intros)
1427    let ?f = "\<lambda>(a, b). (fst a, fst b)"
1428    have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd"
1429      by auto
1430
1431    show "map_pmf fst (map_pmf ?f pq) = A"
1432      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
1433    show "map_pmf snd (map_pmf ?f pq) = B"
1434      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
1435
1436    fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)"
1437    then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq"
1438      by auto
1439    from pq[OF this] show "R a b" ..
1440  qed
1441  show "rel_pmf S A' B'"
1442  proof (rule rel_pmf.intros)
1443    let ?f = "\<lambda>(a, b). (snd a, snd b)"
1444    have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd"
1445      by auto
1446
1447    show "map_pmf fst (map_pmf ?f pq) = A'"
1448      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
1449    show "map_pmf snd (map_pmf ?f pq) = B'"
1450      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
1451
1452    fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)"
1453    then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq"
1454      by auto
1455    from pq[OF this] show "S c d" ..
1456  qed
1457next
1458  assume "rel_pmf R A B" "rel_pmf S A' B'"
1459  then obtain Rpq Spq
1460    where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b"
1461        "map_pmf fst Rpq = A" "map_pmf snd Rpq = B"
1462      and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b"
1463        "map_pmf fst Spq = A'" "map_pmf snd Spq = B'"
1464    by (force elim: rel_pmf.cases)
1465
1466  let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))"
1467  let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)"
1468  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))"
1469    by auto
1470
1471  show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
1472    by (rule rel_pmf.intros[where pq="?pq"])
1473       (auto simp: map_snd_pair_pmf map_fst_pair_pmf map_pmf_comp Rpq Spq
1474                   map_pair)
1475qed
1476
1477lemma rel_pmf_reflI:
1478  assumes "\<And>x. x \<in> set_pmf p \<Longrightarrow> P x x"
1479  shows "rel_pmf P p p"
1480  by (rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. (x, x)) p"])
1481     (auto simp add: pmf.map_comp o_def assms)
1482
1483lemma rel_pmf_bij_betw:
1484  assumes f: "bij_betw f (set_pmf p) (set_pmf q)"
1485  and eq: "\<And>x. x \<in> set_pmf p \<Longrightarrow> pmf p x = pmf q (f x)"
1486  shows "rel_pmf (\<lambda>x y. f x = y) p q"
1487proof(rule rel_pmf.intros)
1488  let ?pq = "map_pmf (\<lambda>x. (x, f x)) p"
1489  show "map_pmf fst ?pq = p" by(simp add: pmf.map_comp o_def)
1490
1491  have "map_pmf f p = q"
1492  proof(rule pmf_eqI)
1493    fix i
1494    show "pmf (map_pmf f p) i = pmf q i"
1495    proof(cases "i \<in> set_pmf q")
1496      case True
1497      with f obtain j where "i = f j" "j \<in> set_pmf p"
1498        by(auto simp add: bij_betw_def image_iff)
1499      thus ?thesis using f by(simp add: bij_betw_def pmf_map_inj eq)
1500    next
1501      case False thus ?thesis
1502        by(subst pmf_map_outside)(auto simp add: set_pmf_iff eq[symmetric])
1503    qed
1504  qed
1505  then show "map_pmf snd ?pq = q" by(simp add: pmf.map_comp o_def)
1506qed auto
1507
1508context
1509begin
1510
1511interpretation pmf_as_measure .
1512
1513definition "join_pmf M = bind_pmf M (\<lambda>x. x)"
1514
1515lemma bind_eq_join_pmf: "bind_pmf M f = join_pmf (map_pmf f M)"
1516  unfolding join_pmf_def bind_map_pmf ..
1517
1518lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id"
1519  by (simp add: join_pmf_def id_def)
1520
1521lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"
1522  unfolding join_pmf_def pmf_bind ..
1523
1524lemma ennreal_pmf_join: "ennreal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"
1525  unfolding join_pmf_def ennreal_pmf_bind ..
1526
1527lemma set_pmf_join_pmf[simp]: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)"
1528  by (simp add: join_pmf_def)
1529
1530lemma join_return_pmf: "join_pmf (return_pmf M) = M"
1531  by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq)
1532
1533lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)"
1534  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf)
1535
1536lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A"
1537  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
1538
1539end
1540
1541lemma rel_pmf_joinI:
1542  assumes "rel_pmf (rel_pmf P) p q"
1543  shows "rel_pmf P (join_pmf p) (join_pmf q)"
1544proof -
1545  from assms obtain pq where p: "p = map_pmf fst pq"
1546    and q: "q = map_pmf snd pq"
1547    and P: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> rel_pmf P x y"
1548    by cases auto
1549  from P obtain PQ
1550    where PQ: "\<And>x y a b. \<lbrakk> (x, y) \<in> set_pmf pq; (a, b) \<in> set_pmf (PQ x y) \<rbrakk> \<Longrightarrow> P a b"
1551    and x: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf fst (PQ x y) = x"
1552    and y: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf snd (PQ x y) = y"
1553    by(metis rel_pmf.simps)
1554
1555  let ?r = "bind_pmf pq (\<lambda>(x, y). PQ x y)"
1556  have "\<And>a b. (a, b) \<in> set_pmf ?r \<Longrightarrow> P a b" by (auto intro: PQ)
1557  moreover have "map_pmf fst ?r = join_pmf p" "map_pmf snd ?r = join_pmf q"
1558    by (simp_all add: p q x y join_pmf_def map_bind_pmf bind_map_pmf split_def cong: bind_pmf_cong)
1559  ultimately show ?thesis ..
1560qed
1561
1562lemma rel_pmf_bindI:
1563  assumes pq: "rel_pmf R p q"
1564  and fg: "\<And>x y. R x y \<Longrightarrow> rel_pmf P (f x) (g y)"
1565  shows "rel_pmf P (bind_pmf p f) (bind_pmf q g)"
1566  unfolding bind_eq_join_pmf
1567  by (rule rel_pmf_joinI)
1568     (auto simp add: pmf.rel_map intro: pmf.rel_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp, OF _ pq] fg)
1569
1570text \<open>
1571  Proof that \<^const>\<open>rel_pmf\<close> preserves orders.
1572  Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism,
1573  Theoretical Computer Science 12(1):19--37, 1980,
1574  \<^url>\<open>https://doi.org/10.1016/0304-3975(80)90003-1\<close>
1575\<close>
1576
1577lemma
1578  assumes *: "rel_pmf R p q"
1579  and refl: "reflp R" and trans: "transp R"
1580  shows measure_Ici: "measure p {y. R x y} \<le> measure q {y. R x y}" (is ?thesis1)
1581  and measure_Ioi: "measure p {y. R x y \<and> \<not> R y x} \<le> measure q {y. R x y \<and> \<not> R y x}" (is ?thesis2)
1582proof -
1583  from * obtain pq
1584    where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
1585    and p: "p = map_pmf fst pq"
1586    and q: "q = map_pmf snd pq"
1587    by cases auto
1588  show ?thesis1 ?thesis2 unfolding p q map_pmf_rep_eq using refl trans
1589    by(auto 4 3 simp add: measure_distr reflpD AE_measure_pmf_iff intro!: measure_pmf.finite_measure_mono_AE dest!: pq elim: transpE)
1590qed
1591
1592lemma rel_pmf_inf:
1593  fixes p q :: "'a pmf"
1594  assumes 1: "rel_pmf R p q"
1595  assumes 2: "rel_pmf R q p"
1596  and refl: "reflp R" and trans: "transp R"
1597  shows "rel_pmf (inf R R\<inverse>\<inverse>) p q"
1598proof (subst rel_pmf_iff_equivp, safe)
1599  show "equivp (inf R R\<inverse>\<inverse>)"
1600    using trans refl by (auto simp: equivp_reflp_symp_transp intro: sympI transpI reflpI dest: transpD reflpD)
1601
1602  fix C assume "C \<in> UNIV // {(x, y). inf R R\<inverse>\<inverse> x y}"
1603  then obtain x where C: "C = {y. R x y \<and> R y x}"
1604    by (auto elim: quotientE)
1605
1606  let ?R = "\<lambda>x y. R x y \<and> R y x"
1607  let ?\<mu>R = "\<lambda>y. measure q {x. ?R x y}"
1608  have "measure p {y. ?R x y} = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})"
1609    by(auto intro!: arg_cong[where f="measure p"])
1610  also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}"
1611    by (rule measure_pmf.finite_measure_Diff) auto
1612  also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}"
1613    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi)
1614  also have "measure p {y. R x y} = measure q {y. R x y}"
1615    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici)
1616  also have "measure q {y. R x y} - measure q {y. R x y \<and> \<not> R y x} =
1617    measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})"
1618    by(rule measure_pmf.finite_measure_Diff[symmetric]) auto
1619  also have "\<dots> = ?\<mu>R x"
1620    by(auto intro!: arg_cong[where f="measure q"])
1621  finally show "measure p C = measure q C"
1622    by (simp add: C conj_commute)
1623qed
1624
1625lemma rel_pmf_antisym:
1626  fixes p q :: "'a pmf"
1627  assumes 1: "rel_pmf R p q"
1628  assumes 2: "rel_pmf R q p"
1629  and refl: "reflp R" and trans: "transp R" and antisym: "antisymp R"
1630  shows "p = q"
1631proof -
1632  from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf)
1633  also have "inf R R\<inverse>\<inverse> = (=)"
1634    using refl antisym by (auto intro!: ext simp add: reflpD dest: antisympD)
1635  finally show ?thesis unfolding pmf.rel_eq .
1636qed
1637
1638lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)"
1639  by (fact pmf.rel_reflp)
1640
1641lemma antisymp_rel_pmf:
1642  "\<lbrakk> reflp R; transp R; antisymp R \<rbrakk>
1643  \<Longrightarrow> antisymp (rel_pmf R)"
1644by(rule antisympI)(blast intro: rel_pmf_antisym)
1645
1646lemma transp_rel_pmf:
1647  assumes "transp R"
1648  shows "transp (rel_pmf R)"
1649  using assms by (fact pmf.rel_transp)
1650
1651
1652subsection \<open> Distributions \<close>
1653
1654context
1655begin
1656
1657interpretation pmf_as_function .
1658
1659subsubsection \<open> Bernoulli Distribution \<close>
1660
1661lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is
1662  "\<lambda>p b. ((\<lambda>p. if b then p else 1 - p) \<circ> min 1 \<circ> max 0) p"
1663  by (auto simp: nn_integral_count_space_finite[where A="{False, True}"] UNIV_bool
1664           split: split_max split_min)
1665
1666lemma pmf_bernoulli_True[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) True = p"
1667  by transfer simp
1668
1669lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p"
1670  by transfer simp
1671
1672lemma set_pmf_bernoulli[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"
1673  by (auto simp add: set_pmf_iff UNIV_bool)
1674
1675lemma nn_integral_bernoulli_pmf[simp]:
1676  assumes [simp]: "0 \<le> p" "p \<le> 1" "\<And>x. 0 \<le> f x"
1677  shows "(\<integral>\<^sup>+x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"
1678  by (subst nn_integral_measure_pmf_support[of UNIV])
1679     (auto simp: UNIV_bool field_simps)
1680
1681lemma integral_bernoulli_pmf[simp]:
1682  assumes [simp]: "0 \<le> p" "p \<le> 1"
1683  shows "(\<integral>x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"
1684  by (subst integral_measure_pmf[of UNIV]) (auto simp: UNIV_bool)
1685
1686lemma pmf_bernoulli_half [simp]: "pmf (bernoulli_pmf (1 / 2)) x = 1 / 2"
1687by(cases x) simp_all
1688
1689lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV"
1690  by (rule measure_eqI)
1691     (simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure ennreal_divide_numeral[symmetric]
1692                    nn_integral_count_space_finite sets_uniform_count_measure divide_ennreal_def mult_ac
1693                    ennreal_of_nat_eq_real_of_nat)
1694
1695subsubsection \<open> Geometric Distribution \<close>
1696
1697context
1698  fixes p :: real assumes p[arith]: "0 < p" "p \<le> 1"
1699begin
1700
1701lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p"
1702proof
1703  have "(\<Sum>i. ennreal (p * (1 - p) ^ i)) = ennreal (p * (1 / (1 - (1 - p))))"
1704    by (intro suminf_ennreal_eq sums_mult geometric_sums) auto
1705  then show "(\<integral>\<^sup>+ x. ennreal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"
1706    by (simp add: nn_integral_count_space_nat field_simps)
1707qed simp
1708
1709lemma pmf_geometric[simp]: "pmf geometric_pmf n = (1 - p)^n * p"
1710  by transfer rule
1711
1712end
1713
1714lemma set_pmf_geometric: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (geometric_pmf p) = UNIV"
1715  by (auto simp: set_pmf_iff)
1716
1717subsubsection \<open> Uniform Multiset Distribution \<close>
1718
1719context
1720  fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
1721begin
1722
1723lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M"
1724proof
1725  show "(\<integral>\<^sup>+ x. ennreal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
1726    using M_not_empty
1727    by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size
1728                  sum_divide_distrib[symmetric])
1729       (auto simp: size_multiset_overloaded_eq intro!: sum.cong)
1730qed simp
1731
1732lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M"
1733  by transfer rule
1734
1735lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_mset M"
1736  by (auto simp: set_pmf_iff)
1737
1738end
1739
1740subsubsection \<open> Uniform Distribution \<close>
1741
1742context
1743  fixes S :: "'a set" assumes S_not_empty: "S \<noteq> {}" and S_finite: "finite S"
1744begin
1745
1746lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S"
1747proof
1748  show "(\<integral>\<^sup>+ x. ennreal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"
1749    using S_not_empty S_finite
1750    by (subst nn_integral_count_space'[of S])
1751       (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_mult[symmetric])
1752qed simp
1753
1754lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S"
1755  by transfer rule
1756
1757lemma set_pmf_of_set[simp]: "set_pmf pmf_of_set = S"
1758  using S_finite S_not_empty by (auto simp: set_pmf_iff)
1759
1760lemma emeasure_pmf_of_set_space[simp]: "emeasure pmf_of_set S = 1"
1761  by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff)
1762
1763lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = sum f S / card S"
1764  by (subst nn_integral_measure_pmf_finite)
1765     (simp_all add: sum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
1766                divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)
1767
1768lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = sum f S / card S"
1769  by (subst integral_measure_pmf[of S]) (auto simp: S_finite sum_divide_distrib)
1770
1771lemma emeasure_pmf_of_set: "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
1772  by (subst nn_integral_indicator[symmetric], simp)
1773     (simp add: S_finite S_not_empty card_gt_0_iff indicator_def sum.If_cases divide_ennreal
1774                ennreal_of_nat_eq_real_of_nat nn_integral_pmf_of_set)
1775
1776lemma measure_pmf_of_set: "measure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
1777  using emeasure_pmf_of_set[of A]
1778  by (simp add: measure_nonneg measure_pmf.emeasure_eq_measure)
1779
1780end
1781
1782lemma pmf_expectation_bind_pmf_of_set:
1783  fixes A :: "'a set" and f :: "'a \<Rightarrow> 'b pmf"
1784    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
1785  assumes "A \<noteq> {}" "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))"
1786  shows "measure_pmf.expectation (pmf_of_set A \<bind> f) h =
1787           (\<Sum>a\<in>A. measure_pmf.expectation (f a) h /\<^sub>R real (card A))"
1788  using assms by (subst pmf_expectation_bind[of A]) (auto simp: field_split_simps)
1789
1790lemma map_pmf_of_set:
1791  assumes "finite A" "A \<noteq> {}"
1792  shows   "map_pmf f (pmf_of_set A) = pmf_of_multiset (image_mset f (mset_set A))"
1793    (is "?lhs = ?rhs")
1794proof (intro pmf_eqI)
1795  fix x
1796  from assms have "ennreal (pmf ?lhs x) = ennreal (pmf ?rhs x)"
1797    by (subst ennreal_pmf_map)
1798       (simp_all add: emeasure_pmf_of_set mset_set_empty_iff count_image_mset Int_commute)
1799  thus "pmf ?lhs x = pmf ?rhs x" by simp
1800qed
1801
1802lemma pmf_bind_pmf_of_set:
1803  assumes "A \<noteq> {}" "finite A"
1804  shows   "pmf (bind_pmf (pmf_of_set A) f) x =
1805             (\<Sum>xa\<in>A. pmf (f xa) x) / real_of_nat (card A)" (is "?lhs = ?rhs")
1806proof -
1807  from assms have "card A > 0" by auto
1808  with assms have "ennreal ?lhs = ennreal ?rhs"
1809    by (subst ennreal_pmf_bind)
1810       (simp_all add: nn_integral_pmf_of_set max_def pmf_nonneg divide_ennreal [symmetric]
1811        sum_nonneg ennreal_of_nat_eq_real_of_nat)
1812  thus ?thesis by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg divide_nonneg_nonneg)
1813qed
1814
1815lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
1816by(rule pmf_eqI)(simp add: indicator_def)
1817
1818lemma map_pmf_of_set_inj:
1819  assumes f: "inj_on f A"
1820  and [simp]: "A \<noteq> {}" "finite A"
1821  shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs")
1822proof(rule pmf_eqI)
1823  fix i
1824  show "pmf ?lhs i = pmf ?rhs i"
1825  proof(cases "i \<in> f ` A")
1826    case True
1827    then obtain i' where "i = f i'" "i' \<in> A" by auto
1828    thus ?thesis using f by(simp add: card_image pmf_map_inj)
1829  next
1830    case False
1831    hence "pmf ?lhs i = 0" by(simp add: pmf_eq_0_set_pmf set_map_pmf)
1832    moreover have "pmf ?rhs i = 0" using False by simp
1833    ultimately show ?thesis by simp
1834  qed
1835qed
1836
1837lemma map_pmf_of_set_bij_betw:
1838  assumes "bij_betw f A B" "A \<noteq> {}" "finite A"
1839  shows   "map_pmf f (pmf_of_set A) = pmf_of_set B"
1840proof -
1841  have "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)"
1842    by (intro map_pmf_of_set_inj assms bij_betw_imp_inj_on[OF assms(1)])
1843  also from assms have "f ` A = B" by (simp add: bij_betw_def)
1844  finally show ?thesis .
1845qed
1846
1847text \<open>
1848  Choosing an element uniformly at random from the union of a disjoint family
1849  of finite non-empty sets with the same size is the same as first choosing a set
1850  from the family uniformly at random and then choosing an element from the chosen set
1851  uniformly at random.
1852\<close>
1853lemma pmf_of_set_UN:
1854  assumes "finite (\<Union>(f ` A))" "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}"
1855          "\<And>x. x \<in> A \<Longrightarrow> card (f x) = n" "disjoint_family_on f A"
1856  shows   "pmf_of_set (\<Union>(f ` A)) = do {x \<leftarrow> pmf_of_set A; pmf_of_set (f x)}"
1857            (is "?lhs = ?rhs")
1858proof (intro pmf_eqI)
1859  fix x
1860  from assms have [simp]: "finite A"
1861    using infinite_disjoint_family_imp_infinite_UNION[of A f] by blast
1862  from assms have "ereal (pmf (pmf_of_set (\<Union>(f ` A))) x) =
1863    ereal (indicator (\<Union>x\<in>A. f x) x / real (card (\<Union>x\<in>A. f x)))"
1864    by (subst pmf_of_set) auto
1865  also from assms have "card (\<Union>x\<in>A. f x) = card A * n"
1866    by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def)
1867  also from assms
1868    have "indicator (\<Union>x\<in>A. f x) x / real \<dots> =
1869              indicator (\<Union>x\<in>A. f x) x / (n * real (card A))"
1870      by (simp add: sum_divide_distrib [symmetric] mult_ac)
1871  also from assms have "indicator (\<Union>x\<in>A. f x) x = (\<Sum>y\<in>A. indicator (f y) x)"
1872    by (intro indicator_UN_disjoint) simp_all
1873  also from assms have "ereal ((\<Sum>y\<in>A. indicator (f y) x) / (real n * real (card A))) =
1874                          ereal (pmf ?rhs x)"
1875    by (subst pmf_bind_pmf_of_set) (simp_all add: sum_divide_distrib)
1876  finally show "pmf ?lhs x = pmf ?rhs x" by simp
1877qed
1878
1879lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV"
1880  by (rule pmf_eqI) simp_all
1881
1882subsubsection \<open> Poisson Distribution \<close>
1883
1884context
1885  fixes rate :: real assumes rate_pos: "0 < rate"
1886begin
1887
1888lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)"
1889proof  (* by Manuel Eberl *)
1890  have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp
1891    by (simp add: field_simps divide_inverse [symmetric])
1892  have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x * exp (-rate) \<partial>count_space UNIV) =
1893          exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)"
1894    by (simp add: field_simps nn_integral_cmult[symmetric] ennreal_mult'[symmetric])
1895  also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)"
1896    by (simp_all add: nn_integral_count_space_nat suminf_ennreal summable ennreal_suminf_neq_top)
1897  also have "... = exp rate" unfolding exp_def
1898    by (simp add: field_simps divide_inverse [symmetric])
1899  also have "ennreal (exp (-rate)) * ennreal (exp rate) = 1"
1900    by (simp add: mult_exp_exp ennreal_mult[symmetric])
1901  finally show "(\<integral>\<^sup>+ x. ennreal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
1902qed (simp add: rate_pos[THEN less_imp_le])
1903
1904lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)"
1905  by transfer rule
1906
1907lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV"
1908  using rate_pos by (auto simp: set_pmf_iff)
1909
1910end
1911
1912subsubsection \<open> Binomial Distribution \<close>
1913
1914context
1915  fixes n :: nat and p :: real assumes p_nonneg: "0 \<le> p" and p_le_1: "p \<le> 1"
1916begin
1917
1918lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)"
1919proof
1920  have "(\<integral>\<^sup>+k. ennreal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =
1921    ennreal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"
1922    using p_le_1 p_nonneg by (subst nn_integral_count_space') auto
1923  also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n"
1924    by (subst binomial_ring) (simp add: atLeast0AtMost)
1925  finally show "(\<integral>\<^sup>+ x. ennreal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"
1926    by simp
1927qed (insert p_nonneg p_le_1, simp)
1928
1929lemma pmf_binomial[simp]: "pmf binomial_pmf k = (n choose k) * p^k * (1 - p)^(n - k)"
1930  by transfer rule
1931
1932lemma set_pmf_binomial_eq: "set_pmf binomial_pmf = (if p = 0 then {0} else if p = 1 then {n} else {.. n})"
1933  using p_nonneg p_le_1 unfolding set_eq_iff set_pmf_iff pmf_binomial by (auto simp: set_pmf_iff)
1934
1935end
1936
1937end
1938
1939lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
1940  by (simp add: set_pmf_binomial_eq)
1941
1942lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
1943  by (simp add: set_pmf_binomial_eq)
1944
1945lemma set_pmf_binomial[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (binomial_pmf n p) = {..n}"
1946  by (simp add: set_pmf_binomial_eq)
1947
1948context includes lifting_syntax
1949begin
1950
1951lemma bind_pmf_parametric [transfer_rule]:
1952  "(rel_pmf A ===> (A ===> rel_pmf B) ===> rel_pmf B) bind_pmf bind_pmf"
1953by(blast intro: rel_pmf_bindI dest: rel_funD)
1954
1955lemma return_pmf_parametric [transfer_rule]: "(A ===> rel_pmf A) return_pmf return_pmf"
1956by(rule rel_funI) simp
1957
1958end
1959
1960
1961primrec replicate_pmf :: "nat \<Rightarrow> 'a pmf \<Rightarrow> 'a list pmf" where
1962  "replicate_pmf 0 _ = return_pmf []"
1963| "replicate_pmf (Suc n) p = do {x \<leftarrow> p; xs \<leftarrow> replicate_pmf n p; return_pmf (x#xs)}"
1964
1965lemma replicate_pmf_1: "replicate_pmf 1 p = map_pmf (\<lambda>x. [x]) p"
1966  by (simp add: map_pmf_def bind_return_pmf)
1967
1968lemma set_replicate_pmf:
1969  "set_pmf (replicate_pmf n p) = {xs\<in>lists (set_pmf p). length xs = n}"
1970  by (induction n) (auto simp: length_Suc_conv)
1971
1972lemma replicate_pmf_distrib:
1973  "replicate_pmf (m + n) p =
1974     do {xs \<leftarrow> replicate_pmf m p; ys \<leftarrow> replicate_pmf n p; return_pmf (xs @ ys)}"
1975  by (induction m) (simp_all add: bind_return_pmf bind_return_pmf' bind_assoc_pmf)
1976
1977lemma power_diff':
1978  assumes "b \<le> a"
1979  shows   "x ^ (a - b) = (if x = 0 \<and> a = b then 1 else x ^ a / (x::'a::field) ^ b)"
1980proof (cases "x = 0")
1981  case True
1982  with assms show ?thesis by (cases "a - b") simp_all
1983qed (insert assms, simp_all add: power_diff)
1984
1985
1986lemma binomial_pmf_Suc:
1987  assumes "p \<in> {0..1}"
1988  shows   "binomial_pmf (Suc n) p =
1989             do {b \<leftarrow> bernoulli_pmf p;
1990                 k \<leftarrow> binomial_pmf n p;
1991                 return_pmf ((if b then 1 else 0) + k)}" (is "_ = ?rhs")
1992proof (intro pmf_eqI)
1993  fix k
1994  have A: "indicator {Suc a} (Suc b) = indicator {a} b" for a b
1995    by (simp add: indicator_def)
1996  show "pmf (binomial_pmf (Suc n) p) k = pmf ?rhs k"
1997    by (cases k; cases "k > n")
1998       (insert assms, auto simp: pmf_bind measure_pmf_single A field_split_simps algebra_simps
1999          not_less less_eq_Suc_le [symmetric] power_diff')
2000qed
2001
2002lemma binomial_pmf_0: "p \<in> {0..1} \<Longrightarrow> binomial_pmf 0 p = return_pmf 0"
2003  by (rule pmf_eqI) (simp_all add: indicator_def)
2004
2005lemma binomial_pmf_altdef:
2006  assumes "p \<in> {0..1}"
2007  shows   "binomial_pmf n p = map_pmf (length \<circ> filter id) (replicate_pmf n (bernoulli_pmf p))"
2008  by (induction n)
2009     (insert assms, auto simp: binomial_pmf_Suc map_pmf_def bind_return_pmf bind_assoc_pmf
2010        bind_return_pmf' binomial_pmf_0 intro!: bind_pmf_cong)
2011
2012
2013subsection \<open>PMFs from association lists\<close>
2014
2015definition pmf_of_list ::" ('a \<times> real) list \<Rightarrow> 'a pmf" where
2016  "pmf_of_list xs = embed_pmf (\<lambda>x. sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
2017
2018definition pmf_of_list_wf where
2019  "pmf_of_list_wf xs \<longleftrightarrow> (\<forall>x\<in>set (map snd xs) . x \<ge> 0) \<and> sum_list (map snd xs) = 1"
2020
2021lemma pmf_of_list_wfI:
2022  "(\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0) \<Longrightarrow> sum_list (map snd xs) = 1 \<Longrightarrow> pmf_of_list_wf xs"
2023  unfolding pmf_of_list_wf_def by simp
2024
2025context
2026begin
2027
2028private lemma pmf_of_list_aux:
2029  assumes "\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0"
2030  assumes "sum_list (map snd xs) = 1"
2031  shows "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])) \<partial>count_space UNIV) = 1"
2032proof -
2033  have "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))) \<partial>count_space UNIV) =
2034            (\<integral>\<^sup>+ x. ennreal (sum_list (map (\<lambda>(x',p). indicator {x'} x * p) xs)) \<partial>count_space UNIV)"
2035    apply (intro nn_integral_cong ennreal_cong, subst sum_list_map_filter')
2036    apply (rule arg_cong[where f = sum_list])
2037    apply (auto cong: map_cong)
2038    done
2039  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. (\<integral>\<^sup>+ x. ennreal (indicator {x'} x * p) \<partial>count_space UNIV))"
2040    using assms(1)
2041  proof (induction xs)
2042    case (Cons x xs)
2043    from Cons.prems have "snd x \<ge> 0" by simp
2044    moreover have "b \<ge> 0" if "(a,b) \<in> set xs" for a b
2045      using Cons.prems[of b] that by force
2046    ultimately have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>x # xs. indicator {x'} y * p) \<partial>count_space UNIV) =
2047            (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) +
2048            ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
2049      by (intro nn_integral_cong, subst ennreal_plus [symmetric])
2050         (auto simp: case_prod_unfold indicator_def intro!: sum_list_nonneg)
2051    also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) \<partial>count_space UNIV) +
2052                      (\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
2053      by (intro nn_integral_add)
2054         (force intro!: sum_list_nonneg AE_I2 intro: Cons simp: indicator_def)+
2055    also have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV) =
2056               (\<Sum>(x', p)\<leftarrow>xs. (\<integral>\<^sup>+ y. ennreal (indicator {x'} y * p) \<partial>count_space UNIV))"
2057      using Cons(1) by (intro Cons) simp_all
2058    finally show ?case by (simp add: case_prod_unfold)
2059  qed simp
2060  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. ennreal p * (\<integral>\<^sup>+ x. indicator {x'} x \<partial>count_space UNIV))"
2061    using assms(1)
2062    by (simp cong: map_cong only: case_prod_unfold, subst nn_integral_cmult [symmetric])
2063       (auto intro!: assms(1) simp: max_def times_ereal.simps [symmetric] mult_ac ereal_indicator
2064             simp del: times_ereal.simps)+
2065  also from assms have "\<dots> = sum_list (map snd xs)" by (simp add: case_prod_unfold sum_list_ennreal)
2066  also have "\<dots> = 1" using assms(2) by simp
2067  finally show ?thesis .
2068qed
2069
2070lemma pmf_pmf_of_list:
2071  assumes "pmf_of_list_wf xs"
2072  shows   "pmf (pmf_of_list xs) x = sum_list (map snd (filter (\<lambda>z. fst z = x) xs))"
2073  using assms pmf_of_list_aux[of xs] unfolding pmf_of_list_def pmf_of_list_wf_def
2074  by (subst pmf_embed_pmf) (auto intro!: sum_list_nonneg)
2075
2076end
2077
2078lemma set_pmf_of_list:
2079  assumes "pmf_of_list_wf xs"
2080  shows   "set_pmf (pmf_of_list xs) \<subseteq> set (map fst xs)"
2081proof clarify
2082  fix x assume A: "x \<in> set_pmf (pmf_of_list xs)"
2083  show "x \<in> set (map fst xs)"
2084  proof (rule ccontr)
2085    assume "x \<notin> set (map fst xs)"
2086    hence "[z\<leftarrow>xs . fst z = x] = []" by (auto simp: filter_empty_conv)
2087    with A assms show False by (simp add: pmf_pmf_of_list set_pmf_eq)
2088  qed
2089qed
2090
2091lemma finite_set_pmf_of_list:
2092  assumes "pmf_of_list_wf xs"
2093  shows   "finite (set_pmf (pmf_of_list xs))"
2094  using assms by (rule finite_subset[OF set_pmf_of_list]) simp_all
2095
2096lemma emeasure_Int_set_pmf:
2097  "emeasure (measure_pmf p) (A \<inter> set_pmf p) = emeasure (measure_pmf p) A"
2098  by (rule emeasure_eq_AE) (auto simp: AE_measure_pmf_iff)
2099
2100lemma measure_Int_set_pmf:
2101  "measure (measure_pmf p) (A \<inter> set_pmf p) = measure (measure_pmf p) A"
2102  using emeasure_Int_set_pmf[of p A] by (simp add: Sigma_Algebra.measure_def)
2103
2104lemma measure_prob_cong_0:
2105  assumes "\<And>x. x \<in> A - B \<Longrightarrow> pmf p x = 0"
2106  assumes "\<And>x. x \<in> B - A \<Longrightarrow> pmf p x = 0"
2107  shows   "measure (measure_pmf p) A = measure (measure_pmf p) B"
2108proof -
2109  have "measure_pmf.prob p A = measure_pmf.prob p (A \<inter> set_pmf p)"
2110    by (simp add: measure_Int_set_pmf)
2111  also have "A \<inter> set_pmf p = B \<inter> set_pmf p"
2112    using assms by (auto simp: set_pmf_eq)
2113  also have "measure_pmf.prob p \<dots> = measure_pmf.prob p B"
2114    by (simp add: measure_Int_set_pmf)
2115  finally show ?thesis .
2116qed
2117
2118lemma emeasure_pmf_of_list:
2119  assumes "pmf_of_list_wf xs"
2120  shows   "emeasure (pmf_of_list xs) A = ennreal (sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs)))"
2121proof -
2122  have "emeasure (pmf_of_list xs) A = nn_integral (measure_pmf (pmf_of_list xs)) (indicator A)"
2123    by simp
2124  also from assms
2125    have "\<dots> = (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])))"
2126    by (subst nn_integral_measure_pmf_finite) (simp_all add: finite_set_pmf_of_list pmf_pmf_of_list Int_def)
2127  also from assms
2128    have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. sum_list (map snd [z\<leftarrow>xs . fst z = x]))"
2129    by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)
2130  also have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A.
2131      indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S")
2132    using assms by (intro ennreal_cong sum.cong) (auto simp: pmf_pmf_of_list)
2133  also have "?S = (\<Sum>x\<in>set_pmf (pmf_of_list xs). indicator A x * pmf (pmf_of_list xs) x)"
2134    using assms by (intro sum.mono_neutral_left set_pmf_of_list finite_set_pmf_of_list) auto
2135  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x * pmf (pmf_of_list xs) x)"
2136    using assms by (intro sum.mono_neutral_left set_pmf_of_list) (auto simp: set_pmf_eq)
2137  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x *
2138                      sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
2139    using assms by (simp add: pmf_pmf_of_list)
2140  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). sum_list (map snd (filter (\<lambda>z. fst z = x \<and> x \<in> A) xs)))"
2141    by (intro sum.cong) (auto simp: indicator_def)
2142  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). (\<Sum>xa = 0..<length xs.
2143                     if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
2144    by (intro sum.cong refl, subst sum_list_map_filter', subst sum_list_sum_nth) simp
2145  also have "\<dots> = (\<Sum>xa = 0..<length xs. (\<Sum>x\<in>set (map fst xs).
2146                     if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
2147    by (rule sum.swap)
2148  also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then
2149                     (\<Sum>x\<in>set (map fst xs). if x = fst (xs ! xa) then snd (xs ! xa) else 0) else 0)"
2150    by (auto intro!: sum.cong sum.neutral simp del: sum.delta)
2151  also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then snd (xs ! xa) else 0)"
2152    by (intro sum.cong refl) (simp_all add: sum.delta)
2153  also have "\<dots> = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"
2154    by (subst sum_list_map_filter', subst sum_list_sum_nth) simp_all
2155  finally show ?thesis .
2156qed
2157
2158lemma measure_pmf_of_list:
2159  assumes "pmf_of_list_wf xs"
2160  shows   "measure (pmf_of_list xs) A = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"
2161  using assms unfolding pmf_of_list_wf_def Sigma_Algebra.measure_def
2162  by (subst emeasure_pmf_of_list [OF assms], subst enn2real_ennreal) (auto intro!: sum_list_nonneg)
2163
2164(* TODO Move? *)
2165lemma sum_list_nonneg_eq_zero_iff:
2166  fixes xs :: "'a :: linordered_ab_group_add list"
2167  shows "(\<And>x. x \<in> set xs \<Longrightarrow> x \<ge> 0) \<Longrightarrow> sum_list xs = 0 \<longleftrightarrow> set xs \<subseteq> {0}"
2168proof (induction xs)
2169  case (Cons x xs)
2170  from Cons.prems have "sum_list (x#xs) = 0 \<longleftrightarrow> x = 0 \<and> sum_list xs = 0"
2171    unfolding sum_list_simps by (subst add_nonneg_eq_0_iff) (auto intro: sum_list_nonneg)
2172  with Cons.IH Cons.prems show ?case by simp
2173qed simp_all
2174
2175lemma sum_list_filter_nonzero:
2176  "sum_list (filter (\<lambda>x. x \<noteq> 0) xs) = sum_list xs"
2177  by (induction xs) simp_all
2178(* END MOVE *)
2179
2180lemma set_pmf_of_list_eq:
2181  assumes "pmf_of_list_wf xs" "\<And>x. x \<in> snd ` set xs \<Longrightarrow> x > 0"
2182  shows   "set_pmf (pmf_of_list xs) = fst ` set xs"
2183proof
2184  {
2185    fix x assume A: "x \<in> fst ` set xs" and B: "x \<notin> set_pmf (pmf_of_list xs)"
2186    then obtain y where y: "(x, y) \<in> set xs" by auto
2187    from B have "sum_list (map snd [z\<leftarrow>xs. fst z = x]) = 0"
2188      by (simp add: pmf_pmf_of_list[OF assms(1)] set_pmf_eq)
2189    moreover from y have "y \<in> snd ` {xa \<in> set xs. fst xa = x}" by force
2190    ultimately have "y = 0" using assms(1)
2191      by (subst (asm) sum_list_nonneg_eq_zero_iff) (auto simp: pmf_of_list_wf_def)
2192    with assms(2) y have False by force
2193  }
2194  thus "fst ` set xs \<subseteq> set_pmf (pmf_of_list xs)" by blast
2195qed (insert set_pmf_of_list[OF assms(1)], simp_all)
2196
2197lemma pmf_of_list_remove_zeros:
2198  assumes "pmf_of_list_wf xs"
2199  defines "xs' \<equiv> filter (\<lambda>z. snd z \<noteq> 0) xs"
2200  shows   "pmf_of_list_wf xs'" "pmf_of_list xs' = pmf_of_list xs"
2201proof -
2202  have "map snd [z\<leftarrow>xs . snd z \<noteq> 0] = filter (\<lambda>x. x \<noteq> 0) (map snd xs)"
2203    by (induction xs) simp_all
2204  with assms(1) show wf: "pmf_of_list_wf xs'"
2205    by (auto simp: pmf_of_list_wf_def xs'_def sum_list_filter_nonzero)
2206  have "sum_list (map snd [z\<leftarrow>xs' . fst z = i]) = sum_list (map snd [z\<leftarrow>xs . fst z = i])" for i
2207    unfolding xs'_def by (induction xs) simp_all
2208  with assms(1) wf show "pmf_of_list xs' = pmf_of_list xs"
2209    by (intro pmf_eqI) (simp_all add: pmf_pmf_of_list)
2210qed
2211
2212end
2213