| author | nipkow | 
| Thu, 01 Nov 2018 11:26:38 +0100 | |
| changeset 69218 | 59aefb3b9e95 | 
| parent 68386 | 98cf1c823c48 | 
| child 69313 | b021008c5397 | 
| permissions | -rw-r--r-- | 
| 58606 | 1  | 
(* Title: HOL/Probability/Probability_Mass_Function.thy  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
2  | 
Author: Johannes Hölzl, TU München  | 
| 59023 | 3  | 
Author: Andreas Lochbihler, ETH Zurich  | 
4  | 
*)  | 
|
| 58606 | 5  | 
|
| 59000 | 6  | 
section \<open> Probability mass function \<close>  | 
7  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
8  | 
theory Probability_Mass_Function  | 
| 59000 | 9  | 
imports  | 
10  | 
Giry_Monad  | 
|
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
66192 
diff
changeset
 | 
11  | 
"HOL-Library.Multiset"  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
13  | 
|
| 59664 | 14  | 
lemma AE_emeasure_singleton:  | 
15  | 
  assumes x: "emeasure M {x} \<noteq> 0" and ae: "AE x in M. P x" shows "P x"
 | 
|
16  | 
proof -  | 
|
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  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
25  | 
by (auto simp:) }  | 
| 59664 | 26  | 
then show "P x" by auto  | 
27  | 
qed  | 
|
28  | 
||
29  | 
lemma 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  | 
||
| 59000 | 32  | 
lemma (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))"
 | 
|
35  | 
proof  | 
|
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  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
39  | 
  then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
 | 
| 59000 | 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"
 | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
54  | 
by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure measure_nonneg set_diff_eq cong: conj_cong)  | 
| 59000 | 55  | 
  then show "AE x in M. measure M {x} \<noteq> 0"
 | 
56  | 
by (auto simp: emeasure_eq_measure)  | 
|
57  | 
qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
 | 
|
58  | 
||
| 59664 | 59  | 
subsection \<open> PMF as measure \<close>  | 
| 59000 | 60  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
61  | 
typedef 'a pmf = "{M :: 'a measure. prob_space M \<and> sets M = UNIV \<and> (AE x in M. measure M {x} \<noteq> 0)}"
 | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
62  | 
morphisms measure_pmf Abs_pmf  | 
| 58606 | 63  | 
  by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
 | 
64  | 
(auto intro!: prob_space_uniform_measure AE_uniform_measureI)  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
66  | 
declare [[coercion measure_pmf]]  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
68  | 
lemma prob_space_measure_pmf: "prob_space (measure_pmf p)"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
69  | 
using pmf.measure_pmf[of p] by auto  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
70  | 
|
| 61605 | 71  | 
interpretation measure_pmf: prob_space "measure_pmf M" for M  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
72  | 
by (rule prob_space_measure_pmf)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
73  | 
|
| 61605 | 74  | 
interpretation measure_pmf: subprob_space "measure_pmf M" for M  | 
| 59000 | 75  | 
by (rule prob_space_imp_subprob_space) unfold_locales  | 
76  | 
||
| 59048 | 77  | 
lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)"  | 
78  | 
by unfold_locales  | 
|
79  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
80  | 
locale pmf_as_measure  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
81  | 
begin  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
82  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
83  | 
setup_lifting type_definition_pmf  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
85  | 
end  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
87  | 
context  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
88  | 
begin  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
90  | 
interpretation pmf_as_measure .  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
91  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
92  | 
lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV"  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
93  | 
by transfer blast  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
94  | 
|
| 59048 | 95  | 
lemma sets_measure_pmf_count_space[measurable_cong]:  | 
96  | 
"sets (measure_pmf M) = sets (count_space UNIV)"  | 
|
| 59000 | 97  | 
by simp  | 
98  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
99  | 
lemma space_measure_pmf[simp]: "space (measure_pmf p) = UNIV"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
100  | 
using sets_eq_imp_space_eq[of "measure_pmf p" "count_space UNIV"] by simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
101  | 
|
| 61634 | 102  | 
lemma measure_pmf_UNIV [simp]: "measure (measure_pmf p) UNIV = 1"  | 
103  | 
using measure_pmf.prob_space[of p] by simp  | 
|
104  | 
||
| 59048 | 105  | 
lemma 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  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
108  | 
lemma measurable_pmf_measure1[simp]: "measurable (M :: 'a pmf) N = UNIV \<rightarrow> space N"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
109  | 
by (auto simp: measurable_def)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
111  | 
lemma measurable_pmf_measure2[simp]: "measurable N (M :: 'a pmf) = measurable N (count_space UNIV)"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
112  | 
by (intro measurable_cong_sets) simp_all  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
113  | 
|
| 59664 | 114  | 
lemma 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 _")  | 
|
118  | 
proof -  | 
|
119  | 
have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"  | 
|
120  | 
by (simp add: restrict_count_space)  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
121  | 
|
| 59664 | 122  | 
show ?thesis  | 
123  | 
by (intro measurable_compose_countable'[where f="\<lambda>a b. f (fst b, a)" and g=snd and I=A,  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
124  | 
unfolded prod.collapse] assms)  | 
| 59664 | 125  | 
measurable  | 
126  | 
qed  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
127  | 
|
| 59664 | 128  | 
lemma 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"  | 
|
132  | 
proof -  | 
|
133  | 
have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"  | 
|
134  | 
by (simp add: restrict_count_space)  | 
|
| 59000 | 135  | 
|
| 59664 | 136  | 
show ?thesis  | 
137  | 
by (intro measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A,  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
138  | 
unfolded prod.collapse] assms)  | 
| 59664 | 139  | 
measurable  | 
140  | 
qed  | 
|
141  | 
||
142  | 
lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
 | 
|
143  | 
||
144  | 
lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
 | 
|
145  | 
declare [[coercion set_pmf]]  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
147  | 
lemma AE_measure_pmf: "AE x in (M::'a pmf). x \<in> M"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
148  | 
by transfer simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
150  | 
lemma emeasure_pmf_single_eq_zero_iff:  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
151  | 
fixes M :: "'a pmf"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
152  | 
  shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
153  | 
unfolding set_pmf.rep_eq by (simp add: measure_pmf.emeasure_eq_measure)  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
155  | 
lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"  | 
| 59664 | 156  | 
using AE_measure_singleton[of M] AE_measure_pmf[of M]  | 
157  | 
by (auto simp: set_pmf.rep_eq)  | 
|
158  | 
||
| 61634 | 159  | 
lemma AE_pmfI: "(\<And>y. y \<in> set_pmf M \<Longrightarrow> P y) \<Longrightarrow> almost_everywhere (measure_pmf M) P"  | 
160  | 
by(simp add: AE_measure_pmf_iff)  | 
|
161  | 
||
| 59664 | 162  | 
lemma countable_set_pmf [simp]: "countable (set_pmf p)"  | 
163  | 
by transfer (metis prob_space.finite_measure finite_measure.countable_support)  | 
|
164  | 
||
165  | 
lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
166  | 
by transfer (simp add: less_le)  | 
| 59664 | 167  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
168  | 
lemma pmf_nonneg[simp]: "0 \<le> pmf p x"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
169  | 
by transfer simp  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
170  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
171  | 
lemma pmf_not_neg [simp]: "\<not>pmf p x < 0"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
172  | 
by (simp add: not_less pmf_nonneg)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
173  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
174  | 
lemma pmf_pos [simp]: "pmf p x \<noteq> 0 \<Longrightarrow> pmf p x > 0"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
175  | 
using pmf_nonneg[of p x] by linarith  | 
| 59664 | 176  | 
|
177  | 
lemma pmf_le_1: "pmf p x \<le> 1"  | 
|
178  | 
by (simp add: pmf.rep_eq)  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
180  | 
lemma set_pmf_not_empty: "set_pmf M \<noteq> {}"
 | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
181  | 
using AE_measure_pmf[of M] by (intro notI) simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
183  | 
lemma set_pmf_iff: "x \<in> set_pmf M \<longleftrightarrow> pmf M x \<noteq> 0"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
184  | 
by transfer simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
185  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
186  | 
lemma pmf_positive_iff: "0 < pmf p x \<longleftrightarrow> x \<in> set_pmf p"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
187  | 
unfolding less_le by (simp add: set_pmf_iff)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
188  | 
|
| 59664 | 189  | 
lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
 | 
190  | 
by (auto simp: set_pmf_iff)  | 
|
191  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
192  | 
lemma set_pmf_eq': "set_pmf p = {x. pmf p x > 0}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
193  | 
proof safe  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
194  | 
fix x assume "x \<in> set_pmf p"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
195  | 
hence "pmf p x \<noteq> 0" by (auto simp: set_pmf_eq)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
196  | 
with pmf_nonneg[of p x] show "pmf p x > 0" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
197  | 
qed (auto simp: set_pmf_eq)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
198  | 
|
| 59664 | 199  | 
lemma 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  | 
||
| 60068 | 204  | 
lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
205  | 
using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure pmf_nonneg measure_nonneg)  | 
| 60068 | 206  | 
|
| 59000 | 207  | 
lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"  | 
| 64267 | 208  | 
by (subst emeasure_eq_sum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)  | 
| 59000 | 209  | 
|
| 64267 | 210  | 
lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = sum (pmf M) S"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
211  | 
using emeasure_measure_pmf_finite[of S M]  | 
| 64267 | 212  | 
by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg sum_nonneg pmf_nonneg)  | 
| 59023 | 213  | 
|
| 64267 | 214  | 
lemma sum_pmf_eq_1:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
215  | 
assumes "finite A" "set_pmf p \<subseteq> A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
216  | 
shows "(\<Sum>x\<in>A. pmf p x) = 1"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
217  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
218  | 
have "(\<Sum>x\<in>A. pmf p x) = measure_pmf.prob p A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
219  | 
by (simp add: measure_measure_pmf_finite assms)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
220  | 
also from assms have "\<dots> = 1"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
221  | 
by (subst measure_pmf.prob_eq_1) (auto simp: AE_measure_pmf_iff)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
222  | 
finally show ?thesis .  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
223  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
224  | 
|
| 59000 | 225  | 
lemma nn_integral_measure_pmf_support:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
226  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
| 59000 | 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)"  | 
|
229  | 
proof -  | 
|
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)  | 
|
236  | 
qed  | 
|
237  | 
||
238  | 
lemma nn_integral_measure_pmf_finite:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
239  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
| 59000 | 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  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
243  | 
|
| 59000 | 244  | 
lemma integrable_measure_pmf_finite:  | 
245  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
|
246  | 
shows "finite (set_pmf M) \<Longrightarrow> integrable M f"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
247  | 
by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite ennreal_mult_less_top)  | 
| 59000 | 248  | 
|
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
249  | 
lemma integral_measure_pmf_real:  | 
| 59000 | 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)"  | 
|
252  | 
proof -  | 
|
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)"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
256  | 
by (subst integral_indicator_finite_real)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
257  | 
(auto simp: measure_def emeasure_measure_pmf_finite pmf_nonneg)  | 
| 59000 | 258  | 
finally show ?thesis .  | 
259  | 
qed  | 
|
260  | 
||
261  | 
lemma integrable_pmf: "integrable (count_space X) (pmf M)"  | 
|
262  | 
proof -  | 
|
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  | 
|
| 59023 | 268  | 
by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def)  | 
| 59000 | 269  | 
qed  | 
270  | 
||
271  | 
lemma integral_pmf: "(\<integral>x. pmf M x \<partial>count_space X) = measure M X"  | 
|
272  | 
proof -  | 
|
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  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
284  | 
by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg integral_nonneg pmf_nonneg)  | 
| 59000 | 285  | 
qed  | 
286  | 
||
287  | 
lemma 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  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
292  | 
lemma emeasure_pmf: "emeasure (M::'a pmf) M = 1"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
293  | 
proof -  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
294  | 
have "emeasure (M::'a pmf) M = emeasure (M::'a pmf) (space M)"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
295  | 
by (intro emeasure_eq_AE) (simp_all add: AE_measure_pmf)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
296  | 
then show ?thesis  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
297  | 
using measure_pmf.emeasure_space_1 by simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
298  | 
qed  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
299  | 
|
| 59490 | 300  | 
lemma emeasure_pmf_UNIV [simp]: "emeasure (measure_pmf M) UNIV = 1"  | 
301  | 
using measure_pmf.emeasure_space_1[of M] by simp  | 
|
302  | 
||
| 59023 | 303  | 
lemma in_null_sets_measure_pmfI:  | 
304  | 
  "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
 | 
|
305  | 
using emeasure_eq_0_AE[where ?P="\<lambda>x. x \<in> A" and M="measure_pmf p"]  | 
|
306  | 
by(auto simp add: null_sets_def AE_measure_pmf_iff)  | 
|
307  | 
||
| 59664 | 308  | 
lemma measure_subprob: "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"  | 
309  | 
by (simp add: space_subprob_algebra subprob_space_measure_pmf)  | 
|
310  | 
||
311  | 
subsection \<open> Monad Interpretation \<close>  | 
|
312  | 
||
313  | 
lemma 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  | 
||
317  | 
lemma 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"  | 
|
321  | 
proof (rule measure_eqI)  | 
|
| 62026 | 322  | 
show "sets (measure_pmf x \<bind> A) = sets (measure_pmf x \<bind> B)"  | 
| 59664 | 323  | 
using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra)  | 
324  | 
next  | 
|
| 62026 | 325  | 
fix X assume "X \<in> sets (measure_pmf x \<bind> A)"  | 
| 59664 | 326  | 
then have X: "X \<in> sets N"  | 
327  | 
using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra)  | 
|
| 62026 | 328  | 
show "emeasure (measure_pmf x \<bind> A) X = emeasure (measure_pmf x \<bind> B) X"  | 
| 59664 | 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)  | 
|
332  | 
qed  | 
|
333  | 
||
334  | 
lift_definition bind_pmf :: "'a pmf \<Rightarrow> ('a \<Rightarrow> 'b pmf ) \<Rightarrow> 'b pmf" is bind
 | 
|
335  | 
proof (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)  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
349  | 
|
| 62026 | 350  | 
show "prob_space (f \<bind> g)"  | 
| 59664 | 351  | 
using g by (intro f.prob_space_bind[where S="count_space UNIV"]) auto  | 
| 62026 | 352  | 
then interpret fg: prob_space "f \<bind> g" .  | 
353  | 
show [simp]: "sets (f \<bind> g) = UNIV"  | 
|
| 59664 | 354  | 
using sets_eq_imp_space_eq[OF s_f]  | 
355  | 
by (subst sets_bind[where N="count_space UNIV"]) auto  | 
|
| 62026 | 356  | 
  show "AE x in f \<bind> g. measure (f \<bind> g) {x} \<noteq> 0"
 | 
| 59664 | 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  | 
|
364  | 
qed  | 
|
365  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
366  | 
adhoc_overloading Monad_Syntax.bind bind_pmf  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
367  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
368  | 
lemma ennreal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"  | 
| 59664 | 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  | 
||
373  | 
lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
374  | 
using ennreal_pmf_bind[of N f i]  | 
| 59664 | 375  | 
by (subst (asm) nn_integral_eq_integral)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
376  | 
(auto simp: pmf_nonneg pmf_le_1 pmf_nonneg integral_nonneg  | 
| 59664 | 377  | 
intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])  | 
378  | 
||
379  | 
lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"  | 
|
380  | 
by transfer (simp add: bind_const' prob_space_imp_subprob_space)  | 
|
381  | 
||
| 59665 | 382  | 
lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
383  | 
proof -  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
384  | 
  have "set_pmf (bind_pmf M N) = {x. ennreal (pmf (bind_pmf M N) x) \<noteq> 0}"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
385  | 
by (simp add: set_pmf_eq pmf_nonneg)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
386  | 
also have "\<dots> = (\<Union>M\<in>set_pmf M. set_pmf (N M))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
387  | 
unfolding ennreal_pmf_bind  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
388  | 
by (subst nn_integral_0_iff_AE) (auto simp: AE_measure_pmf_iff pmf_nonneg set_pmf_eq)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
389  | 
finally show ?thesis .  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
390  | 
qed  | 
| 59664 | 391  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
392  | 
lemma bind_pmf_cong [fundef_cong]:  | 
| 59664 | 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"  | 
|
| 61808 | 395  | 
unfolding \<open>p = q\<close>[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq  | 
| 59664 | 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  | 
||
400  | 
lemma 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  | 
||
| 62026 | 404  | 
lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<bind> (\<lambda>x. measure_pmf (f x)))"  | 
| 59664 | 405  | 
by transfer simp  | 
406  | 
||
407  | 
lemma nn_integral_bind_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>bind_pmf M N) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)"  | 
|
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  | 
||
414  | 
lemma 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  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
418  | 
|
| 59664 | 419  | 
lift_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  | 
||
422  | 
lemma 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  | 
||
| 59665 | 427  | 
lemma set_return_pmf[simp]: "set_pmf (return_pmf x) = {x}"
 | 
| 59664 | 428  | 
by transfer (auto simp add: measure_return split: split_indicator)  | 
429  | 
||
430  | 
lemma bind_return_pmf': "bind_pmf N return_pmf = N"  | 
|
431  | 
proof (transfer, clarify)  | 
|
| 62026 | 432  | 
fix N :: "'a measure" assume "sets N = UNIV" then show "N \<bind> return (count_space UNIV) = N"  | 
| 59664 | 433  | 
by (subst return_sets_cong[where N=N]) (simp_all add: bind_return')  | 
434  | 
qed  | 
|
435  | 
||
436  | 
lemma 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  | 
||
441  | 
definition "map_pmf f M = bind_pmf M (\<lambda>x. return_pmf (f x))"  | 
|
442  | 
||
443  | 
lemma 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  | 
||
446  | 
lemma 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  | 
||
449  | 
lemma map_pmf_transfer[transfer_rule]:  | 
|
| 67399 | 450  | 
"rel_fun (=) (rel_fun cr_pmf cr_pmf) (\<lambda>f M. distr M (count_space UNIV) f) map_pmf"  | 
| 59664 | 451  | 
proof -  | 
| 67399 | 452  | 
have "rel_fun (=) (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf)  | 
| 62026 | 453  | 
(\<lambda>f M. M \<bind> (return (count_space UNIV) o f)) map_pmf"  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
454  | 
unfolding map_pmf_def[abs_def] comp_def by transfer_prover  | 
| 59664 | 455  | 
then show ?thesis  | 
456  | 
by (force simp: rel_fun_def cr_pmf_def bind_return_distr)  | 
|
457  | 
qed  | 
|
458  | 
||
459  | 
lemma 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  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
464  | 
lemma map_pmf_id[simp]: "map_pmf id = id"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
465  | 
by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
466  | 
|
| 59053 | 467  | 
lemma map_pmf_ident[simp]: "map_pmf (\<lambda>x. x) = (\<lambda>x. x)"  | 
468  | 
using map_pmf_id unfolding id_def .  | 
|
469  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
470  | 
lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g"  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
471  | 
by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def)  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
472  | 
|
| 59000 | 473  | 
lemma 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  | 
||
| 59664 | 476  | 
lemma 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  | 
||
| 67399 | 479  | 
lemma pmf_set_map: "set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"  | 
| 59665 | 480  | 
by (auto simp add: comp_def fun_eq_iff map_pmf_def)  | 
| 59664 | 481  | 
|
| 59665 | 482  | 
lemma set_map_pmf[simp]: "set_pmf (map_pmf f M) = f`set_pmf M"  | 
| 59664 | 483  | 
using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff)  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
484  | 
|
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
485  | 
lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)"  | 
| 59664 | 486  | 
unfolding map_pmf_rep_eq by (subst emeasure_distr) auto  | 
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
487  | 
|
| 61634 | 488  | 
lemma measure_map_pmf[simp]: "measure (map_pmf f M) X = measure M (f -` X)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
489  | 
using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure measure_nonneg)  | 
| 61634 | 490  | 
|
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
491  | 
lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"  | 
| 59664 | 492  | 
unfolding map_pmf_rep_eq by (intro nn_integral_distr) auto  | 
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
493  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
494  | 
lemma ennreal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
 | 
| 59664 | 495  | 
proof (transfer fixing: f x)  | 
| 59023 | 496  | 
fix p :: "'b measure"  | 
497  | 
presume "prob_space p"  | 
|
498  | 
then interpret prob_space p .  | 
|
499  | 
presume "sets p = UNIV"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
500  | 
  then show "ennreal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
501  | 
by(simp add: measure_distr measurable_def emeasure_eq_measure)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
502  | 
qed simp_all  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
503  | 
|
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
504  | 
lemma pmf_map: "pmf (map_pmf f p) x = measure p (f -` {x})"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
505  | 
proof (transfer fixing: f x)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
506  | 
fix p :: "'b measure"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
507  | 
presume "prob_space p"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
508  | 
then interpret prob_space p .  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
509  | 
presume "sets p = UNIV"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
510  | 
  then show "measure (distr p (count_space UNIV) f) {x} = measure p (f -` {x})"
 | 
| 59023 | 511  | 
by(simp add: measure_distr measurable_def emeasure_eq_measure)  | 
512  | 
qed simp_all  | 
|
513  | 
||
514  | 
lemma nn_integral_pmf: "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = emeasure (measure_pmf p) A"  | 
|
515  | 
proof -  | 
|
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 .  | 
|
525  | 
qed  | 
|
526  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
527  | 
lemma integral_map_pmf[simp]:  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
528  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
529  | 
shows "integral\<^sup>L (map_pmf g p) f = integral\<^sup>L p (\<lambda>x. f (g x))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
530  | 
by (simp add: integral_distr map_pmf_rep_eq)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
531  | 
|
| 
66568
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
532  | 
lemma pmf_abs_summable [intro]: "pmf p abs_summable_on A"  | 
| 67486 | 533  | 
by (rule abs_summable_on_subset[OF _ subset_UNIV])  | 
| 
66568
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
534  | 
(auto simp: abs_summable_on_def integrable_iff_bounded nn_integral_pmf)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
535  | 
|
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
536  | 
lemma measure_pmf_conv_infsetsum: "measure (measure_pmf p) A = infsetsum (pmf p) A"  | 
| 67486 | 537  | 
unfolding infsetsum_def by (simp add: integral_eq_nn_integral nn_integral_pmf measure_def)  | 
| 
66568
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
538  | 
|
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
539  | 
lemma infsetsum_pmf_eq_1:  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
540  | 
assumes "set_pmf p \<subseteq> A"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
541  | 
shows "infsetsum (pmf p) A = 1"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
542  | 
proof -  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
543  | 
have "infsetsum (pmf p) A = lebesgue_integral (count_space UNIV) (pmf p)"  | 
| 
67977
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67601 
diff
changeset
 | 
544  | 
using assms unfolding infsetsum_altdef set_lebesgue_integral_def  | 
| 
66568
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
545  | 
by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def set_pmf_eq)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
546  | 
also have "\<dots> = 1"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
547  | 
by (subst integral_eq_nn_integral) (auto simp: nn_integral_pmf)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
548  | 
finally show ?thesis .  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
549  | 
qed  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
550  | 
|
| 60068 | 551  | 
lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)"  | 
| 59664 | 552  | 
by transfer (simp add: distr_return)  | 
553  | 
||
554  | 
lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c"  | 
|
555  | 
by transfer (auto simp: prob_space.distr_const)  | 
|
556  | 
||
| 60068 | 557  | 
lemma pmf_return [simp]: "pmf (return_pmf x) y = indicator {y} x"
 | 
| 59664 | 558  | 
by transfer (simp add: measure_return)  | 
559  | 
||
560  | 
lemma 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  | 
||
563  | 
lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x"  | 
|
564  | 
unfolding return_pmf.rep_eq by (intro emeasure_return) auto  | 
|
565  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
566  | 
lemma measure_return_pmf [simp]: "measure_pmf.prob (return_pmf x) A = indicator A x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
567  | 
proof -  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
568  | 
have "ennreal (measure_pmf.prob (return_pmf x) A) =  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
569  | 
emeasure (measure_pmf (return_pmf x)) A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
570  | 
by (simp add: measure_pmf.emeasure_eq_measure)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
571  | 
also have "\<dots> = ennreal (indicator A x)" by (simp add: ennreal_indicator)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
572  | 
finally show ?thesis by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
573  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
574  | 
|
| 59664 | 575  | 
lemma return_pmf_inj[simp]: "return_pmf x = return_pmf y \<longleftrightarrow> x = y"  | 
576  | 
by (metis insertI1 set_return_pmf singletonD)  | 
|
577  | 
||
| 59665 | 578  | 
lemma map_pmf_eq_return_pmf_iff:  | 
579  | 
"map_pmf f p = return_pmf x \<longleftrightarrow> (\<forall>y \<in> set_pmf p. f y = x)"  | 
|
580  | 
proof  | 
|
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  | 
|
584  | 
next  | 
|
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  | 
|
588  | 
qed  | 
|
589  | 
||
| 59664 | 590  | 
definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))"  | 
591  | 
||
592  | 
lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b"  | 
|
593  | 
unfolding pair_pmf_def pmf_bind pmf_return  | 
|
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
594  | 
  apply (subst integral_measure_pmf_real[where A="{b}"])
 | 
| 59664 | 595  | 
apply (auto simp: indicator_eq_0_iff)  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
596  | 
  apply (subst integral_measure_pmf_real[where A="{a}"])
 | 
| 64267 | 597  | 
apply (auto simp: indicator_eq_0_iff sum_nonneg_eq_0_iff pmf_nonneg)  | 
| 59664 | 598  | 
done  | 
599  | 
||
| 59665 | 600  | 
lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"  | 
| 59664 | 601  | 
unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto  | 
602  | 
||
603  | 
lemma 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  | 
||
607  | 
lemma nn_integral_pair_pmf': "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"  | 
|
608  | 
proof -  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
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)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
610  | 
by (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
611  | 
also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"  | 
| 59664 | 612  | 
by (simp add: pair_pmf_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
613  | 
also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"  | 
| 59664 | 614  | 
by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
615  | 
finally show ?thesis .  | 
| 59664 | 616  | 
qed  | 
617  | 
||
618  | 
lemma bind_pair_pmf:  | 
|
619  | 
assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)"  | 
|
| 62026 | 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))))"  | 
| 59664 | 621  | 
(is "?L = ?R")  | 
622  | 
proof (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]  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
640  | 
nn_integral_measure_pmf_finite)  | 
| 59664 | 641  | 
apply (subst emeasure_bind[OF _ _ X])  | 
642  | 
apply measurable  | 
|
643  | 
apply (subst emeasure_bind[OF _ _ X])  | 
|
644  | 
apply measurable  | 
|
645  | 
done  | 
|
646  | 
qed  | 
|
647  | 
||
648  | 
lemma 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  | 
||
651  | 
lemma 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  | 
||
654  | 
lemma 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  | 
||
659  | 
lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
660  | 
using pmf_nonneg[of M p] by arith  | 
| 59664 | 661  | 
|
662  | 
lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
663  | 
using pmf_nonneg[of M p] by arith+  | 
| 59664 | 664  | 
|
665  | 
lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"  | 
|
666  | 
unfolding set_pmf_iff by simp  | 
|
667  | 
||
668  | 
lemma pmf_map_inj: "inj_on f (set_pmf M) \<Longrightarrow> x \<in> set_pmf M \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"  | 
|
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  | 
||
| 60068 | 672  | 
lemma pmf_map_inj': "inj f \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"  | 
673  | 
apply(cases "x \<in> set_pmf M")  | 
|
674  | 
apply(simp add: pmf_map_inj[OF subset_inj_on])  | 
|
675  | 
apply(simp add: pmf_eq_0_set_pmf[symmetric])  | 
|
676  | 
apply(auto simp add: pmf_eq_0_set_pmf dest: injD)  | 
|
677  | 
done  | 
|
678  | 
||
679  | 
lemma pmf_map_outside: "x \<notin> f ` set_pmf M \<Longrightarrow> pmf (map_pmf f M) x = 0"  | 
|
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
680  | 
unfolding pmf_eq_0_set_pmf by simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
681  | 
|
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
682  | 
lemma measurable_set_pmf[measurable]: "Measurable.pred (count_space UNIV) (\<lambda>x. x \<in> set_pmf M)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
683  | 
by simp  | 
| 60068 | 684  | 
|
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
685  | 
|
| 59664 | 686  | 
subsection \<open> PMFs as function \<close>  | 
| 59000 | 687  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
688  | 
context  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
689  | 
fixes f :: "'a \<Rightarrow> real"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
690  | 
assumes nonneg: "\<And>x. 0 \<le> f x"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
691  | 
assumes prob: "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
692  | 
begin  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
693  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
694  | 
lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ennreal \<circ> f)"  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
695  | 
proof (intro conjI)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
696  | 
  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
 | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
697  | 
by (simp split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
698  | 
show "AE x in density (count_space UNIV) (ennreal \<circ> f).  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
699  | 
    measure (density (count_space UNIV) (ennreal \<circ> f)) {x} \<noteq> 0"
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents: 
59053 
diff
changeset
 | 
700  | 
by (simp add: AE_density nonneg measure_def emeasure_density max_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
701  | 
show "prob_space (density (count_space UNIV) (ennreal \<circ> f))"  | 
| 61169 | 702  | 
by standard (simp add: emeasure_density prob)  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
703  | 
qed simp  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
704  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
705  | 
lemma pmf_embed_pmf: "pmf embed_pmf x = f x"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
706  | 
proof transfer  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
707  | 
  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
 | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
708  | 
by (simp split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
709  | 
  fix x show "measure (density (count_space UNIV) (ennreal \<circ> f)) {x} = f x"
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents: 
59053 
diff
changeset
 | 
710  | 
by transfer (simp add: measure_def emeasure_density nonneg max_def)  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
711  | 
qed  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
712  | 
|
| 60068 | 713  | 
lemma set_embed_pmf: "set_pmf embed_pmf = {x. f x \<noteq> 0}"
 | 
| 63092 | 714  | 
by(auto simp add: set_pmf_eq pmf_embed_pmf)  | 
| 60068 | 715  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
716  | 
end  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
717  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
718  | 
lemma embed_pmf_transfer:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
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"  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
720  | 
by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
721  | 
|
| 59000 | 722  | 
lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)"  | 
723  | 
proof (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 .  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
726  | 
  show "M = density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))"
 | 
| 59000 | 727  | 
proof (rule measure_eqI)  | 
728  | 
fix A :: "'a set"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
729  | 
    have "(\<integral>\<^sup>+ x. ennreal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
 | 
| 59000 | 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  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
739  | 
    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))) A"
 | 
| 59000 | 740  | 
using emeasure_space_1 by (simp add: emeasure_density)  | 
741  | 
qed simp  | 
|
742  | 
qed  | 
|
743  | 
||
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
744  | 
lemma td_pmf_embed_pmf:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
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}"
 | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
746  | 
unfolding type_definition_def  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
747  | 
proof safe  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
748  | 
fix p :: "'a pmf"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
749  | 
have "(\<integral>\<^sup>+ x. 1 \<partial>measure_pmf p) = 1"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
750  | 
using measure_pmf.emeasure_space_1[of p] by simp  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
751  | 
then show *: "(\<integral>\<^sup>+ x. ennreal (pmf p x) \<partial>count_space UNIV) = 1"  | 
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
752  | 
by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
753  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
754  | 
show "embed_pmf (pmf p) = p"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
755  | 
by (intro measure_pmf_inject[THEN iffD1])  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
756  | 
(simp add: * embed_pmf.rep_eq pmf_nonneg measure_pmf_eq_density[of p] comp_def)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
757  | 
next  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
758  | 
fix f :: "'a \<Rightarrow> real" assume "\<forall>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
759  | 
then show "pmf (embed_pmf f) = f"  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
760  | 
by (auto intro!: pmf_embed_pmf)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
761  | 
qed (rule pmf_nonneg)  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
762  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
763  | 
end  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
764  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
765  | 
lemma 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"  | 
| 60068 | 766  | 
by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg)  | 
767  | 
||
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
768  | 
lemma integral_measure_pmf:  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
769  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
770  | 
assumes A: "finite A"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
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)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
772  | 
unfolding measure_pmf_eq_density  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
773  | 
apply (simp add: integral_density)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
774  | 
apply (subst lebesgue_integral_count_space_finite_support)  | 
| 64267 | 775  | 
apply (auto intro!: finite_subset[OF _ \<open>finite A\<close>] sum.mono_neutral_left simp: pmf_eq_0_set_pmf)  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
776  | 
done  | 
| 67486 | 777  | 
|
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
778  | 
lemma expectation_return_pmf [simp]:  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
779  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
780  | 
shows "measure_pmf.expectation (return_pmf x) f = f x"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
781  | 
  by (subst integral_measure_pmf[of "{x}"]) simp_all
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
782  | 
|
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
783  | 
lemma pmf_expectation_bind:  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
784  | 
fixes p :: "'a pmf" and f :: "'a \<Rightarrow> 'b pmf"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
785  | 
    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
786  | 
assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))" "set_pmf p \<subseteq> A"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
787  | 
shows "measure_pmf.expectation (p \<bind> f) h =  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
788  | 
(\<Sum>a\<in>A. pmf p a *\<^sub>R measure_pmf.expectation (f a) h)"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
789  | 
proof -  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
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)"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
791  | 
using assms by (intro integral_measure_pmf) auto  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
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))"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
793  | 
proof (intro sum.cong refl, goal_cases)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
794  | 
case (1 x)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
795  | 
thus ?case  | 
| 67486 | 796  | 
by (subst pmf_bind, subst integral_measure_pmf[of A])  | 
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
797  | 
(insert assms, auto simp: scaleR_sum_left)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
798  | 
qed  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
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))"  | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66568 
diff
changeset
 | 
800  | 
by (subst sum.swap) (simp add: scaleR_sum_right)  | 
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
801  | 
also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R measure_pmf.expectation (f j) h)"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
802  | 
proof (intro sum.cong refl, goal_cases)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
803  | 
case (1 x)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
804  | 
thus ?case  | 
| 67486 | 805  | 
by (subst integral_measure_pmf[of "(\<Union>x\<in>A. set_pmf (f x))"])  | 
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
806  | 
(insert assms, auto simp: scaleR_sum_left)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
807  | 
qed  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
808  | 
finally show ?thesis .  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
809  | 
qed  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
810  | 
|
| 67226 | 811  | 
lemma continuous_on_LINT_pmf: \<comment> \<open>This is dominated convergence!?\<close>  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
812  | 
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
813  | 
assumes f: "\<And>i. i \<in> set_pmf M \<Longrightarrow> continuous_on A (f i)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
814  | 
and bnd: "\<And>a i. a \<in> A \<Longrightarrow> i \<in> set_pmf M \<Longrightarrow> norm (f i a) \<le> B"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
815  | 
shows "continuous_on A (\<lambda>a. LINT i|M. f i a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
816  | 
proof cases  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
817  | 
assume "finite M" with f show ?thesis  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
818  | 
using integral_measure_pmf[OF \<open>finite M\<close>]  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
819  | 
by (subst integral_measure_pmf[OF \<open>finite M\<close>])  | 
| 64267 | 820  | 
(auto intro!: continuous_on_sum continuous_on_scaleR continuous_on_const)  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
821  | 
next  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
822  | 
assume "infinite M"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
823  | 
let ?f = "\<lambda>i x. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) x"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
824  | 
|
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
825  | 
show ?thesis  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
826  | 
proof (rule uniform_limit_theorem)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
827  | 
show "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>a. \<Sum>i<n. ?f i a)"  | 
| 64267 | 828  | 
by (intro always_eventually allI continuous_on_sum continuous_on_scaleR continuous_on_const f  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
829  | 
from_nat_into set_pmf_not_empty)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
830  | 
show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. LINT i|M. f i a) sequentially"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
831  | 
proof (subst uniform_limit_cong[where g="\<lambda>n a. \<Sum>i<n. ?f i a"])  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
832  | 
fix a assume "a \<in> A"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
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)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
834  | 
by (auto intro!: integral_cong_AE AE_pmfI)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
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)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
836  | 
by (simp add: measure_pmf_eq_density integral_density)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
837  | 
have "(\<lambda>n. ?f n a) sums (LINT i|M. f i a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
838  | 
unfolding 1 2  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
839  | 
proof (intro sums_integral_count_space_nat)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
840  | 
have A: "integrable M (\<lambda>i. f i a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
841  | 
using \<open>a\<in>A\<close> by (auto intro!: measure_pmf.integrable_const_bound AE_pmfI bnd)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
842  | 
have "integrable (map_pmf (to_nat_on M) M) (\<lambda>i. f (from_nat_into M i) a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
843  | 
by (auto simp add: map_pmf_rep_eq integrable_distr_eq intro!: AE_pmfI integrable_cong_AE_imp[OF A])  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
844  | 
then show "integrable (count_space UNIV) (\<lambda>n. ?f n a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
845  | 
by (simp add: measure_pmf_eq_density integrable_density)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
846  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
847  | 
then show "(LINT i|M. f i a) = (\<Sum> n. ?f n a)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
848  | 
by (simp add: sums_unique)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
849  | 
next  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
850  | 
show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. (\<Sum> n. ?f n a)) sequentially"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
851  | 
proof (rule weierstrass_m_test)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
852  | 
fix n a assume "a\<in>A"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
853  | 
then show "norm (?f n a) \<le> pmf (map_pmf (to_nat_on M) M) n * B"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
854  | 
using bnd by (auto intro!: mult_mono simp: from_nat_into set_pmf_not_empty)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
855  | 
next  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
856  | 
have "integrable (map_pmf (to_nat_on M) M) (\<lambda>n. B)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
857  | 
by auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
858  | 
then show "summable (\<lambda>n. pmf (map_pmf (to_nat_on (set_pmf M)) M) n * B)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
859  | 
by (simp add: measure_pmf_eq_density integrable_density integrable_count_space_nat_iff summable_rabs_cancel)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
860  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
861  | 
qed simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
862  | 
qed simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
863  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
864  | 
|
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
865  | 
lemma continuous_on_LBINT:  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
866  | 
fixes f :: "real \<Rightarrow> real"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
867  | 
  assumes f: "\<And>b. a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
868  | 
  shows "continuous_on UNIV (\<lambda>b. LBINT x:{a..b}. f x)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
869  | 
proof (subst set_borel_integral_eq_integral)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
870  | 
  { fix b :: real assume "a \<le> b"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
871  | 
    from f[OF this] have "continuous_on {a..b} (\<lambda>b. integral {a..b} f)"
 | 
| 
66192
 
e5b84854baa4
A few renamings and several tidied-up proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
66089 
diff
changeset
 | 
872  | 
by (intro indefinite_integral_continuous_1 set_borel_integral_eq_integral) }  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
873  | 
note * = this  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
874  | 
|
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
875  | 
  have "continuous_on (\<Union>b\<in>{a..}. {a <..< b}) (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
876  | 
proof (intro continuous_on_open_UN)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
877  | 
    show "b \<in> {a..} \<Longrightarrow> continuous_on {a<..<b} (\<lambda>b. integral {a..b} f)" for b
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
878  | 
using *[of b] by (rule continuous_on_subset) auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
879  | 
qed simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
880  | 
  also have "(\<Union>b\<in>{a..}. {a <..< b}) = {a <..}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
881  | 
by (auto simp: lt_ex gt_ex less_imp_le) (simp add: Bex_def less_imp_le gt_ex cong: rev_conj_cong)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
882  | 
  finally have "continuous_on {a+1 ..} (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
883  | 
by (rule continuous_on_subset) auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
884  | 
  moreover have "continuous_on {a..a+1} (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
885  | 
by (rule *) simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
886  | 
moreover  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
887  | 
  have "x \<le> a \<Longrightarrow> {a..x} = (if a = x then {a} else {})" for x
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
888  | 
by auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
889  | 
  then have "continuous_on {..a} (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
890  | 
by (subst continuous_on_cong[OF refl, where g="\<lambda>x. 0"]) (auto intro!: continuous_on_const)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
891  | 
  ultimately have "continuous_on ({..a} \<union> {a..a+1} \<union> {a+1 ..}) (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
892  | 
by (intro continuous_on_closed_Un) auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
893  | 
  also have "{..a} \<union> {a..a+1} \<union> {a+1 ..} = UNIV"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
894  | 
by auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
895  | 
  finally show "continuous_on UNIV (\<lambda>b. integral {a..b} f)"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
896  | 
by auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
897  | 
next  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
898  | 
  show "set_integrable lborel {a..b} f" for b
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
899  | 
using f by (cases "a \<le> b") auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
900  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63973 
diff
changeset
 | 
901  | 
|
| 
58587
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
902  | 
locale pmf_as_function  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
903  | 
begin  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
904  | 
|
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
905  | 
setup_lifting td_pmf_embed_pmf  | 
| 
 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 
hoelzl 
parents:  
diff
changeset
 | 
906  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
907  | 
lemma set_pmf_transfer[transfer_rule]:  | 
| 58730 | 908  | 
assumes "bi_total A"  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
909  | 
  shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
 | 
| 61808 | 910  | 
using \<open>bi_total A\<close>  | 
| 58730 | 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  | 
||
| 59000 | 914  | 
end  | 
915  | 
||
916  | 
context  | 
|
917  | 
begin  | 
|
918  | 
||
919  | 
interpretation pmf_as_function .  | 
|
920  | 
||
921  | 
lemma pmf_eqI: "(\<And>i. pmf M i = pmf N i) \<Longrightarrow> M = N"  | 
|
922  | 
by transfer auto  | 
|
923  | 
||
924  | 
lemma pmf_eq_iff: "M = N \<longleftrightarrow> (\<forall>i. pmf M i = pmf N i)"  | 
|
925  | 
by (auto intro: pmf_eqI)  | 
|
926  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
927  | 
lemma pmf_neq_exists_less:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
928  | 
assumes "M \<noteq> N"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
929  | 
shows "\<exists>x. pmf M x < pmf N x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
930  | 
proof (rule ccontr)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
931  | 
assume "\<not>(\<exists>x. pmf M x < pmf N x)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
932  | 
hence ge: "pmf M x \<ge> pmf N x" for x by (auto simp: not_less)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
933  | 
from assms obtain x where "pmf M x \<noteq> pmf N x" by (auto simp: pmf_eq_iff)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
934  | 
with ge[of x] have gt: "pmf M x > pmf N x" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
935  | 
have "1 = measure (measure_pmf M) UNIV" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
936  | 
  also have "\<dots> = measure (measure_pmf N) {x} + measure (measure_pmf N) (UNIV - {x})"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
937  | 
by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
938  | 
  also from gt have "measure (measure_pmf N) {x} < measure (measure_pmf M) {x}"
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
939  | 
by (simp add: measure_pmf_single)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
940  | 
  also have "measure (measure_pmf N) (UNIV - {x}) \<le> measure (measure_pmf M) (UNIV - {x})"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
941  | 
by (subst (1 2) integral_pmf [symmetric])  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
942  | 
(intro integral_mono integrable_pmf, simp_all add: ge)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
943  | 
  also have "measure (measure_pmf M) {x} + \<dots> = 1"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
944  | 
by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
945  | 
finally show False by simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
946  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
947  | 
|
| 59664 | 948  | 
lemma bind_commute_pmf: "bind_pmf A (\<lambda>x. bind_pmf B (C x)) = bind_pmf B (\<lambda>y. bind_pmf A (\<lambda>x. C x y))"  | 
949  | 
unfolding pmf_eq_iff pmf_bind  | 
|
950  | 
proof  | 
|
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)"  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
963  | 
by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict)  | 
| 59664 | 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)"  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
975  | 
by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict[symmetric])  | 
| 59664 | 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)" .  | 
977  | 
qed  | 
|
978  | 
||
979  | 
lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"  | 
|
980  | 
proof (safe intro!: pmf_eqI)  | 
|
981  | 
fix a :: "'a" and b :: "'b"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
982  | 
  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ennreal)"
 | 
| 59664 | 983  | 
by (auto split: split_indicator)  | 
984  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
985  | 
have "ennreal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
986  | 
ennreal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
987  | 
unfolding pmf_pair ennreal_pmf_map  | 
| 59664 | 988  | 
by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
989  | 
emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)  | 
| 59664 | 990  | 
then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
991  | 
by (simp add: pmf_nonneg)  | 
| 59664 | 992  | 
qed  | 
993  | 
||
994  | 
lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"  | 
|
995  | 
proof (safe intro!: pmf_eqI)  | 
|
996  | 
fix a :: "'a" and b :: "'b"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
997  | 
  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ennreal)"
 | 
| 59664 | 998  | 
by (auto split: split_indicator)  | 
999  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1000  | 
have "ennreal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1001  | 
ennreal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1002  | 
unfolding pmf_pair ennreal_pmf_map  | 
| 59664 | 1003  | 
by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1004  | 
emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)  | 
| 59664 | 1005  | 
then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1006  | 
by (simp add: pmf_nonneg)  | 
| 59664 | 1007  | 
qed  | 
1008  | 
||
1009  | 
lemma map_pair: "map_pmf (\<lambda>(a, b). (f a, g b)) (pair_pmf A B) = pair_pmf (map_pmf f A) (map_pmf g B)"  | 
|
1010  | 
by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta')  | 
|
1011  | 
||
| 59000 | 1012  | 
end  | 
1013  | 
||
| 61634 | 1014  | 
lemma pair_return_pmf1: "pair_pmf (return_pmf x) y = map_pmf (Pair x) y"  | 
1015  | 
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)  | 
|
1016  | 
||
1017  | 
lemma pair_return_pmf2: "pair_pmf x (return_pmf y) = map_pmf (\<lambda>x. (x, y)) x"  | 
|
1018  | 
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)  | 
|
1019  | 
||
1020  | 
lemma 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))"  | 
|
1021  | 
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf)  | 
|
1022  | 
||
1023  | 
lemma pair_commute_pmf: "pair_pmf x y = map_pmf (\<lambda>(x, y). (y, x)) (pair_pmf y x)"  | 
|
1024  | 
unfolding pair_pmf_def by(subst bind_commute_pmf)(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)  | 
|
1025  | 
||
1026  | 
lemma set_pmf_subset_singleton: "set_pmf p \<subseteq> {x} \<longleftrightarrow> p = return_pmf x"
 | 
|
1027  | 
proof(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
 | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1031  | 
  have "ennreal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
 | 
| 61634 | 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)  | 
|
1036  | 
qed auto  | 
|
1037  | 
||
1038  | 
lemma 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")  | 
|
1041  | 
proof(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)  | 
|
1049  | 
next  | 
|
1050  | 
assume *: ?rhs  | 
|
1051  | 
show ?lhs  | 
|
1052  | 
proof(rule pmf_eqI)  | 
|
1053  | 
fix i  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1054  | 
have "ennreal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ennreal (pmf (f y) i) \<partial>p"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1055  | 
by (simp add: ennreal_pmf_bind)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1056  | 
also have "\<dots> = \<integral>\<^sup>+ y. ennreal (pmf (return_pmf x) i) \<partial>p"  | 
| 61634 | 1057  | 
by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1058  | 
also have "\<dots> = ennreal (pmf (return_pmf x) i)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1059  | 
by simp  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1060  | 
finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1061  | 
by (simp add: pmf_nonneg)  | 
| 61634 | 1062  | 
qed  | 
1063  | 
qed  | 
|
1064  | 
||
1065  | 
lemma pmf_False_conv_True: "pmf p False = 1 - pmf p True"  | 
|
1066  | 
proof -  | 
|
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  | 
|
1073  | 
qed  | 
|
1074  | 
||
1075  | 
lemma pmf_True_conv_False: "pmf p True = 1 - pmf p False"  | 
|
1076  | 
by(simp add: pmf_False_conv_True)  | 
|
1077  | 
||
| 59664 | 1078  | 
subsection \<open> Conditional Probabilities \<close>  | 
1079  | 
||
| 59670 | 1080  | 
lemma 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  | 
||
| 59664 | 1083  | 
context  | 
1084  | 
fixes p :: "'a pmf" and s :: "'a set"  | 
|
1085  | 
  assumes not_empty: "set_pmf p \<inter> s \<noteq> {}"
 | 
|
1086  | 
begin  | 
|
1087  | 
||
1088  | 
interpretation pmf_as_measure .  | 
|
1089  | 
||
1090  | 
lemma emeasure_measure_pmf_not_zero: "emeasure (measure_pmf p) s \<noteq> 0"  | 
|
1091  | 
proof  | 
|
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)  | 
|
1097  | 
qed  | 
|
1098  | 
||
1099  | 
lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1100  | 
using emeasure_measure_pmf_not_zero by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg)  | 
| 59664 | 1101  | 
|
1102  | 
lift_definition cond_pmf :: "'a pmf" is  | 
|
1103  | 
"uniform_measure (measure_pmf p) s"  | 
|
1104  | 
proof (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  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1109  | 
AE_measure_pmf_iff set_pmf.rep_eq less_top[symmetric])  | 
| 59664 | 1110  | 
qed simp  | 
1111  | 
||
1112  | 
lemma 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  | 
||
| 59665 | 1115  | 
lemma set_cond_pmf[simp]: "set_pmf cond_pmf = set_pmf p \<inter> s"  | 
| 62390 | 1116  | 
by (auto simp add: set_pmf_iff pmf_cond measure_measure_pmf_not_zero split: if_split_asm)  | 
| 59664 | 1117  | 
|
1118  | 
end  | 
|
1119  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1120  | 
lemma measure_pmf_posI: "x \<in> set_pmf p \<Longrightarrow> x \<in> A \<Longrightarrow> measure_pmf.prob p A > 0"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1121  | 
using measure_measure_pmf_not_zero[of p A] by (subst zero_less_measure_iff) blast  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1122  | 
|
| 59664 | 1123  | 
lemma 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))"  | 
|
1126  | 
proof -  | 
|
1127  | 
  have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
 | 
|
| 59665 | 1128  | 
using assms by auto  | 
| 59664 | 1129  | 
  { fix x
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1130  | 
have "ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =  | 
| 59664 | 1131  | 
      emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1132  | 
unfolding ennreal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)  | 
| 59664 | 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) =
 | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1136  | 
ennreal (pmf (cond_pmf (map_pmf f p) s) x)"  | 
| 59664 | 1137  | 
using measure_measure_pmf_not_zero[OF *]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1138  | 
by (simp add: pmf_cond[OF *] ennreal_pmf_map measure_pmf.emeasure_eq_measure  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1139  | 
divide_ennreal pmf_nonneg measure_nonneg zero_less_measure_iff pmf_map)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1140  | 
finally have "ennreal (pmf (cond_pmf (map_pmf f p) s) x) = ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"  | 
| 59664 | 1141  | 
by simp }  | 
1142  | 
then show ?thesis  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1143  | 
by (intro pmf_eqI) (simp add: pmf_nonneg)  | 
| 59664 | 1144  | 
qed  | 
1145  | 
||
1146  | 
lemma bind_cond_pmf_cancel:  | 
|
| 59670 | 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"
 | 
|
| 59664 | 1151  | 
proof (rule pmf_eqI)  | 
| 59670 | 1152  | 
fix i  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1153  | 
  have "ennreal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1154  | 
    (\<integral>\<^sup>+x. ennreal (pmf q i / measure p {x. R x i}) * ennreal (indicator {x. R x i} x) \<partial>p)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1155  | 
by (auto simp add: ennreal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf pmf_nonneg measure_nonneg  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1156  | 
intro!: nn_integral_cong_AE)  | 
| 59670 | 1157  | 
  also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1158  | 
by (simp add: pmf_nonneg measure_nonneg zero_ennreal_def[symmetric] ennreal_indicator  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1159  | 
nn_integral_cmult measure_pmf.emeasure_eq_measure ennreal_mult[symmetric])  | 
| 59670 | 1160  | 
also have "\<dots> = pmf q i"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1161  | 
by (cases "pmf q i = 0")  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1162  | 
(simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero pmf_nonneg)  | 
| 59670 | 1163  | 
  finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1164  | 
by (simp add: pmf_nonneg)  | 
| 59664 | 1165  | 
qed  | 
1166  | 
||
1167  | 
subsection \<open> Relator \<close>  | 
|
1168  | 
||
1169  | 
inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
 | 
|
1170  | 
for R p q  | 
|
1171  | 
where  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1172  | 
"\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y;  | 
| 59664 | 1173  | 
map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk>  | 
1174  | 
\<Longrightarrow> rel_pmf R p q"  | 
|
1175  | 
||
| 59681 | 1176  | 
lemma 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"  | 
|
1181  | 
proof  | 
|
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"  | 
|
| 60068 | 1188  | 
by (simp add: map_bind_pmf bind_return_pmf')  | 
| 59681 | 1189  | 
|
1190  | 
show "map_pmf snd ?pq = q"  | 
|
1191  | 
using R eq  | 
|
| 60068 | 1192  | 
apply (simp add: bind_cond_pmf_cancel map_bind_pmf bind_return_pmf')  | 
| 59681 | 1193  | 
apply (rule bind_cond_pmf_cancel)  | 
1194  | 
apply (auto simp: rel_set_def)  | 
|
1195  | 
done  | 
|
1196  | 
qed  | 
|
1197  | 
||
1198  | 
lemma 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  | 
||
1201  | 
lemma 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}"
 | 
|
1205  | 
proof -  | 
|
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}" .
 | 
|
1217  | 
qed  | 
|
1218  | 
||
| 61634 | 1219  | 
lemma 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")
 | 
|
1222  | 
using assms  | 
|
1223  | 
proof 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 .  | 
|
1233  | 
qed  | 
|
1234  | 
||
| 59681 | 1235  | 
lemma 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  | 
||
1243  | 
lemma 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 = {}"
 | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1245  | 
  using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C]
 | 
| 59681 | 1246  | 
by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE)  | 
1247  | 
(blast dest: equivp_symp)+  | 
|
1248  | 
||
1249  | 
lemma 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  | 
||
1252  | 
lemma 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. _)")  | 
|
1256  | 
proof (subst rel_pmf_iff_measure, safe)  | 
|
1257  | 
show "symp R" "transp R"  | 
|
1258  | 
using assms by (auto simp: equivp_reflp_symp_transp)  | 
|
1259  | 
next  | 
|
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}"
 | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1262  | 
|
| 59681 | 1263  | 
show "measure p C = measure q C"  | 
| 63540 | 1264  | 
  proof (cases "p \<inter> C = {}")
 | 
1265  | 
case True  | 
|
1266  | 
    then have "q \<inter> C = {}"
 | 
|
| 59681 | 1267  | 
using quotient_rel_set_disjoint[OF assms C R] by simp  | 
| 63540 | 1268  | 
with True show ?thesis  | 
| 59681 | 1269  | 
unfolding measure_pmf_zero_iff[symmetric] by simp  | 
1270  | 
next  | 
|
| 63540 | 1271  | 
case False  | 
1272  | 
    then have "q \<inter> C \<noteq> {}"
 | 
|
| 59681 | 1273  | 
using quotient_rel_set_disjoint[OF assms C R] by simp  | 
| 63540 | 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"  | 
| 59681 | 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"
 | 
|
| 61808 | 1282  | 
using assms \<open>x \<in> C\<close> C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)  | 
| 59681 | 1283  | 
    moreover have "{x. R x y} = C"
 | 
| 61808 | 1284  | 
using assms \<open>y \<in> C\<close> C quotientD[of UNIV "?R" C y] sympD[of R]  | 
| 59681 | 1285  | 
by (auto simp add: equivp_equiv elim: equivpE)  | 
1286  | 
ultimately show ?thesis  | 
|
1287  | 
by auto  | 
|
1288  | 
qed  | 
|
1289  | 
next  | 
|
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}"
 | 
|
| 61808 | 1317  | 
using assms \<open>R x y\<close> by (auto simp: quotient_def dest: equivp_symp equivp_transp)  | 
| 59681 | 1318  | 
  with eq show "measure p {x. R x y} = measure q {y. R x y}"
 | 
1319  | 
by auto  | 
|
1320  | 
qed  | 
|
1321  | 
||
| 59664 | 1322  | 
bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf  | 
1323  | 
proof -  | 
|
1324  | 
show "map_pmf id = id" by (rule map_pmf_id)  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1325  | 
show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose)  | 
| 59664 | 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  | 
||
| 67399 | 1329  | 
show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"  | 
| 59664 | 1330  | 
by (rule pmf_set_map)  | 
1331  | 
||
| 60595 | 1332  | 
show "(card_of (set_pmf p), natLeq) \<in> ordLeq" for p :: "'s pmf"  | 
1333  | 
proof -  | 
|
| 59664 | 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)  | 
|
| 60595 | 1339  | 
finally show ?thesis .  | 
1340  | 
qed  | 
|
| 59664 | 1341  | 
|
| 62324 | 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)  | 
|
| 59664 | 1345  | 
|
| 60595 | 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  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1356  | 
|
| 63040 | 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)))"  | 
|
| 60595 | 1360  | 
      have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
 | 
1361  | 
by (force simp: q')  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1362  | 
|
| 60595 | 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  | 
|
| 59664 | 1380  | 
qed (fact natLeq_card_order natLeq_cinfinite)+  | 
1381  | 
||
| 61634 | 1382  | 
lemma map_pmf_idI: "(\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_pmf f p = p"  | 
1383  | 
by(simp cong: pmf.map_cong)  | 
|
1384  | 
||
| 59665 | 1385  | 
lemma 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  | 
||
1390  | 
lemma 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  | 
||
| 59664 | 1394  | 
lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)"  | 
1395  | 
proof 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}"
 | 
|
| 59665 | 1401  | 
by simp  | 
| 61808 | 1402  | 
with \<open>a \<in> M\<close> have "(x, a) \<in> pq"  | 
| 59665 | 1403  | 
by (force simp: eq)  | 
| 59664 | 1404  | 
with * show "R x a"  | 
1405  | 
by auto  | 
|
1406  | 
qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"]  | 
|
| 59665 | 1407  | 
simp: map_fst_pair_pmf map_snd_pair_pmf)  | 
| 59664 | 1408  | 
|
1409  | 
lemma 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  | 
||
1412  | 
lemma 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  | 
||
1415  | 
lemma 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  | 
||
1418  | 
lemma 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'"  | 
|
1420  | 
proof 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"  | 
|
| 59665 | 1438  | 
by auto  | 
| 59664 | 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"  | 
|
| 59665 | 1454  | 
by auto  | 
| 59664 | 1455  | 
from pq[OF this] show "S c d" ..  | 
1456  | 
qed  | 
|
1457  | 
next  | 
|
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"])  | 
|
| 59665 | 1473  | 
(auto simp: map_snd_pair_pmf map_fst_pair_pmf map_pmf_comp Rpq Spq  | 
| 59664 | 1474  | 
map_pair)  | 
1475  | 
qed  | 
|
1476  | 
||
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1477  | 
lemma rel_pmf_reflI:  | 
| 59664 | 1478  | 
assumes "\<And>x. x \<in> set_pmf p \<Longrightarrow> P x x"  | 
1479  | 
shows "rel_pmf P p p"  | 
|
| 59665 | 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)  | 
|
| 59664 | 1482  | 
|
| 61634 | 1483  | 
lemma 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"  | 
|
1487  | 
proof(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)  | 
|
1506  | 
qed auto  | 
|
1507  | 
||
| 59664 | 1508  | 
context  | 
1509  | 
begin  | 
|
1510  | 
||
1511  | 
interpretation pmf_as_measure .  | 
|
1512  | 
||
1513  | 
definition "join_pmf M = bind_pmf M (\<lambda>x. x)"  | 
|
1514  | 
||
1515  | 
lemma bind_eq_join_pmf: "bind_pmf M f = join_pmf (map_pmf f M)"  | 
|
1516  | 
unfolding join_pmf_def bind_map_pmf ..  | 
|
1517  | 
||
1518  | 
lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id"  | 
|
1519  | 
by (simp add: join_pmf_def id_def)  | 
|
1520  | 
||
1521  | 
lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"  | 
|
1522  | 
unfolding join_pmf_def pmf_bind ..  | 
|
1523  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1524  | 
lemma ennreal_pmf_join: "ennreal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1525  | 
unfolding join_pmf_def ennreal_pmf_bind ..  | 
| 59664 | 1526  | 
|
| 59665 | 1527  | 
lemma 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)  | 
|
| 59664 | 1529  | 
|
1530  | 
lemma 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  | 
||
1533  | 
lemma 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  | 
||
1536  | 
lemma 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  | 
||
1539  | 
end  | 
|
1540  | 
||
1541  | 
lemma rel_pmf_joinI:  | 
|
1542  | 
assumes "rel_pmf (rel_pmf P) p q"  | 
|
1543  | 
shows "rel_pmf P (join_pmf p) (join_pmf q)"  | 
|
1544  | 
proof -  | 
|
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  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1549  | 
from P obtain PQ  | 
| 59664 | 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)"  | 
|
| 59665 | 1556  | 
have "\<And>a b. (a, b) \<in> set_pmf ?r \<Longrightarrow> P a b" by (auto intro: PQ)  | 
| 59664 | 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 ..  | 
|
1560  | 
qed  | 
|
1561  | 
||
1562  | 
lemma 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  | 
||
| 61808 | 1570  | 
text \<open>  | 
| 59664 | 1571  | 
  Proof that @{const rel_pmf} preserves orders.
 | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1572  | 
Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism,  | 
| 
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1573  | 
Theoretical Computer Science 12(1):19--37, 1980,  | 
| 
67601
 
b34be3010273
use preferred resolver according to DOI Handbook §3.8
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
67489 
diff
changeset
 | 
1574  | 
\<^url>\<open>https://doi.org/10.1016/0304-3975(80)90003-1\<close>  | 
| 61808 | 1575  | 
\<close>  | 
| 59664 | 1576  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1577  | 
lemma  | 
| 59664 | 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)
 | 
|
1582  | 
proof -  | 
|
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)  | 
|
1590  | 
qed  | 
|
1591  | 
||
1592  | 
lemma 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"  | 
|
| 59681 | 1598  | 
proof (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)  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1601  | 
|
| 59681 | 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  | 
||
| 59670 | 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}"
 | 
|
| 59681 | 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)  | 
|
| 59664 | 1623  | 
qed  | 
1624  | 
||
1625  | 
lemma 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"  | 
|
| 64634 | 1629  | 
and refl: "reflp R" and trans: "transp R" and antisym: "antisymp R"  | 
| 59664 | 1630  | 
shows "p = q"  | 
1631  | 
proof -  | 
|
1632  | 
from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf)  | 
|
| 67399 | 1633  | 
also have "inf R R\<inverse>\<inverse> = (=)"  | 
| 64634 | 1634  | 
using refl antisym by (auto intro!: ext simp add: reflpD dest: antisympD)  | 
| 59664 | 1635  | 
finally show ?thesis unfolding pmf.rel_eq .  | 
1636  | 
qed  | 
|
1637  | 
||
1638  | 
lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)"  | 
|
| 64634 | 1639  | 
by (fact pmf.rel_reflp)  | 
| 59664 | 1640  | 
|
| 64634 | 1641  | 
lemma antisymp_rel_pmf:  | 
1642  | 
"\<lbrakk> reflp R; transp R; antisymp R \<rbrakk>  | 
|
1643  | 
\<Longrightarrow> antisymp (rel_pmf R)"  | 
|
1644  | 
by(rule antisympI)(blast intro: rel_pmf_antisym)  | 
|
| 59664 | 1645  | 
|
1646  | 
lemma transp_rel_pmf:  | 
|
1647  | 
assumes "transp R"  | 
|
1648  | 
shows "transp (rel_pmf R)"  | 
|
| 64634 | 1649  | 
using assms by (fact pmf.rel_transp)  | 
| 59664 | 1650  | 
|
| 67486 | 1651  | 
|
| 59664 | 1652  | 
subsection \<open> Distributions \<close>  | 
1653  | 
||
| 59000 | 1654  | 
context  | 
1655  | 
begin  | 
|
1656  | 
||
1657  | 
interpretation pmf_as_function .  | 
|
1658  | 
||
| 59093 | 1659  | 
subsubsection \<open> Bernoulli Distribution \<close>  | 
1660  | 
||
| 59000 | 1661  | 
lift_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  | 
||
1666  | 
lemma pmf_bernoulli_True[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) True = p"  | 
|
1667  | 
by transfer simp  | 
|
1668  | 
||
1669  | 
lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p"  | 
|
1670  | 
by transfer simp  | 
|
1671  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1672  | 
lemma set_pmf_bernoulli[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"  | 
| 59000 | 1673  | 
by (auto simp add: set_pmf_iff UNIV_bool)  | 
1674  | 
||
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1675  | 
lemma nn_integral_bernoulli_pmf[simp]:  | 
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1676  | 
assumes [simp]: "0 \<le> p" "p \<le> 1" "\<And>x. 0 \<le> f x"  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1677  | 
shows "(\<integral>\<^sup>+x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1678  | 
by (subst nn_integral_measure_pmf_support[of UNIV])  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1679  | 
(auto simp: UNIV_bool field_simps)  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1680  | 
|
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59665 
diff
changeset
 | 
1681  | 
lemma integral_bernoulli_pmf[simp]:  | 
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1682  | 
assumes [simp]: "0 \<le> p" "p \<le> 1"  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1683  | 
shows "(\<integral>x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1684  | 
by (subst integral_measure_pmf[of UNIV]) (auto simp: UNIV_bool)  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1685  | 
|
| 59525 | 1686  | 
lemma pmf_bernoulli_half [simp]: "pmf (bernoulli_pmf (1 / 2)) x = 1 / 2"  | 
1687  | 
by(cases x) simp_all  | 
|
1688  | 
||
1689  | 
lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1690  | 
by (rule measure_eqI)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1691  | 
(simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure ennreal_divide_numeral[symmetric]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1692  | 
nn_integral_count_space_finite sets_uniform_count_measure divide_ennreal_def mult_ac  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1693  | 
ennreal_of_nat_eq_real_of_nat)  | 
| 59525 | 1694  | 
|
| 59093 | 1695  | 
subsubsection \<open> Geometric Distribution \<close>  | 
1696  | 
||
| 60602 | 1697  | 
context  | 
1698  | 
fixes p :: real assumes p[arith]: "0 < p" "p \<le> 1"  | 
|
1699  | 
begin  | 
|
1700  | 
||
1701  | 
lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p"  | 
|
| 59000 | 1702  | 
proof  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1703  | 
have "(\<Sum>i. ennreal (p * (1 - p) ^ i)) = ennreal (p * (1 / (1 - (1 - p))))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1704  | 
by (intro suminf_ennreal_eq sums_mult geometric_sums) auto  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1705  | 
then show "(\<integral>\<^sup>+ x. ennreal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"  | 
| 59000 | 1706  | 
by (simp add: nn_integral_count_space_nat field_simps)  | 
1707  | 
qed simp  | 
|
1708  | 
||
| 60602 | 1709  | 
lemma pmf_geometric[simp]: "pmf geometric_pmf n = (1 - p)^n * p"  | 
| 59000 | 1710  | 
by transfer rule  | 
1711  | 
||
| 60602 | 1712  | 
end  | 
1713  | 
||
1714  | 
lemma set_pmf_geometric: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (geometric_pmf p) = UNIV"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1715  | 
by (auto simp: set_pmf_iff)  | 
| 59000 | 1716  | 
|
| 59093 | 1717  | 
subsubsection \<open> Uniform Multiset Distribution \<close>  | 
1718  | 
||
| 59000 | 1719  | 
context  | 
1720  | 
  fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
 | 
|
1721  | 
begin  | 
|
1722  | 
||
1723  | 
lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M"  | 
|
1724  | 
proof  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1725  | 
show "(\<integral>\<^sup>+ x. ennreal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"  | 
| 59000 | 1726  | 
using M_not_empty  | 
1727  | 
by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size  | 
|
| 64267 | 1728  | 
sum_divide_distrib[symmetric])  | 
1729  | 
(auto simp: size_multiset_overloaded_eq intro!: sum.cong)  | 
|
| 59000 | 1730  | 
qed simp  | 
1731  | 
||
1732  | 
lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M"  | 
|
1733  | 
by transfer rule  | 
|
1734  | 
||
| 60495 | 1735  | 
lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_mset M"  | 
| 59000 | 1736  | 
by (auto simp: set_pmf_iff)  | 
1737  | 
||
1738  | 
end  | 
|
1739  | 
||
| 59093 | 1740  | 
subsubsection \<open> Uniform Distribution \<close>  | 
1741  | 
||
| 59000 | 1742  | 
context  | 
1743  | 
  fixes S :: "'a set" assumes S_not_empty: "S \<noteq> {}" and S_finite: "finite S"
 | 
|
1744  | 
begin  | 
|
1745  | 
||
1746  | 
lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S"  | 
|
1747  | 
proof  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1748  | 
show "(\<integral>\<^sup>+ x. ennreal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1749  | 
using S_not_empty S_finite  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1750  | 
by (subst nn_integral_count_space'[of S])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1751  | 
(auto simp: ennreal_of_nat_eq_real_of_nat ennreal_mult[symmetric])  | 
| 59000 | 1752  | 
qed simp  | 
1753  | 
||
1754  | 
lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S"  | 
|
1755  | 
by transfer rule  | 
|
1756  | 
||
1757  | 
lemma 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  | 
||
| 61634 | 1760  | 
lemma emeasure_pmf_of_set_space[simp]: "emeasure pmf_of_set S = 1"  | 
| 
59002
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1761  | 
by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff)  | 
| 
 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
1762  | 
|
| 64267 | 1763  | 
lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = sum f S / card S"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1764  | 
by (subst nn_integral_measure_pmf_finite)  | 
| 64267 | 1765  | 
(simp_all add: sum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1766  | 
divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)  | 
| 60068 | 1767  | 
|
| 64267 | 1768  | 
lemma 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)  | 
|
| 60068 | 1770  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1771  | 
lemma emeasure_pmf_of_set: "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1772  | 
by (subst nn_integral_indicator[symmetric], simp)  | 
| 64267 | 1773  | 
(simp add: S_finite S_not_empty card_gt_0_iff indicator_def sum.If_cases divide_ennreal  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1774  | 
ennreal_of_nat_eq_real_of_nat nn_integral_pmf_of_set)  | 
| 60068 | 1775  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1776  | 
lemma measure_pmf_of_set: "measure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"  | 
| 63092 | 1777  | 
using emeasure_pmf_of_set[of A]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1778  | 
by (simp add: measure_nonneg measure_pmf.emeasure_eq_measure)  | 
| 61634 | 1779  | 
|
| 59000 | 1780  | 
end  | 
| 67486 | 1781  | 
|
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1782  | 
lemma pmf_expectation_bind_pmf_of_set:  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1783  | 
fixes A :: "'a set" and f :: "'a \<Rightarrow> 'b pmf"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1784  | 
    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1785  | 
  assumes "A \<noteq> {}" "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))"
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1786  | 
shows "measure_pmf.expectation (pmf_of_set A \<bind> f) h =  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1787  | 
(\<Sum>a\<in>A. measure_pmf.expectation (f a) h /\<^sub>R real (card A))"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1788  | 
using assms by (subst pmf_expectation_bind[of A]) (auto simp: divide_simps)  | 
| 59000 | 1789  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1790  | 
lemma map_pmf_of_set:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1791  | 
  assumes "finite A" "A \<noteq> {}"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1792  | 
shows "map_pmf f (pmf_of_set A) = pmf_of_multiset (image_mset f (mset_set A))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1793  | 
(is "?lhs = ?rhs")  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1794  | 
proof (intro pmf_eqI)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1795  | 
fix x  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1796  | 
from assms have "ennreal (pmf ?lhs x) = ennreal (pmf ?rhs x)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1797  | 
by (subst ennreal_pmf_map)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1798  | 
(simp_all add: emeasure_pmf_of_set mset_set_empty_iff count_image_mset Int_commute)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1799  | 
thus "pmf ?lhs x = pmf ?rhs x" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1800  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1801  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1802  | 
lemma pmf_bind_pmf_of_set:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1803  | 
  assumes "A \<noteq> {}" "finite A"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1804  | 
shows "pmf (bind_pmf (pmf_of_set A) f) x =  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1805  | 
(\<Sum>xa\<in>A. pmf (f xa) x) / real_of_nat (card A)" (is "?lhs = ?rhs")  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1806  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1807  | 
from assms have "card A > 0" by auto  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1808  | 
with assms have "ennreal ?lhs = ennreal ?rhs"  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1809  | 
by (subst ennreal_pmf_bind)  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1810  | 
(simp_all add: nn_integral_pmf_of_set max_def pmf_nonneg divide_ennreal [symmetric]  | 
| 64267 | 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)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1813  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1814  | 
|
| 60068 | 1815  | 
lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
 | 
1816  | 
by(rule pmf_eqI)(simp add: indicator_def)  | 
|
1817  | 
||
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1818  | 
lemma map_pmf_of_set_inj:  | 
| 60068 | 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")  | 
|
1822  | 
proof(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  | 
|
1835  | 
qed  | 
|
1836  | 
||
| 
65395
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1837  | 
lemma map_pmf_of_set_bij_betw:  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1838  | 
  assumes "bij_betw f A B" "A \<noteq> {}" "finite A"
 | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1839  | 
shows "map_pmf f (pmf_of_set A) = pmf_of_set B"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1840  | 
proof -  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1841  | 
have "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)"  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1842  | 
by (intro map_pmf_of_set_inj assms bij_betw_imp_inj_on[OF assms(1)])  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1843  | 
also from assms have "f ` A = B" by (simp add: bij_betw_def)  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1844  | 
finally show ?thesis .  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1845  | 
qed  | 
| 
 
7504569a73c7
moved material from AFP to distribution
 
eberlm <eberlm@in.tum.de> 
parents: 
64634 
diff
changeset
 | 
1846  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1847  | 
text \<open>  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1848  | 
Choosing an element uniformly at random from the union of a disjoint family  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1849  | 
of finite non-empty sets with the same size is the same as first choosing a set  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1850  | 
from the family uniformly at random and then choosing an element from the chosen set  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1851  | 
uniformly at random.  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1852  | 
\<close>  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1853  | 
lemma pmf_of_set_UN:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1854  | 
  assumes "finite (UNION A f)" "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1855  | 
"\<And>x. x \<in> A \<Longrightarrow> card (f x) = n" "disjoint_family_on f A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1856  | 
  shows   "pmf_of_set (UNION A f) = do {x \<leftarrow> pmf_of_set A; pmf_of_set (f x)}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1857  | 
(is "?lhs = ?rhs")  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1858  | 
proof (intro pmf_eqI)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1859  | 
fix x  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1860  | 
from assms have [simp]: "finite A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1861  | 
using infinite_disjoint_family_imp_infinite_UNION[of A f] by blast  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1862  | 
from assms have "ereal (pmf (pmf_of_set (UNION A f)) x) =  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1863  | 
ereal (indicator (\<Union>x\<in>A. f x) x / real (card (\<Union>x\<in>A. f x)))"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1864  | 
by (subst pmf_of_set) auto  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1865  | 
also from assms have "card (\<Union>x\<in>A. f x) = card A * n"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1866  | 
by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1867  | 
also from assms  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1868  | 
have "indicator (\<Union>x\<in>A. f x) x / real \<dots> =  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1869  | 
indicator (\<Union>x\<in>A. f x) x / (n * real (card A))"  | 
| 64267 | 1870  | 
by (simp add: sum_divide_distrib [symmetric] mult_ac)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1871  | 
also from assms have "indicator (\<Union>x\<in>A. f x) x = (\<Sum>y\<in>A. indicator (f y) x)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1872  | 
by (intro indicator_UN_disjoint) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1873  | 
also from assms have "ereal ((\<Sum>y\<in>A. indicator (f y) x) / (real n * real (card A))) =  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1874  | 
ereal (pmf ?rhs x)"  | 
| 64267 | 1875  | 
by (subst pmf_bind_pmf_of_set) (simp_all add: sum_divide_distrib)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1876  | 
finally show "pmf ?lhs x = pmf ?rhs x" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1877  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1878  | 
|
| 60068 | 1879  | 
lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1880  | 
by (rule pmf_eqI) simp_all  | 
| 61634 | 1881  | 
|
| 59093 | 1882  | 
subsubsection \<open> Poisson Distribution \<close>  | 
1883  | 
||
1884  | 
context  | 
|
1885  | 
fixes rate :: real assumes rate_pos: "0 < rate"  | 
|
1886  | 
begin  | 
|
1887  | 
||
1888  | 
lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)"  | 
|
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59667 
diff
changeset
 | 
1889  | 
proof (* by Manuel Eberl *)  | 
| 59093 | 1890  | 
have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp  | 
| 59557 | 1891  | 
by (simp add: field_simps divide_inverse [symmetric])  | 
| 59093 | 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)"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1894  | 
by (simp add: field_simps nn_integral_cmult[symmetric] ennreal_mult'[symmetric])  | 
| 59093 | 1895  | 
also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1896  | 
by (simp_all add: nn_integral_count_space_nat suminf_ennreal summable ennreal_suminf_neq_top)  | 
| 59093 | 1897  | 
also have "... = exp rate" unfolding exp_def  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59667 
diff
changeset
 | 
1898  | 
by (simp add: field_simps divide_inverse [symmetric])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1899  | 
also have "ennreal (exp (-rate)) * ennreal (exp rate) = 1"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1900  | 
by (simp add: mult_exp_exp ennreal_mult[symmetric])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1901  | 
finally show "(\<integral>\<^sup>+ x. ennreal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .  | 
| 59093 | 1902  | 
qed (simp add: rate_pos[THEN less_imp_le])  | 
1903  | 
||
1904  | 
lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)"  | 
|
1905  | 
by transfer rule  | 
|
1906  | 
||
1907  | 
lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV"  | 
|
1908  | 
using rate_pos by (auto simp: set_pmf_iff)  | 
|
1909  | 
||
| 59000 | 1910  | 
end  | 
1911  | 
||
| 59093 | 1912  | 
subsubsection \<open> Binomial Distribution \<close>  | 
1913  | 
||
1914  | 
context  | 
|
1915  | 
fixes n :: nat and p :: real assumes p_nonneg: "0 \<le> p" and p_le_1: "p \<le> 1"  | 
|
1916  | 
begin  | 
|
1917  | 
||
1918  | 
lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)"  | 
|
1919  | 
proof  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1920  | 
have "(\<integral>\<^sup>+k. ennreal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1921  | 
ennreal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"  | 
| 59093 | 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"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1924  | 
by (subst binomial_ring) (simp add: atLeast0AtMost)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1925  | 
finally show "(\<integral>\<^sup>+ x. ennreal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"  | 
| 59093 | 1926  | 
by simp  | 
1927  | 
qed (insert p_nonneg p_le_1, simp)  | 
|
1928  | 
||
1929  | 
lemma pmf_binomial[simp]: "pmf binomial_pmf k = (n choose k) * p^k * (1 - p)^(n - k)"  | 
|
1930  | 
by transfer rule  | 
|
1931  | 
||
1932  | 
lemma 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  | 
||
1935  | 
end  | 
|
1936  | 
||
1937  | 
end  | 
|
1938  | 
||
1939  | 
lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
 | 
|
1940  | 
by (simp add: set_pmf_binomial_eq)  | 
|
1941  | 
||
1942  | 
lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
 | 
|
1943  | 
by (simp add: set_pmf_binomial_eq)  | 
|
1944  | 
||
1945  | 
lemma 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  | 
||
| 63343 | 1948  | 
context includes lifting_syntax  | 
1949  | 
begin  | 
|
| 61634 | 1950  | 
|
1951  | 
lemma bind_pmf_parametric [transfer_rule]:  | 
|
1952  | 
"(rel_pmf A ===> (A ===> rel_pmf B) ===> rel_pmf B) bind_pmf bind_pmf"  | 
|
1953  | 
by(blast intro: rel_pmf_bindI dest: rel_funD)  | 
|
1954  | 
||
1955  | 
lemma return_pmf_parametric [transfer_rule]: "(A ===> rel_pmf A) return_pmf return_pmf"  | 
|
1956  | 
by(rule rel_funI) simp  | 
|
1957  | 
||
| 59000 | 1958  | 
end  | 
| 61634 | 1959  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1960  | 
|
| 63194 | 1961  | 
primrec 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  | 
||
1965  | 
lemma replicate_pmf_1: "replicate_pmf 1 p = map_pmf (\<lambda>x. [x]) p"  | 
|
1966  | 
by (simp add: map_pmf_def bind_return_pmf)  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1967  | 
|
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1968  | 
lemma set_replicate_pmf:  | 
| 63194 | 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  | 
||
1972  | 
lemma replicate_pmf_distrib:  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1973  | 
"replicate_pmf (m + n) p =  | 
| 63194 | 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  | 
||
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1977  | 
lemma power_diff':  | 
| 63194 | 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)"  | 
|
1980  | 
proof (cases "x = 0")  | 
|
1981  | 
case True  | 
|
1982  | 
with assms show ?thesis by (cases "a - b") simp_all  | 
|
1983  | 
qed (insert assms, simp_all add: power_diff)  | 
|
1984  | 
||
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1985  | 
|
| 63194 | 1986  | 
lemma binomial_pmf_Suc:  | 
1987  | 
  assumes "p \<in> {0..1}"
 | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1988  | 
shows "binomial_pmf (Suc n) p =  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1989  | 
             do {b \<leftarrow> bernoulli_pmf p;
 | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
1990  | 
k \<leftarrow> binomial_pmf n p;  | 
| 63194 | 1991  | 
return_pmf ((if b then 1 else 0) + k)}" (is "_ = ?rhs")  | 
1992  | 
proof (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 divide_simps algebra_simps  | 
|
1999  | 
not_less less_eq_Suc_le [symmetric] power_diff')  | 
|
2000  | 
qed  | 
|
2001  | 
||
2002  | 
lemma 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  | 
||
2005  | 
lemma 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))"  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2008  | 
by (induction n)  | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2009  | 
(insert assms, auto simp: binomial_pmf_Suc map_pmf_def bind_return_pmf bind_assoc_pmf  | 
| 63194 | 2010  | 
bind_return_pmf' binomial_pmf_0 intro!: bind_pmf_cong)  | 
2011  | 
||
2012  | 
||
| 67486 | 2013  | 
subsection \<open>PMFs from association lists\<close>  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2014  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2015  | 
definition pmf_of_list ::" ('a \<times> real) list \<Rightarrow> 'a pmf" where
 | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2016  | 
"pmf_of_list xs = embed_pmf (\<lambda>x. sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2017  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2018  | 
definition pmf_of_list_wf where  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2019  | 
"pmf_of_list_wf xs \<longleftrightarrow> (\<forall>x\<in>set (map snd xs) . x \<ge> 0) \<and> sum_list (map snd xs) = 1"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2020  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2021  | 
lemma pmf_of_list_wfI:  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
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"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2023  | 
unfolding pmf_of_list_wf_def by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2024  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2025  | 
context  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2026  | 
begin  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2027  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2028  | 
private lemma pmf_of_list_aux:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2029  | 
assumes "\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0"  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2030  | 
assumes "sum_list (map snd xs) = 1"  | 
| 68386 | 2031  | 
shows "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])) \<partial>count_space UNIV) = 1"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2032  | 
proof -  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2033  | 
have "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))) \<partial>count_space UNIV) =  | 
| 
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2034  | 
            (\<integral>\<^sup>+ x. ennreal (sum_list (map (\<lambda>(x',p). indicator {x'} x * p) xs)) \<partial>count_space UNIV)"
 | 
| 67486 | 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  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2039  | 
  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. (\<integral>\<^sup>+ x. ennreal (indicator {x'} x * p) \<partial>count_space UNIV))"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2040  | 
using assms(1)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2041  | 
proof (induction xs)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2042  | 
case (Cons x xs)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2043  | 
from Cons.prems have "snd x \<ge> 0" by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2044  | 
moreover have "b \<ge> 0" if "(a,b) \<in> set xs" for a b  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2045  | 
using Cons.prems[of b] that by force  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2046  | 
    ultimately have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>x # xs. indicator {x'} y * p) \<partial>count_space UNIV) =
 | 
| 
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2047  | 
            (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) +
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2048  | 
            ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2049  | 
by (intro nn_integral_cong, subst ennreal_plus [symmetric])  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2050  | 
(auto simp: case_prod_unfold indicator_def intro!: sum_list_nonneg)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2051  | 
    also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) \<partial>count_space UNIV) +
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2052  | 
                      (\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2053  | 
by (intro nn_integral_add)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2054  | 
(force intro!: sum_list_nonneg AE_I2 intro: Cons simp: indicator_def)+  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2055  | 
    also have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV) =
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2056  | 
               (\<Sum>(x', p)\<leftarrow>xs. (\<integral>\<^sup>+ y. ennreal (indicator {x'} y * p) \<partial>count_space UNIV))"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2057  | 
using Cons(1) by (intro Cons) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2058  | 
finally show ?case by (simp add: case_prod_unfold)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2059  | 
qed simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2060  | 
  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. ennreal p * (\<integral>\<^sup>+ x. indicator {x'} x \<partial>count_space UNIV))"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2061  | 
using assms(1)  | 
| 
67489
 
f1ba59ddd9a6
drop redundant cong rules
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
67486 
diff
changeset
 | 
2062  | 
by (simp cong: map_cong only: case_prod_unfold, subst nn_integral_cmult [symmetric])  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2063  | 
(auto intro!: assms(1) simp: max_def times_ereal.simps [symmetric] mult_ac ereal_indicator  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2064  | 
simp del: times_ereal.simps)+  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2065  | 
also from assms have "\<dots> = sum_list (map snd xs)" by (simp add: case_prod_unfold sum_list_ennreal)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2066  | 
also have "\<dots> = 1" using assms(2) by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2067  | 
finally show ?thesis .  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2068  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2069  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2070  | 
lemma pmf_pmf_of_list:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2071  | 
assumes "pmf_of_list_wf xs"  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2072  | 
shows "pmf (pmf_of_list xs) x = sum_list (map snd (filter (\<lambda>z. fst z = x) xs))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2073  | 
using assms pmf_of_list_aux[of xs] unfolding pmf_of_list_def pmf_of_list_wf_def  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2074  | 
by (subst pmf_embed_pmf) (auto intro!: sum_list_nonneg)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2075  | 
|
| 61634 | 2076  | 
end  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2077  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2078  | 
lemma set_pmf_of_list:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2079  | 
assumes "pmf_of_list_wf xs"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2080  | 
shows "set_pmf (pmf_of_list xs) \<subseteq> set (map fst xs)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2081  | 
proof clarify  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2082  | 
fix x assume A: "x \<in> set_pmf (pmf_of_list xs)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2083  | 
show "x \<in> set (map fst xs)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2084  | 
proof (rule ccontr)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2085  | 
assume "x \<notin> set (map fst xs)"  | 
| 68386 | 2086  | 
hence "[z\<leftarrow>xs . fst z = x] = []" by (auto simp: filter_empty_conv)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2087  | 
with A assms show False by (simp add: pmf_pmf_of_list set_pmf_eq)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2088  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2089  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2090  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2091  | 
lemma finite_set_pmf_of_list:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2092  | 
assumes "pmf_of_list_wf xs"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2093  | 
shows "finite (set_pmf (pmf_of_list xs))"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2094  | 
using assms by (rule finite_subset[OF set_pmf_of_list]) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2095  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2096  | 
lemma emeasure_Int_set_pmf:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2097  | 
"emeasure (measure_pmf p) (A \<inter> set_pmf p) = emeasure (measure_pmf p) A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2098  | 
by (rule emeasure_eq_AE) (auto simp: AE_measure_pmf_iff)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2099  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2100  | 
lemma measure_Int_set_pmf:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2101  | 
"measure (measure_pmf p) (A \<inter> set_pmf p) = measure (measure_pmf p) A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2102  | 
using emeasure_Int_set_pmf[of p A] by (simp add: Sigma_Algebra.measure_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2103  | 
|
| 
66568
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2104  | 
lemma measure_prob_cong_0:  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2105  | 
assumes "\<And>x. x \<in> A - B \<Longrightarrow> pmf p x = 0"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2106  | 
assumes "\<And>x. x \<in> B - A \<Longrightarrow> pmf p x = 0"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2107  | 
shows "measure (measure_pmf p) A = measure (measure_pmf p) B"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2108  | 
proof -  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2109  | 
have "measure_pmf.prob p A = measure_pmf.prob p (A \<inter> set_pmf p)"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2110  | 
by (simp add: measure_Int_set_pmf)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2111  | 
also have "A \<inter> set_pmf p = B \<inter> set_pmf p"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2112  | 
using assms by (auto simp: set_pmf_eq)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2113  | 
also have "measure_pmf.prob p \<dots> = measure_pmf.prob p B"  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2114  | 
by (simp add: measure_Int_set_pmf)  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2115  | 
finally show ?thesis .  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2116  | 
qed  | 
| 
 
52b5cf533fd6
Connecting PMFs to infinite sums
 
eberlm <eberlm@in.tum.de> 
parents: 
66453 
diff
changeset
 | 
2117  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2118  | 
lemma emeasure_pmf_of_list:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2119  | 
assumes "pmf_of_list_wf xs"  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2120  | 
shows "emeasure (pmf_of_list xs) A = ennreal (sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs)))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2121  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2122  | 
have "emeasure (pmf_of_list xs) A = nn_integral (measure_pmf (pmf_of_list xs)) (indicator A)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2123  | 
by simp  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2124  | 
also from assms  | 
| 68386 | 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])))"  | 
| 63973 | 2126  | 
by (subst nn_integral_measure_pmf_finite) (simp_all add: finite_set_pmf_of_list pmf_pmf_of_list Int_def)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2127  | 
also from assms  | 
| 68386 | 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]))"  | 
| 64267 | 2129  | 
by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2130  | 
also have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A.  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2131  | 
indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S")  | 
| 64267 | 2132  | 
using assms by (intro ennreal_cong sum.cong) (auto simp: pmf_pmf_of_list)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2133  | 
also have "?S = (\<Sum>x\<in>set_pmf (pmf_of_list xs). indicator A x * pmf (pmf_of_list xs) x)"  | 
| 64267 | 2134  | 
using assms by (intro sum.mono_neutral_left set_pmf_of_list finite_set_pmf_of_list) auto  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2135  | 
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x * pmf (pmf_of_list xs) x)"  | 
| 64267 | 2136  | 
using assms by (intro sum.mono_neutral_left set_pmf_of_list) (auto simp: set_pmf_eq)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2137  | 
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x *  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2138  | 
sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2139  | 
using assms by (simp add: pmf_pmf_of_list)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
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)))"  | 
| 64267 | 2141  | 
by (intro sum.cong) (auto simp: indicator_def)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2142  | 
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). (\<Sum>xa = 0..<length xs.  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2143  | 
if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"  | 
| 64267 | 2144  | 
by (intro sum.cong refl, subst sum_list_map_filter', subst sum_list_sum_nth) simp  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2145  | 
also have "\<dots> = (\<Sum>xa = 0..<length xs. (\<Sum>x\<in>set (map fst xs).  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2146  | 
if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"  | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66568 
diff
changeset
 | 
2147  | 
by (rule sum.swap)  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2148  | 
also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2149  | 
(\<Sum>x\<in>set (map fst xs). if x = fst (xs ! xa) then snd (xs ! xa) else 0) else 0)"  | 
| 
66089
 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 
paulson <lp15@cam.ac.uk> 
parents: 
65395 
diff
changeset
 | 
2150  | 
by (auto intro!: sum.cong sum.neutral simp del: sum.delta)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2151  | 
also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then snd (xs ! xa) else 0)"  | 
| 64267 | 2152  | 
by (intro sum.cong refl) (simp_all add: sum.delta)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2153  | 
also have "\<dots> = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"  | 
| 64267 | 2154  | 
by (subst sum_list_map_filter', subst sum_list_sum_nth) simp_all  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2155  | 
finally show ?thesis .  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2156  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2157  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2158  | 
lemma measure_pmf_of_list:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2159  | 
assumes "pmf_of_list_wf xs"  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2160  | 
shows "measure (pmf_of_list xs) A = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2161  | 
using assms unfolding pmf_of_list_wf_def Sigma_Algebra.measure_def  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2162  | 
by (subst emeasure_pmf_of_list [OF assms], subst enn2real_ennreal) (auto intro!: sum_list_nonneg)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2163  | 
|
| 63194 | 2164  | 
(* TODO Move? *)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2165  | 
lemma sum_list_nonneg_eq_zero_iff:  | 
| 63194 | 2166  | 
fixes xs :: "'a :: linordered_ab_group_add list"  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2167  | 
  shows "(\<And>x. x \<in> set xs \<Longrightarrow> x \<ge> 0) \<Longrightarrow> sum_list xs = 0 \<longleftrightarrow> set xs \<subseteq> {0}"
 | 
| 63194 | 2168  | 
proof (induction xs)  | 
2169  | 
case (Cons x xs)  | 
|
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2170  | 
from Cons.prems have "sum_list (x#xs) = 0 \<longleftrightarrow> x = 0 \<and> sum_list xs = 0"  | 
| 
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2171  | 
unfolding sum_list_simps by (subst add_nonneg_eq_0_iff) (auto intro: sum_list_nonneg)  | 
| 63194 | 2172  | 
with Cons.IH Cons.prems show ?case by simp  | 
2173  | 
qed simp_all  | 
|
2174  | 
||
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2175  | 
lemma sum_list_filter_nonzero:  | 
| 
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2176  | 
"sum_list (filter (\<lambda>x. x \<noteq> 0) xs) = sum_list xs"  | 
| 63194 | 2177  | 
by (induction xs) simp_all  | 
2178  | 
(* END MOVE *)  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2179  | 
|
| 63194 | 2180  | 
lemma 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"  | 
|
2183  | 
proof  | 
|
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  | 
|
| 68386 | 2187  | 
from B have "sum_list (map snd [z\<leftarrow>xs. fst z = x]) = 0"  | 
| 63194 | 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
 | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2190  | 
ultimately have "y = 0" using assms(1)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2191  | 
by (subst (asm) sum_list_nonneg_eq_zero_iff) (auto simp: pmf_of_list_wf_def)  | 
| 63194 | 2192  | 
with assms(2) y have False by force  | 
2193  | 
}  | 
|
2194  | 
thus "fst ` set xs \<subseteq> set_pmf (pmf_of_list xs)" by blast  | 
|
2195  | 
qed (insert set_pmf_of_list[OF assms(1)], simp_all)  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63882 
diff
changeset
 | 
2196  | 
|
| 63194 | 2197  | 
lemma 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"  | 
|
2201  | 
proof -  | 
|
| 68386 | 2202  | 
have "map snd [z\<leftarrow>xs . snd z \<noteq> 0] = filter (\<lambda>x. x \<noteq> 0) (map snd xs)"  | 
| 63194 | 2203  | 
by (induction xs) simp_all  | 
2204  | 
with assms(1) show wf: "pmf_of_list_wf xs'"  | 
|
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63680 
diff
changeset
 | 
2205  | 
by (auto simp: pmf_of_list_wf_def xs'_def sum_list_filter_nonzero)  | 
| 68386 | 2206  | 
have "sum_list (map snd [z\<leftarrow>xs' . fst z = i]) = sum_list (map snd [z\<leftarrow>xs . fst z = i])" for i  | 
| 63194 | 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)  | 
|
2210  | 
qed  | 
|
2211  | 
||
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2212  | 
end  |