| author | paulson <lp15@cam.ac.uk> | 
| Wed, 09 Dec 2015 17:35:22 +0000 | |
| changeset 61810 | 3c5040d5694a | 
| parent 61808 | fc1556774cfe | 
| child 62026 | ea3b1b0413b4 | 
| 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: 
59665diff
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 | |
| 11 | "~~/src/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 | |
| 25 | by (auto simp: emeasure_le_0_iff) } | |
| 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 | ||
| 59494 | 32 | lemma ereal_divide': "b \<noteq> 0 \<Longrightarrow> ereal (a / b) = ereal a / ereal b" | 
| 33 | using ereal_divide[of a b] by simp | |
| 34 | ||
| 59052 | 35 | lemma (in finite_measure) countable_support: | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 36 |   "countable {x. measure M {x} \<noteq> 0}"
 | 
| 59000 | 37 | proof cases | 
| 38 | assume "measure M (space M) = 0" | |
| 39 |   with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
 | |
| 40 | by auto | |
| 41 | then show ?thesis | |
| 42 | by simp | |
| 43 | next | |
| 44 |   let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
 | |
| 45 | assume "?M \<noteq> 0" | |
| 46 |   then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
 | |
| 47 | using reals_Archimedean[of "?m x / ?M" for x] | |
| 48 | by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff) | |
| 49 |   have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
 | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 50 | proof (rule ccontr) | 
| 59000 | 51 |     fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
 | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 52 | then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X" | 
| 58606 | 53 | by (metis infinite_arbitrarily_large) | 
| 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: 
59665diff
changeset | 54 | from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x" | 
| 59000 | 55 | by auto | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 56 |     { fix x assume "x \<in> X"
 | 
| 61808 | 57 | from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff) | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 58 |       then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
 | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 59 | note singleton_sets = this | 
| 59000 | 60 | have "?M < (\<Sum>x\<in>X. ?M / Suc n)" | 
| 61808 | 61 | using \<open>?M \<noteq> 0\<close> | 
| 62 | by (simp add: \<open>card X = Suc (Suc n)\<close> of_nat_Suc field_simps less_le measure_nonneg) | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 63 | also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 64 | by (rule setsum_mono) fact | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 65 |     also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
 | 
| 61808 | 66 | using singleton_sets \<open>finite X\<close> | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 67 | by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) | 
| 59000 | 68 |     finally have "?M < measure M (\<Union>x\<in>X. {x})" .
 | 
| 69 |     moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
 | |
| 70 | using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto | |
| 71 | ultimately show False by simp | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 72 | qed | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 73 | show ?thesis | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 74 | unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 75 | qed | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 76 | |
| 59000 | 77 | lemma (in finite_measure) AE_support_countable: | 
| 78 | assumes [simp]: "sets M = UNIV" | |
| 79 |   shows "(AE x in M. measure M {x} \<noteq> 0) \<longleftrightarrow> (\<exists>S. countable S \<and> (AE x in M. x \<in> S))"
 | |
| 80 | proof | |
| 81 | assume "\<exists>S. countable S \<and> (AE x in M. x \<in> S)" | |
| 82 | then obtain S where S[intro]: "countable S" and ae: "AE x in M. x \<in> S" | |
| 83 | 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: 
59665diff
changeset | 84 |   then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
 | 
| 59000 | 85 |     (\<integral>\<^sup>+ x. emeasure M {x} * indicator {x\<in>S. emeasure M {x} \<noteq> 0} x \<partial>count_space UNIV)"
 | 
| 86 | by (subst emeasure_UN_countable) | |
| 87 | (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space) | |
| 88 |   also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} * indicator S x \<partial>count_space UNIV)"
 | |
| 89 | by (auto intro!: nn_integral_cong split: split_indicator) | |
| 90 |   also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
 | |
| 91 | by (subst emeasure_UN_countable) | |
| 92 | (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space) | |
| 93 | also have "\<dots> = emeasure M (space M)" | |
| 94 | using ae by (intro emeasure_eq_AE) auto | |
| 95 |   finally have "emeasure M {x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0} = emeasure M (space M)"
 | |
| 96 | by (simp add: emeasure_single_in_space cong: rev_conj_cong) | |
| 97 |   with finite_measure_compl[of "{x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0}"]
 | |
| 98 |   have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
 | |
| 99 | by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure set_diff_eq cong: conj_cong) | |
| 100 |   then show "AE x in M. measure M {x} \<noteq> 0"
 | |
| 101 | by (auto simp: emeasure_eq_measure) | |
| 102 | qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
 | |
| 103 | ||
| 59664 | 104 | subsection \<open> PMF as measure \<close> | 
| 59000 | 105 | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 106 | 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 | 107 | morphisms measure_pmf Abs_pmf | 
| 58606 | 108 |   by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
 | 
| 109 | (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 | 110 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 111 | declare [[coercion measure_pmf]] | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 112 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 113 | 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 | 114 | using pmf.measure_pmf[of p] by auto | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 115 | |
| 61605 | 116 | 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 | 117 | by (rule prob_space_measure_pmf) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 118 | |
| 61605 | 119 | interpretation measure_pmf: subprob_space "measure_pmf M" for M | 
| 59000 | 120 | by (rule prob_space_imp_subprob_space) unfold_locales | 
| 121 | ||
| 59048 | 122 | lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)" | 
| 123 | by unfold_locales | |
| 124 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 125 | locale pmf_as_measure | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 126 | begin | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 127 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 128 | setup_lifting type_definition_pmf | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 129 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 130 | end | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 131 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 132 | context | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 133 | begin | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 134 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 135 | interpretation pmf_as_measure . | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 136 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 137 | 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: 
59665diff
changeset | 138 | by transfer blast | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 139 | |
| 59048 | 140 | lemma sets_measure_pmf_count_space[measurable_cong]: | 
| 141 | "sets (measure_pmf M) = sets (count_space UNIV)" | |
| 59000 | 142 | by simp | 
| 143 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 144 | 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 | 145 | 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 | 146 | |
| 61634 | 147 | lemma measure_pmf_UNIV [simp]: "measure (measure_pmf p) UNIV = 1" | 
| 148 | using measure_pmf.prob_space[of p] by simp | |
| 149 | ||
| 59048 | 150 | lemma measure_pmf_in_subprob_algebra[measurable (raw)]: "measure_pmf x \<in> space (subprob_algebra (count_space UNIV))" | 
| 151 | by (simp add: space_subprob_algebra subprob_space_measure_pmf) | |
| 152 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 153 | 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 | 154 | by (auto simp: measurable_def) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 155 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 156 | 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 | 157 | by (intro measurable_cong_sets) simp_all | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 158 | |
| 59664 | 159 | lemma measurable_pair_restrict_pmf2: | 
| 160 | assumes "countable A" | |
| 161 | assumes [measurable]: "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L" | |
| 162 | shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L" (is "f \<in> measurable ?M _") | |
| 163 | proof - | |
| 164 | have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)" | |
| 165 | by (simp add: restrict_count_space) | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 166 | |
| 59664 | 167 | show ?thesis | 
| 168 | 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: 
61169diff
changeset | 169 | unfolded prod.collapse] assms) | 
| 59664 | 170 | measurable | 
| 171 | qed | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 172 | |
| 59664 | 173 | lemma measurable_pair_restrict_pmf1: | 
| 174 | assumes "countable A" | |
| 175 | assumes [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L" | |
| 176 | shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L" | |
| 177 | proof - | |
| 178 | have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)" | |
| 179 | by (simp add: restrict_count_space) | |
| 59000 | 180 | |
| 59664 | 181 | show ?thesis | 
| 182 | 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: 
61169diff
changeset | 183 | unfolded prod.collapse] assms) | 
| 59664 | 184 | measurable | 
| 185 | qed | |
| 186 | ||
| 187 | lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
 | |
| 188 | ||
| 189 | lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
 | |
| 190 | declare [[coercion set_pmf]] | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 191 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 192 | 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 | 193 | by transfer simp | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 194 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 195 | lemma emeasure_pmf_single_eq_zero_iff: | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 196 | fixes M :: "'a pmf" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 197 |   shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
 | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 198 | by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 199 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 200 | lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)" | 
| 59664 | 201 | using AE_measure_singleton[of M] AE_measure_pmf[of M] | 
| 202 | by (auto simp: set_pmf.rep_eq) | |
| 203 | ||
| 61634 | 204 | lemma AE_pmfI: "(\<And>y. y \<in> set_pmf M \<Longrightarrow> P y) \<Longrightarrow> almost_everywhere (measure_pmf M) P" | 
| 205 | by(simp add: AE_measure_pmf_iff) | |
| 206 | ||
| 59664 | 207 | lemma countable_set_pmf [simp]: "countable (set_pmf p)" | 
| 208 | by transfer (metis prob_space.finite_measure finite_measure.countable_support) | |
| 209 | ||
| 210 | lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x" | |
| 211 | by transfer (simp add: less_le measure_nonneg) | |
| 212 | ||
| 213 | lemma pmf_nonneg: "0 \<le> pmf p x" | |
| 214 | by transfer (simp add: measure_nonneg) | |
| 215 | ||
| 216 | lemma pmf_le_1: "pmf p x \<le> 1" | |
| 217 | by (simp add: pmf.rep_eq) | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 218 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 219 | 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 | 220 | 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 | 221 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 222 | 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 | 223 | by transfer simp | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 224 | |
| 59664 | 225 | lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
 | 
| 226 | by (auto simp: set_pmf_iff) | |
| 227 | ||
| 228 | lemma emeasure_pmf_single: | |
| 229 | fixes M :: "'a pmf" | |
| 230 |   shows "emeasure M {x} = pmf M x"
 | |
| 231 | by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) | |
| 232 | ||
| 60068 | 233 | lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
 | 
| 234 | using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure) | |
| 235 | ||
| 59000 | 236 | lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)" | 
| 237 | by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single) | |
| 238 | ||
| 59023 | 239 | lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S" | 
| 59425 | 240 | using emeasure_measure_pmf_finite[of S M] by(simp add: measure_pmf.emeasure_eq_measure) | 
| 59023 | 241 | |
| 59000 | 242 | lemma nn_integral_measure_pmf_support: | 
| 243 | fixes f :: "'a \<Rightarrow> ereal" | |
| 244 | 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" | |
| 245 | shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>A. f x * pmf M x)" | |
| 246 | proof - | |
| 247 | have "(\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)" | |
| 248 | using nn by (intro nn_integral_cong_AE) (auto simp: AE_measure_pmf_iff split: split_indicator) | |
| 249 |   also have "\<dots> = (\<Sum>x\<in>A. f x * emeasure M {x})"
 | |
| 250 | using assms by (intro nn_integral_indicator_finite) auto | |
| 251 | finally show ?thesis | |
| 252 | by (simp add: emeasure_measure_pmf_finite) | |
| 253 | qed | |
| 254 | ||
| 255 | lemma nn_integral_measure_pmf_finite: | |
| 256 | fixes f :: "'a \<Rightarrow> ereal" | |
| 257 | assumes f: "finite (set_pmf M)" and nn: "\<And>x. x \<in> set_pmf M \<Longrightarrow> 0 \<le> f x" | |
| 258 | shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>set_pmf M. f x * pmf M x)" | |
| 259 | using assms by (intro nn_integral_measure_pmf_support) auto | |
| 260 | lemma integrable_measure_pmf_finite: | |
| 261 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | |
| 262 | shows "finite (set_pmf M) \<Longrightarrow> integrable M f" | |
| 263 | by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite) | |
| 264 | ||
| 265 | lemma integral_measure_pmf: | |
| 266 | assumes [simp]: "finite A" and "\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A" | |
| 267 | shows "(\<integral>x. f x \<partial>measure_pmf M) = (\<Sum>a\<in>A. f a * pmf M a)" | |
| 268 | proof - | |
| 269 | have "(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x * indicator A x \<partial>measure_pmf M)" | |
| 270 | using assms(2) by (intro integral_cong_AE) (auto split: split_indicator simp: AE_measure_pmf_iff) | |
| 271 | also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)" | |
| 272 | by (subst integral_indicator_finite_real) (auto simp: measure_def emeasure_measure_pmf_finite) | |
| 273 | finally show ?thesis . | |
| 274 | qed | |
| 275 | ||
| 276 | lemma integrable_pmf: "integrable (count_space X) (pmf M)" | |
| 277 | proof - | |
| 278 | have " (\<integral>\<^sup>+ x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+ x. pmf M x \<partial>count_space (M \<inter> X))" | |
| 279 | by (auto simp add: nn_integral_count_space_indicator set_pmf_iff intro!: nn_integral_cong split: split_indicator) | |
| 280 | then have "integrable (count_space X) (pmf M) = integrable (count_space (M \<inter> X)) (pmf M)" | |
| 281 | by (simp add: integrable_iff_bounded pmf_nonneg) | |
| 282 | then show ?thesis | |
| 59023 | 283 | by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def) | 
| 59000 | 284 | qed | 
| 285 | ||
| 286 | lemma integral_pmf: "(\<integral>x. pmf M x \<partial>count_space X) = measure M X" | |
| 287 | proof - | |
| 288 | have "(\<integral>x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+x. pmf M x \<partial>count_space X)" | |
| 289 | by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral) | |
| 290 |   also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space (X \<inter> M))"
 | |
| 291 | by (auto intro!: nn_integral_cong_AE split: split_indicator | |
| 292 | simp: pmf.rep_eq measure_pmf.emeasure_eq_measure nn_integral_count_space_indicator | |
| 293 | AE_count_space set_pmf_iff) | |
| 294 | also have "\<dots> = emeasure M (X \<inter> M)" | |
| 295 | by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf) | |
| 296 | also have "\<dots> = emeasure M X" | |
| 297 | by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff) | |
| 298 | finally show ?thesis | |
| 299 | by (simp add: measure_pmf.emeasure_eq_measure) | |
| 300 | qed | |
| 301 | ||
| 302 | lemma integral_pmf_restrict: | |
| 303 |   "(f::'a \<Rightarrow> 'b::{banach, second_countable_topology}) \<in> borel_measurable (count_space UNIV) \<Longrightarrow>
 | |
| 304 | (\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x \<partial>restrict_space M M)" | |
| 305 | by (auto intro!: integral_cong_AE simp add: integral_restrict_space AE_measure_pmf_iff) | |
| 306 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 307 | 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 | 308 | proof - | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 309 | 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 | 310 | 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 | 311 | then show ?thesis | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 312 | using measure_pmf.emeasure_space_1 by simp | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 313 | qed | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 314 | |
| 59490 | 315 | lemma emeasure_pmf_UNIV [simp]: "emeasure (measure_pmf M) UNIV = 1" | 
| 316 | using measure_pmf.emeasure_space_1[of M] by simp | |
| 317 | ||
| 59023 | 318 | lemma in_null_sets_measure_pmfI: | 
| 319 |   "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
 | |
| 320 | using emeasure_eq_0_AE[where ?P="\<lambda>x. x \<in> A" and M="measure_pmf p"] | |
| 321 | by(auto simp add: null_sets_def AE_measure_pmf_iff) | |
| 322 | ||
| 59664 | 323 | lemma measure_subprob: "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))" | 
| 324 | by (simp add: space_subprob_algebra subprob_space_measure_pmf) | |
| 325 | ||
| 326 | subsection \<open> Monad Interpretation \<close> | |
| 327 | ||
| 328 | lemma measurable_measure_pmf[measurable]: | |
| 329 | "(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))" | |
| 330 | by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales | |
| 331 | ||
| 332 | lemma bind_measure_pmf_cong: | |
| 333 | assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)" | |
| 334 | assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i" | |
| 335 | shows "bind (measure_pmf x) A = bind (measure_pmf x) B" | |
| 336 | proof (rule measure_eqI) | |
| 337 | show "sets (measure_pmf x \<guillemotright>= A) = sets (measure_pmf x \<guillemotright>= B)" | |
| 338 | using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra) | |
| 339 | next | |
| 340 | fix X assume "X \<in> sets (measure_pmf x \<guillemotright>= A)" | |
| 341 | then have X: "X \<in> sets N" | |
| 342 | using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra) | |
| 343 | show "emeasure (measure_pmf x \<guillemotright>= A) X = emeasure (measure_pmf x \<guillemotright>= B) X" | |
| 344 | using assms | |
| 345 | by (subst (1 2) emeasure_bind[where N=N, OF _ _ X]) | |
| 346 | (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff) | |
| 347 | qed | |
| 348 | ||
| 349 | lift_definition bind_pmf :: "'a pmf \<Rightarrow> ('a \<Rightarrow> 'b pmf ) \<Rightarrow> 'b pmf" is bind
 | |
| 350 | proof (clarify, intro conjI) | |
| 351 | fix f :: "'a measure" and g :: "'a \<Rightarrow> 'b measure" | |
| 352 | assume "prob_space f" | |
| 353 | then interpret f: prob_space f . | |
| 354 |   assume "sets f = UNIV" and ae_f: "AE x in f. measure f {x} \<noteq> 0"
 | |
| 355 | then have s_f[simp]: "sets f = sets (count_space UNIV)" | |
| 356 | by simp | |
| 357 |   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)"
 | |
| 358 | then have g: "\<And>x. prob_space (g x)" and s_g[simp]: "\<And>x. sets (g x) = sets (count_space UNIV)" | |
| 359 |     and ae_g: "\<And>x. AE y in g x. measure (g x) {y} \<noteq> 0"
 | |
| 360 | by auto | |
| 361 | ||
| 362 | have [measurable]: "g \<in> measurable f (subprob_algebra (count_space UNIV))" | |
| 363 | 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: 
59665diff
changeset | 364 | |
| 59664 | 365 | show "prob_space (f \<guillemotright>= g)" | 
| 366 | using g by (intro f.prob_space_bind[where S="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: 
59665diff
changeset | 367 | then interpret fg: prob_space "f \<guillemotright>= g" . | 
| 59664 | 368 | show [simp]: "sets (f \<guillemotright>= g) = UNIV" | 
| 369 | using sets_eq_imp_space_eq[OF s_f] | |
| 370 | by (subst sets_bind[where N="count_space UNIV"]) auto | |
| 371 |   show "AE x in f \<guillemotright>= g. measure (f \<guillemotright>= g) {x} \<noteq> 0"
 | |
| 372 | apply (simp add: fg.prob_eq_0 AE_bind[where B="count_space UNIV"]) | |
| 373 | using ae_f | |
| 374 | apply eventually_elim | |
| 375 | using ae_g | |
| 376 | apply eventually_elim | |
| 377 | apply (auto dest: AE_measure_singleton) | |
| 378 | done | |
| 379 | qed | |
| 380 | ||
| 381 | lemma ereal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)" | |
| 382 | unfolding pmf.rep_eq bind_pmf.rep_eq | |
| 383 | by (auto simp: measure_pmf.measure_bind[where N="count_space UNIV"] measure_subprob measure_nonneg | |
| 384 | intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1]) | |
| 385 | ||
| 386 | lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)" | |
| 387 | using ereal_pmf_bind[of N f i] | |
| 388 | by (subst (asm) nn_integral_eq_integral) | |
| 389 | (auto simp: pmf_nonneg pmf_le_1 | |
| 390 | intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1]) | |
| 391 | ||
| 392 | lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c" | |
| 393 | by transfer (simp add: bind_const' prob_space_imp_subprob_space) | |
| 394 | ||
| 59665 | 395 | lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))" | 
| 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: 
59665diff
changeset | 396 | unfolding set_pmf_eq ereal_eq_0(1)[symmetric] ereal_pmf_bind | 
| 59664 | 397 | by (auto simp add: nn_integral_0_iff_AE AE_measure_pmf_iff set_pmf_eq not_le less_le pmf_nonneg) | 
| 398 | ||
| 399 | lemma bind_pmf_cong: | |
| 400 | assumes "p = q" | |
| 401 | shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g" | |
| 61808 | 402 | unfolding \<open>p = q\<close>[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq | 
| 59664 | 403 | by (auto simp: AE_measure_pmf_iff Pi_iff space_subprob_algebra subprob_space_measure_pmf | 
| 404 | sets_bind[where N="count_space UNIV"] emeasure_bind[where N="count_space UNIV"] | |
| 405 | intro!: nn_integral_cong_AE measure_eqI) | |
| 406 | ||
| 407 | lemma bind_pmf_cong_simp: | |
| 408 | "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q =simp=> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g" | |
| 409 | by (simp add: simp_implies_def cong: bind_pmf_cong) | |
| 410 | ||
| 411 | lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<guillemotright>= (\<lambda>x. measure_pmf (f x)))" | |
| 412 | by transfer simp | |
| 413 | ||
| 414 | 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)" | |
| 415 | using measurable_measure_pmf[of N] | |
| 416 | unfolding measure_pmf_bind | |
| 417 | apply (subst (1 3) nn_integral_max_0[symmetric]) | |
| 418 | apply (intro nn_integral_bind[where B="count_space UNIV"]) | |
| 419 | apply auto | |
| 420 | done | |
| 421 | ||
| 422 | lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)" | |
| 423 | using measurable_measure_pmf[of N] | |
| 424 | unfolding measure_pmf_bind | |
| 425 | 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: 
59665diff
changeset | 426 | |
| 59664 | 427 | lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)" | 
| 428 | by (auto intro!: prob_space_return simp: AE_return measure_return) | |
| 429 | ||
| 430 | lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x" | |
| 431 | by transfer | |
| 432 | (auto intro!: prob_space_imp_subprob_space bind_return[where N="count_space UNIV"] | |
| 433 | simp: space_subprob_algebra) | |
| 434 | ||
| 59665 | 435 | lemma set_return_pmf[simp]: "set_pmf (return_pmf x) = {x}"
 | 
| 59664 | 436 | by transfer (auto simp add: measure_return split: split_indicator) | 
| 437 | ||
| 438 | lemma bind_return_pmf': "bind_pmf N return_pmf = N" | |
| 439 | proof (transfer, clarify) | |
| 440 | fix N :: "'a measure" assume "sets N = UNIV" then show "N \<guillemotright>= return (count_space UNIV) = N" | |
| 441 | by (subst return_sets_cong[where N=N]) (simp_all add: bind_return') | |
| 442 | qed | |
| 443 | ||
| 444 | lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)" | |
| 445 | by transfer | |
| 446 | (auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"] | |
| 447 | simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space) | |
| 448 | ||
| 449 | definition "map_pmf f M = bind_pmf M (\<lambda>x. return_pmf (f x))" | |
| 450 | ||
| 451 | lemma map_bind_pmf: "map_pmf f (bind_pmf M g) = bind_pmf M (\<lambda>x. map_pmf f (g x))" | |
| 452 | by (simp add: map_pmf_def bind_assoc_pmf) | |
| 453 | ||
| 454 | lemma bind_map_pmf: "bind_pmf (map_pmf f M) g = bind_pmf M (\<lambda>x. g (f x))" | |
| 455 | by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf) | |
| 456 | ||
| 457 | lemma map_pmf_transfer[transfer_rule]: | |
| 458 | "rel_fun op = (rel_fun cr_pmf cr_pmf) (\<lambda>f M. distr M (count_space UNIV) f) map_pmf" | |
| 459 | proof - | |
| 460 | have "rel_fun op = (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf) | |
| 461 | (\<lambda>f M. M \<guillemotright>= (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: 
59665diff
changeset | 462 | unfolding map_pmf_def[abs_def] comp_def by transfer_prover | 
| 59664 | 463 | then show ?thesis | 
| 464 | by (force simp: rel_fun_def cr_pmf_def bind_return_distr) | |
| 465 | qed | |
| 466 | ||
| 467 | lemma map_pmf_rep_eq: | |
| 468 | "measure_pmf (map_pmf f M) = distr (measure_pmf M) (count_space UNIV) f" | |
| 469 | unfolding map_pmf_def bind_pmf.rep_eq comp_def return_pmf.rep_eq | |
| 470 | using bind_return_distr[of M f "count_space UNIV"] by (simp add: comp_def) | |
| 471 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 472 | 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 | 473 | 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 | 474 | |
| 59053 | 475 | lemma map_pmf_ident[simp]: "map_pmf (\<lambda>x. x) = (\<lambda>x. x)" | 
| 476 | using map_pmf_id unfolding id_def . | |
| 477 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 478 | 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: 
59665diff
changeset | 479 | 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 | 480 | |
| 59000 | 481 | lemma map_pmf_comp: "map_pmf f (map_pmf g M) = map_pmf (\<lambda>x. f (g x)) M" | 
| 482 | using map_pmf_compose[of f g] by (simp add: comp_def) | |
| 483 | ||
| 59664 | 484 | 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" | 
| 485 | unfolding map_pmf_def by (rule bind_pmf_cong) auto | |
| 486 | ||
| 487 | lemma pmf_set_map: "set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf" | |
| 59665 | 488 | by (auto simp add: comp_def fun_eq_iff map_pmf_def) | 
| 59664 | 489 | |
| 59665 | 490 | lemma set_map_pmf[simp]: "set_pmf (map_pmf f M) = f`set_pmf M" | 
| 59664 | 491 | 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 | 492 | |
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 493 | lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)" | 
| 59664 | 494 | 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: 
59000diff
changeset | 495 | |
| 61634 | 496 | lemma measure_map_pmf[simp]: "measure (map_pmf f M) X = measure M (f -` X)" | 
| 497 | using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure) | |
| 498 | ||
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 499 | 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 | 500 | 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: 
59000diff
changeset | 501 | |
| 59023 | 502 | lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
 | 
| 59664 | 503 | proof (transfer fixing: f x) | 
| 59023 | 504 | fix p :: "'b measure" | 
| 505 | presume "prob_space p" | |
| 506 | then interpret prob_space p . | |
| 507 | presume "sets p = UNIV" | |
| 508 |   then show "ereal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
 | |
| 509 | by(simp add: measure_distr measurable_def emeasure_eq_measure) | |
| 510 | qed simp_all | |
| 511 | ||
| 512 | lemma nn_integral_pmf: "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = emeasure (measure_pmf p) A" | |
| 513 | proof - | |
| 514 | 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))" | |
| 515 | by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong) | |
| 516 |   also have "\<dots> = emeasure (measure_pmf p) (\<Union>x\<in>A \<inter> set_pmf p. {x})"
 | |
| 517 | by(subst emeasure_UN_countable)(auto simp add: emeasure_pmf_single disjoint_family_on_def) | |
| 518 |   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})"
 | |
| 519 | by(rule emeasure_Un_null_set[symmetric])(auto intro: in_null_sets_measure_pmfI) | |
| 520 | also have "\<dots> = emeasure (measure_pmf p) A" | |
| 521 | by(auto intro: arg_cong2[where f=emeasure]) | |
| 522 | finally show ?thesis . | |
| 523 | qed | |
| 524 | ||
| 60068 | 525 | lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)" | 
| 59664 | 526 | by transfer (simp add: distr_return) | 
| 527 | ||
| 528 | lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c" | |
| 529 | by transfer (auto simp: prob_space.distr_const) | |
| 530 | ||
| 60068 | 531 | lemma pmf_return [simp]: "pmf (return_pmf x) y = indicator {y} x"
 | 
| 59664 | 532 | by transfer (simp add: measure_return) | 
| 533 | ||
| 534 | lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x" | |
| 535 | unfolding return_pmf.rep_eq by (intro nn_integral_return) auto | |
| 536 | ||
| 537 | lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x" | |
| 538 | unfolding return_pmf.rep_eq by (intro emeasure_return) auto | |
| 539 | ||
| 540 | lemma return_pmf_inj[simp]: "return_pmf x = return_pmf y \<longleftrightarrow> x = y" | |
| 541 | by (metis insertI1 set_return_pmf singletonD) | |
| 542 | ||
| 59665 | 543 | lemma map_pmf_eq_return_pmf_iff: | 
| 544 | "map_pmf f p = return_pmf x \<longleftrightarrow> (\<forall>y \<in> set_pmf p. f y = x)" | |
| 545 | proof | |
| 546 | assume "map_pmf f p = return_pmf x" | |
| 547 | then have "set_pmf (map_pmf f p) = set_pmf (return_pmf x)" by simp | |
| 548 | then show "\<forall>y \<in> set_pmf p. f y = x" by auto | |
| 549 | next | |
| 550 | assume "\<forall>y \<in> set_pmf p. f y = x" | |
| 551 | then show "map_pmf f p = return_pmf x" | |
| 552 | unfolding map_pmf_const[symmetric, of _ p] by (intro map_pmf_cong) auto | |
| 553 | qed | |
| 554 | ||
| 59664 | 555 | definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))" | 
| 556 | ||
| 557 | lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b" | |
| 558 | unfolding pair_pmf_def pmf_bind pmf_return | |
| 559 |   apply (subst integral_measure_pmf[where A="{b}"])
 | |
| 560 | apply (auto simp: indicator_eq_0_iff) | |
| 561 |   apply (subst integral_measure_pmf[where A="{a}"])
 | |
| 562 | apply (auto simp: indicator_eq_0_iff setsum_nonneg_eq_0_iff pmf_nonneg) | |
| 563 | done | |
| 564 | ||
| 59665 | 565 | lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B" | 
| 59664 | 566 | unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto | 
| 567 | ||
| 568 | lemma measure_pmf_in_subprob_space[measurable (raw)]: | |
| 569 | "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))" | |
| 570 | by (simp add: space_subprob_algebra) intro_locales | |
| 571 | ||
| 572 | 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)" | |
| 573 | proof - | |
| 574 | have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. max 0 (f x) * indicator (A \<times> B) x \<partial>pair_pmf A B)" | |
| 575 | by (subst nn_integral_max_0[symmetric]) | |
| 59665 | 576 | (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE) | 
| 59664 | 577 | also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)" | 
| 578 | by (simp add: pair_pmf_def) | |
| 579 | also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) \<partial>B \<partial>A)" | |
| 580 | by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff) | |
| 581 | finally show ?thesis | |
| 582 | unfolding nn_integral_max_0 . | |
| 583 | qed | |
| 584 | ||
| 585 | lemma bind_pair_pmf: | |
| 586 | assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)" | |
| 587 | shows "measure_pmf (pair_pmf A B) \<guillemotright>= M = (measure_pmf A \<guillemotright>= (\<lambda>x. measure_pmf B \<guillemotright>= (\<lambda>y. M (x, y))))" | |
| 588 | (is "?L = ?R") | |
| 589 | proof (rule measure_eqI) | |
| 590 | have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)" | |
| 591 | using M[THEN measurable_space] by (simp_all add: space_pair_measure) | |
| 592 | ||
| 593 | note measurable_bind[where N="count_space UNIV", measurable] | |
| 594 | note measure_pmf_in_subprob_space[simp] | |
| 595 | ||
| 596 | have sets_eq_N: "sets ?L = N" | |
| 597 | by (subst sets_bind[OF sets_kernel[OF M']]) auto | |
| 598 | show "sets ?L = sets ?R" | |
| 599 | using measurable_space[OF M] | |
| 600 | by (simp add: sets_eq_N space_pair_measure space_subprob_algebra) | |
| 601 | fix X assume "X \<in> sets ?L" | |
| 602 | then have X[measurable]: "X \<in> sets N" | |
| 603 | unfolding sets_eq_N . | |
| 604 | then show "emeasure ?L X = emeasure ?R X" | |
| 605 | apply (simp add: emeasure_bind[OF _ M' X]) | |
| 606 | apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A] | |
| 60068 | 607 | nn_integral_measure_pmf_finite emeasure_nonneg one_ereal_def[symmetric]) | 
| 59664 | 608 | apply (subst emeasure_bind[OF _ _ X]) | 
| 609 | apply measurable | |
| 610 | apply (subst emeasure_bind[OF _ _ X]) | |
| 611 | apply measurable | |
| 612 | done | |
| 613 | qed | |
| 614 | ||
| 615 | lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A" | |
| 616 | by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') | |
| 617 | ||
| 618 | lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B" | |
| 619 | by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') | |
| 620 | ||
| 621 | lemma nn_integral_pmf': | |
| 622 | "inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)" | |
| 623 | by (subst nn_integral_bij_count_space[where g=f and B="f`A"]) | |
| 624 | (auto simp: bij_betw_def nn_integral_pmf) | |
| 625 | ||
| 626 | lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0" | |
| 627 | using pmf_nonneg[of M p] by simp | |
| 628 | ||
| 629 | lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0" | |
| 630 | using pmf_nonneg[of M p] by simp_all | |
| 631 | ||
| 632 | lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M" | |
| 633 | unfolding set_pmf_iff by simp | |
| 634 | ||
| 635 | 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" | |
| 636 | by (auto simp: pmf.rep_eq map_pmf_rep_eq measure_distr AE_measure_pmf_iff inj_onD | |
| 637 | intro!: measure_pmf.finite_measure_eq_AE) | |
| 638 | ||
| 60068 | 639 | lemma pmf_map_inj': "inj f \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x" | 
| 640 | apply(cases "x \<in> set_pmf M") | |
| 641 | apply(simp add: pmf_map_inj[OF subset_inj_on]) | |
| 642 | apply(simp add: pmf_eq_0_set_pmf[symmetric]) | |
| 643 | apply(auto simp add: pmf_eq_0_set_pmf dest: injD) | |
| 644 | done | |
| 645 | ||
| 646 | lemma pmf_map_outside: "x \<notin> f ` set_pmf M \<Longrightarrow> pmf (map_pmf f M) x = 0" | |
| 647 | unfolding pmf_eq_0_set_pmf by simp | |
| 648 | ||
| 59664 | 649 | subsection \<open> PMFs as function \<close> | 
| 59000 | 650 | |
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 651 | context | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 652 | fixes f :: "'a \<Rightarrow> real" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 653 | assumes nonneg: "\<And>x. 0 \<le> f x" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 654 | 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 | 655 | begin | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 656 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 657 | lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ereal \<circ> f)" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 658 | proof (intro conjI) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 659 |   have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
 | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 660 | by (simp split: split_indicator) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 661 | show "AE x in density (count_space UNIV) (ereal \<circ> f). | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 662 |     measure (density (count_space UNIV) (ereal \<circ> f)) {x} \<noteq> 0"
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59053diff
changeset | 663 | by (simp add: AE_density nonneg measure_def emeasure_density max_def) | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 664 | show "prob_space (density (count_space UNIV) (ereal \<circ> f))" | 
| 61169 | 665 | by standard (simp add: emeasure_density prob) | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 666 | qed simp | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 667 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 668 | 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 | 669 | proof transfer | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 670 |   have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
 | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 671 | by (simp split: split_indicator) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 672 |   fix x show "measure (density (count_space UNIV) (ereal \<circ> f)) {x} = f x"
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: 
59053diff
changeset | 673 | 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 | 674 | qed | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 675 | |
| 60068 | 676 | lemma set_embed_pmf: "set_pmf embed_pmf = {x. f x \<noteq> 0}"
 | 
| 677 | by(auto simp add: set_pmf_eq assms pmf_embed_pmf) | |
| 678 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 679 | end | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 680 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 681 | lemma embed_pmf_transfer: | 
| 58730 | 682 | "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ereal \<circ> f)) embed_pmf" | 
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 683 | 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 | 684 | |
| 59000 | 685 | lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)" | 
| 686 | proof (transfer, elim conjE) | |
| 687 |   fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
 | |
| 688 | assume "prob_space M" then interpret prob_space M . | |
| 689 |   show "M = density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))"
 | |
| 690 | proof (rule measure_eqI) | |
| 691 | fix A :: "'a set" | |
| 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: 
59665diff
changeset | 692 |     have "(\<integral>\<^sup>+ x. ereal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
 | 
| 59000 | 693 |       (\<integral>\<^sup>+ x. emeasure M {x} * indicator (A \<inter> {x. measure M {x} \<noteq> 0}) x \<partial>count_space UNIV)"
 | 
| 694 | by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator) | |
| 695 |     also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} \<partial>count_space (A \<inter> {x. measure M {x} \<noteq> 0}))"
 | |
| 696 | by (subst nn_integral_restrict_space[symmetric]) (auto simp: restrict_count_space) | |
| 697 |     also have "\<dots> = emeasure M (\<Union>x\<in>(A \<inter> {x. measure M {x} \<noteq> 0}). {x})"
 | |
| 698 | by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support) | |
| 699 | (auto simp: disjoint_family_on_def) | |
| 700 | also have "\<dots> = emeasure M A" | |
| 701 | using ae by (intro emeasure_eq_AE) auto | |
| 702 |     finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))) A"
 | |
| 703 | using emeasure_space_1 by (simp add: emeasure_density) | |
| 704 | qed simp | |
| 705 | qed | |
| 706 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 707 | lemma td_pmf_embed_pmf: | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 708 |   "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1}"
 | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 709 | unfolding type_definition_def | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 710 | proof safe | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 711 | fix p :: "'a pmf" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 712 | 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 | 713 | using measure_pmf.emeasure_space_1[of p] by simp | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 714 | then show *: "(\<integral>\<^sup>+ x. ereal (pmf p x) \<partial>count_space UNIV) = 1" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 715 | 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 | 716 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 717 | show "embed_pmf (pmf p) = p" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 718 | by (intro measure_pmf_inject[THEN iffD1]) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 719 | (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 | 720 | next | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 721 | 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 | 722 | then show "pmf (embed_pmf f) = f" | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 723 | by (auto intro!: pmf_embed_pmf) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 724 | qed (rule pmf_nonneg) | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 725 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 726 | end | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 727 | |
| 60068 | 728 | lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ereal (pmf p x) * f x \<partial>count_space UNIV" | 
| 729 | by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg) | |
| 730 | ||
| 58587 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 731 | locale pmf_as_function | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 732 | begin | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 733 | |
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 734 | setup_lifting td_pmf_embed_pmf | 
| 
5484f6079bcd
add type for probability mass functions, i.e. discrete probability distribution
 hoelzl parents: diff
changeset | 735 | |
| 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: 
59665diff
changeset | 736 | lemma set_pmf_transfer[transfer_rule]: | 
| 58730 | 737 | 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: 
59665diff
changeset | 738 |   shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
 | 
| 61808 | 739 | using \<open>bi_total A\<close> | 
| 58730 | 740 | by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff) | 
| 741 | metis+ | |
| 742 | ||
| 59000 | 743 | end | 
| 744 | ||
| 745 | context | |
| 746 | begin | |
| 747 | ||
| 748 | interpretation pmf_as_function . | |
| 749 | ||
| 750 | lemma pmf_eqI: "(\<And>i. pmf M i = pmf N i) \<Longrightarrow> M = N" | |
| 751 | by transfer auto | |
| 752 | ||
| 753 | lemma pmf_eq_iff: "M = N \<longleftrightarrow> (\<forall>i. pmf M i = pmf N i)" | |
| 754 | by (auto intro: pmf_eqI) | |
| 755 | ||
| 59664 | 756 | 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))" | 
| 757 | unfolding pmf_eq_iff pmf_bind | |
| 758 | proof | |
| 759 | fix i | |
| 760 | interpret B: prob_space "restrict_space B B" | |
| 761 | by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE) | |
| 762 | (auto simp: AE_measure_pmf_iff) | |
| 763 | interpret A: prob_space "restrict_space A A" | |
| 764 | by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE) | |
| 765 | (auto simp: AE_measure_pmf_iff) | |
| 766 | ||
| 767 | interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B" | |
| 768 | by unfold_locales | |
| 769 | ||
| 770 | 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)" | |
| 771 | by (rule integral_cong) (auto intro!: integral_pmf_restrict) | |
| 772 | also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)" | |
| 773 | by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2 | |
| 774 | countable_set_pmf borel_measurable_count_space) | |
| 775 | also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)" | |
| 776 | by (rule AB.Fubini_integral[symmetric]) | |
| 777 | (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2 | |
| 778 | simp: pmf_nonneg pmf_le_1 measurable_restrict_space1) | |
| 779 | also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)" | |
| 780 | by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2 | |
| 781 | countable_set_pmf borel_measurable_count_space) | |
| 782 | also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)" | |
| 783 | by (rule integral_cong) (auto intro!: integral_pmf_restrict[symmetric]) | |
| 784 | 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)" . | |
| 785 | qed | |
| 786 | ||
| 787 | lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)" | |
| 788 | proof (safe intro!: pmf_eqI) | |
| 789 | fix a :: "'a" and b :: "'b" | |
| 790 |   have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ereal)"
 | |
| 791 | by (auto split: split_indicator) | |
| 792 | ||
| 793 | have "ereal (pmf (pair_pmf (map_pmf f A) B) (a, b)) = | |
| 794 | ereal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))" | |
| 795 | unfolding pmf_pair ereal_pmf_map | |
| 796 | by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg | |
| 797 | emeasure_map_pmf[symmetric] del: emeasure_map_pmf) | |
| 798 | then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)" | |
| 799 | by simp | |
| 800 | qed | |
| 801 | ||
| 802 | lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)" | |
| 803 | proof (safe intro!: pmf_eqI) | |
| 804 | fix a :: "'a" and b :: "'b" | |
| 805 |   have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ereal)"
 | |
| 806 | by (auto split: split_indicator) | |
| 807 | ||
| 808 | have "ereal (pmf (pair_pmf A (map_pmf f B)) (a, b)) = | |
| 809 | ereal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))" | |
| 810 | unfolding pmf_pair ereal_pmf_map | |
| 811 | by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg | |
| 812 | emeasure_map_pmf[symmetric] del: emeasure_map_pmf) | |
| 813 | then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)" | |
| 814 | by simp | |
| 815 | qed | |
| 816 | ||
| 817 | 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)" | |
| 818 | by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta') | |
| 819 | ||
| 59000 | 820 | end | 
| 821 | ||
| 61634 | 822 | lemma pair_return_pmf1: "pair_pmf (return_pmf x) y = map_pmf (Pair x) y" | 
| 823 | by(simp add: pair_pmf_def bind_return_pmf map_pmf_def) | |
| 824 | ||
| 825 | lemma pair_return_pmf2: "pair_pmf x (return_pmf y) = map_pmf (\<lambda>x. (x, y)) x" | |
| 826 | by(simp add: pair_pmf_def bind_return_pmf map_pmf_def) | |
| 827 | ||
| 828 | 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))" | |
| 829 | by(simp add: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf) | |
| 830 | ||
| 831 | lemma pair_commute_pmf: "pair_pmf x y = map_pmf (\<lambda>(x, y). (y, x)) (pair_pmf y x)" | |
| 832 | unfolding pair_pmf_def by(subst bind_commute_pmf)(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf) | |
| 833 | ||
| 834 | lemma set_pmf_subset_singleton: "set_pmf p \<subseteq> {x} \<longleftrightarrow> p = return_pmf x"
 | |
| 835 | proof(intro iffI pmf_eqI) | |
| 836 | fix i | |
| 837 |   assume x: "set_pmf p \<subseteq> {x}"
 | |
| 838 |   hence *: "set_pmf p = {x}" using set_pmf_not_empty[of p] by auto
 | |
| 839 |   have "ereal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
 | |
| 840 | also have "\<dots> = \<integral>\<^sup>+ i. 1 \<partial>p" by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * ) | |
| 841 | also have "\<dots> = 1" by simp | |
| 842 | finally show "pmf p i = pmf (return_pmf x) i" using x | |
| 843 | by(auto split: split_indicator simp add: pmf_eq_0_set_pmf) | |
| 844 | qed auto | |
| 845 | ||
| 846 | lemma bind_eq_return_pmf: | |
| 847 | "bind_pmf p f = return_pmf x \<longleftrightarrow> (\<forall>y\<in>set_pmf p. f y = return_pmf x)" | |
| 848 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 849 | proof(intro iffI strip) | |
| 850 | fix y | |
| 851 | assume y: "y \<in> set_pmf p" | |
| 852 | assume "?lhs" | |
| 853 |   hence "set_pmf (bind_pmf p f) = {x}" by simp
 | |
| 854 |   hence "(\<Union>y\<in>set_pmf p. set_pmf (f y)) = {x}" by simp
 | |
| 855 |   hence "set_pmf (f y) \<subseteq> {x}" using y by auto
 | |
| 856 | thus "f y = return_pmf x" by(simp add: set_pmf_subset_singleton) | |
| 857 | next | |
| 858 | assume *: ?rhs | |
| 859 | show ?lhs | |
| 860 | proof(rule pmf_eqI) | |
| 861 | fix i | |
| 862 | have "ereal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ereal (pmf (f y) i) \<partial>p" by(simp add: ereal_pmf_bind) | |
| 863 | also have "\<dots> = \<integral>\<^sup>+ y. ereal (pmf (return_pmf x) i) \<partial>p" | |
| 864 | by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * ) | |
| 865 | also have "\<dots> = ereal (pmf (return_pmf x) i)" by simp | |
| 866 | finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i" by simp | |
| 867 | qed | |
| 868 | qed | |
| 869 | ||
| 870 | lemma pmf_False_conv_True: "pmf p False = 1 - pmf p True" | |
| 871 | proof - | |
| 872 |   have "pmf p False + pmf p True = measure p {False} + measure p {True}"
 | |
| 873 | by(simp add: measure_pmf_single) | |
| 874 |   also have "\<dots> = measure p ({False} \<union> {True})"
 | |
| 875 | by(subst measure_pmf.finite_measure_Union) simp_all | |
| 876 |   also have "{False} \<union> {True} = space p" by auto
 | |
| 877 | finally show ?thesis by simp | |
| 878 | qed | |
| 879 | ||
| 880 | lemma pmf_True_conv_False: "pmf p True = 1 - pmf p False" | |
| 881 | by(simp add: pmf_False_conv_True) | |
| 882 | ||
| 59664 | 883 | subsection \<open> Conditional Probabilities \<close> | 
| 884 | ||
| 59670 | 885 | lemma measure_pmf_zero_iff: "measure (measure_pmf p) s = 0 \<longleftrightarrow> set_pmf p \<inter> s = {}"
 | 
| 886 | by (subst measure_pmf.prob_eq_0) (auto simp: AE_measure_pmf_iff) | |
| 887 | ||
| 59664 | 888 | context | 
| 889 | fixes p :: "'a pmf" and s :: "'a set" | |
| 890 |   assumes not_empty: "set_pmf p \<inter> s \<noteq> {}"
 | |
| 891 | begin | |
| 892 | ||
| 893 | interpretation pmf_as_measure . | |
| 894 | ||
| 895 | lemma emeasure_measure_pmf_not_zero: "emeasure (measure_pmf p) s \<noteq> 0" | |
| 896 | proof | |
| 897 | assume "emeasure (measure_pmf p) s = 0" | |
| 898 | then have "AE x in measure_pmf p. x \<notin> s" | |
| 899 | by (rule AE_I[rotated]) auto | |
| 900 | with not_empty show False | |
| 901 | by (auto simp: AE_measure_pmf_iff) | |
| 902 | qed | |
| 903 | ||
| 904 | lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0" | |
| 905 | using emeasure_measure_pmf_not_zero unfolding measure_pmf.emeasure_eq_measure by simp | |
| 906 | ||
| 907 | lift_definition cond_pmf :: "'a pmf" is | |
| 908 | "uniform_measure (measure_pmf p) s" | |
| 909 | proof (intro conjI) | |
| 910 | show "prob_space (uniform_measure (measure_pmf p) s)" | |
| 911 | by (intro prob_space_uniform_measure) (auto simp: emeasure_measure_pmf_not_zero) | |
| 912 |   show "AE x in uniform_measure (measure_pmf p) s. measure (uniform_measure (measure_pmf p) s) {x} \<noteq> 0"
 | |
| 913 | by (simp add: emeasure_measure_pmf_not_zero measure_measure_pmf_not_zero AE_uniform_measure | |
| 914 | AE_measure_pmf_iff set_pmf.rep_eq) | |
| 915 | qed simp | |
| 916 | ||
| 917 | lemma pmf_cond: "pmf cond_pmf x = (if x \<in> s then pmf p x / measure p s else 0)" | |
| 918 | by transfer (simp add: emeasure_measure_pmf_not_zero pmf.rep_eq) | |
| 919 | ||
| 59665 | 920 | lemma set_cond_pmf[simp]: "set_pmf cond_pmf = set_pmf p \<inter> s" | 
| 59664 | 921 | by (auto simp add: set_pmf_iff pmf_cond measure_measure_pmf_not_zero split: split_if_asm) | 
| 922 | ||
| 923 | end | |
| 924 | ||
| 925 | lemma cond_map_pmf: | |
| 926 |   assumes "set_pmf p \<inter> f -` s \<noteq> {}"
 | |
| 927 | shows "cond_pmf (map_pmf f p) s = map_pmf f (cond_pmf p (f -` s))" | |
| 928 | proof - | |
| 929 |   have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
 | |
| 59665 | 930 | using assms by auto | 
| 59664 | 931 |   { fix x
 | 
| 932 | have "ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x) = | |
| 933 |       emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
 | |
| 934 | unfolding ereal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure) | |
| 935 |     also have "f -` s \<inter> f -` {x} = (if x \<in> s then f -` {x} else {})"
 | |
| 936 | by auto | |
| 937 |     also have "emeasure p (if x \<in> s then f -` {x} else {}) / emeasure p (f -` s) =
 | |
| 938 | ereal (pmf (cond_pmf (map_pmf f p) s) x)" | |
| 939 | using measure_measure_pmf_not_zero[OF *] | |
| 940 | by (simp add: pmf_cond[OF *] ereal_divide' ereal_pmf_map measure_pmf.emeasure_eq_measure[symmetric] | |
| 941 | del: ereal_divide) | |
| 942 | finally have "ereal (pmf (cond_pmf (map_pmf f p) s) x) = ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x)" | |
| 943 | by simp } | |
| 944 | then show ?thesis | |
| 945 | by (intro pmf_eqI) simp | |
| 946 | qed | |
| 947 | ||
| 948 | lemma bind_cond_pmf_cancel: | |
| 59670 | 949 |   assumes [simp]: "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
 | 
| 950 |   assumes [simp]: "\<And>y. y \<in> set_pmf q \<Longrightarrow> set_pmf p \<inter> {x. R x y} \<noteq> {}"
 | |
| 951 |   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}"
 | |
| 952 |   shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q"
 | |
| 59664 | 953 | proof (rule pmf_eqI) | 
| 59670 | 954 | fix i | 
| 955 |   have "ereal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
 | |
| 956 |     (\<integral>\<^sup>+x. ereal (pmf q i / measure p {x. R x i}) * ereal (indicator {x. R x i} x) \<partial>p)"
 | |
| 957 | by (auto simp add: ereal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf intro!: nn_integral_cong_AE) | |
| 958 |   also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
 | |
| 959 | by (simp add: pmf_nonneg measure_nonneg zero_ereal_def[symmetric] ereal_indicator | |
| 960 | nn_integral_cmult measure_pmf.emeasure_eq_measure) | |
| 961 | also have "\<dots> = pmf q i" | |
| 962 | by (cases "pmf q i = 0") (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero) | |
| 963 |   finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
 | |
| 964 | by simp | |
| 59664 | 965 | qed | 
| 966 | ||
| 967 | subsection \<open> Relator \<close> | |
| 968 | ||
| 969 | inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
 | |
| 970 | for R p q | |
| 971 | 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: 
59665diff
changeset | 972 | "\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y; | 
| 59664 | 973 | map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk> | 
| 974 | \<Longrightarrow> rel_pmf R p q" | |
| 975 | ||
| 59681 | 976 | lemma rel_pmfI: | 
| 977 | assumes R: "rel_set R (set_pmf p) (set_pmf q)" | |
| 978 | assumes eq: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow> | |
| 979 |     measure p {x. R x y} = measure q {y. R x y}"
 | |
| 980 | shows "rel_pmf R p q" | |
| 981 | proof | |
| 982 |   let ?pq = "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q {y. R x y}) (\<lambda>y. return_pmf (x, y)))"
 | |
| 983 |   have "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
 | |
| 984 | using R by (auto simp: rel_set_def) | |
| 985 | then show "\<And>x y. (x, y) \<in> set_pmf ?pq \<Longrightarrow> R x y" | |
| 986 | by auto | |
| 987 | show "map_pmf fst ?pq = p" | |
| 60068 | 988 | by (simp add: map_bind_pmf bind_return_pmf') | 
| 59681 | 989 | |
| 990 | show "map_pmf snd ?pq = q" | |
| 991 | using R eq | |
| 60068 | 992 | apply (simp add: bind_cond_pmf_cancel map_bind_pmf bind_return_pmf') | 
| 59681 | 993 | apply (rule bind_cond_pmf_cancel) | 
| 994 | apply (auto simp: rel_set_def) | |
| 995 | done | |
| 996 | qed | |
| 997 | ||
| 998 | lemma rel_pmf_imp_rel_set: "rel_pmf R p q \<Longrightarrow> rel_set R (set_pmf p) (set_pmf q)" | |
| 999 | by (force simp add: rel_pmf.simps rel_set_def) | |
| 1000 | ||
| 1001 | lemma rel_pmfD_measure: | |
| 1002 | assumes rel_R: "rel_pmf R p q" and R: "\<And>a b. R a b \<Longrightarrow> R a y \<longleftrightarrow> R x b" | |
| 1003 | assumes "x \<in> set_pmf p" "y \<in> set_pmf q" | |
| 1004 |   shows "measure p {x. R x y} = measure q {y. R x y}"
 | |
| 1005 | proof - | |
| 1006 | from rel_R obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" | |
| 1007 | and eq: "p = map_pmf fst pq" "q = map_pmf snd pq" | |
| 1008 | by (auto elim: rel_pmf.cases) | |
| 1009 |   have "measure p {x. R x y} = measure pq {x. R (fst x) y}"
 | |
| 1010 | by (simp add: eq map_pmf_rep_eq measure_distr) | |
| 1011 |   also have "\<dots> = measure pq {y. R x (snd y)}"
 | |
| 1012 | by (intro measure_pmf.finite_measure_eq_AE) | |
| 1013 | (auto simp: AE_measure_pmf_iff R dest!: pq) | |
| 1014 |   also have "\<dots> = measure q {y. R x y}"
 | |
| 1015 | by (simp add: eq map_pmf_rep_eq measure_distr) | |
| 1016 |   finally show "measure p {x. R x y} = measure q {y. R x y}" .
 | |
| 1017 | qed | |
| 1018 | ||
| 61634 | 1019 | lemma rel_pmf_measureD: | 
| 1020 | assumes "rel_pmf R p q" | |
| 1021 |   shows "measure (measure_pmf p) A \<le> measure (measure_pmf q) {y. \<exists>x\<in>A. R x y}" (is "?lhs \<le> ?rhs")
 | |
| 1022 | using assms | |
| 1023 | proof cases | |
| 1024 | fix pq | |
| 1025 | assume R: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" | |
| 1026 | and p[symmetric]: "map_pmf fst pq = p" | |
| 1027 | and q[symmetric]: "map_pmf snd pq = q" | |
| 1028 | have "?lhs = measure (measure_pmf pq) (fst -` A)" by(simp add: p) | |
| 1029 |   also have "\<dots> \<le> measure (measure_pmf pq) {y. \<exists>x\<in>A. R x (snd y)}"
 | |
| 1030 | by(rule measure_pmf.finite_measure_mono_AE)(auto 4 3 simp add: AE_measure_pmf_iff dest: R) | |
| 1031 | also have "\<dots> = ?rhs" by(simp add: q) | |
| 1032 | finally show ?thesis . | |
| 1033 | qed | |
| 1034 | ||
| 59681 | 1035 | lemma rel_pmf_iff_measure: | 
| 1036 | assumes "symp R" "transp R" | |
| 1037 | shows "rel_pmf R p q \<longleftrightarrow> | |
| 1038 | rel_set R (set_pmf p) (set_pmf q) \<and> | |
| 1039 |     (\<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})"
 | |
| 1040 | by (safe intro!: rel_pmf_imp_rel_set rel_pmfI) | |
| 1041 | (auto intro!: rel_pmfD_measure dest: sympD[OF \<open>symp R\<close>] transpD[OF \<open>transp R\<close>]) | |
| 1042 | ||
| 1043 | lemma quotient_rel_set_disjoint: | |
| 1044 |   "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: 
61424diff
changeset | 1045 |   using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C]
 | 
| 59681 | 1046 | by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE) | 
| 1047 | (blast dest: equivp_symp)+ | |
| 1048 | ||
| 1049 | lemma quotientD: "equiv X R \<Longrightarrow> A \<in> X // R \<Longrightarrow> x \<in> A \<Longrightarrow> A = R `` {x}"
 | |
| 1050 | by (metis Image_singleton_iff equiv_class_eq_iff quotientE) | |
| 1051 | ||
| 1052 | lemma rel_pmf_iff_equivp: | |
| 1053 | assumes "equivp R" | |
| 1054 |   shows "rel_pmf R p q \<longleftrightarrow> (\<forall>C\<in>UNIV // {(x, y). R x y}. measure p C = measure q C)"
 | |
| 1055 | (is "_ \<longleftrightarrow> (\<forall>C\<in>_//?R. _)") | |
| 1056 | proof (subst rel_pmf_iff_measure, safe) | |
| 1057 | show "symp R" "transp R" | |
| 1058 | using assms by (auto simp: equivp_reflp_symp_transp) | |
| 1059 | next | |
| 1060 | fix C assume C: "C \<in> UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)" | |
| 1061 |   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: 
61424diff
changeset | 1062 | |
| 59681 | 1063 | show "measure p C = measure q C" | 
| 1064 | proof cases | |
| 1065 |     assume "p \<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: 
61424diff
changeset | 1066 |     moreover then have "q \<inter> C = {}"
 | 
| 59681 | 1067 | using quotient_rel_set_disjoint[OF assms C R] by simp | 
| 1068 | ultimately show ?thesis | |
| 1069 | unfolding measure_pmf_zero_iff[symmetric] by simp | |
| 1070 | next | |
| 1071 |     assume "p \<inter> C \<noteq> {}"
 | |
| 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: 
61424diff
changeset | 1072 |     moreover then have "q \<inter> C \<noteq> {}"
 | 
| 59681 | 1073 | using quotient_rel_set_disjoint[OF assms C R] by simp | 
| 1074 | ultimately 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" | |
| 1075 | by auto | |
| 1076 | then have "R x y" | |
| 1077 | using in_quotient_imp_in_rel[of UNIV ?R C x y] C assms | |
| 1078 | by (simp add: equivp_equiv) | |
| 1079 |     with in_set eq have "measure p {x. R x y} = measure q {y. R x y}"
 | |
| 1080 | by auto | |
| 1081 |     moreover have "{y. R x y} = C"
 | |
| 61808 | 1082 | using assms \<open>x \<in> C\<close> C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv) | 
| 59681 | 1083 |     moreover have "{x. R x y} = C"
 | 
| 61808 | 1084 | using assms \<open>y \<in> C\<close> C quotientD[of UNIV "?R" C y] sympD[of R] | 
| 59681 | 1085 | by (auto simp add: equivp_equiv elim: equivpE) | 
| 1086 | ultimately show ?thesis | |
| 1087 | by auto | |
| 1088 | qed | |
| 1089 | next | |
| 1090 | assume eq: "\<forall>C\<in>UNIV // ?R. measure p C = measure q C" | |
| 1091 | show "rel_set R (set_pmf p) (set_pmf q)" | |
| 1092 | unfolding rel_set_def | |
| 1093 | proof safe | |
| 1094 | fix x assume x: "x \<in> set_pmf p" | |
| 1095 |     have "{y. R x y} \<in> UNIV // ?R"
 | |
| 1096 | by (auto simp: quotient_def) | |
| 1097 |     with eq have *: "measure q {y. R x y} = measure p {y. R x y}"
 | |
| 1098 | by auto | |
| 1099 |     have "measure q {y. R x y} \<noteq> 0"
 | |
| 1100 | using x assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp) | |
| 1101 | then show "\<exists>y\<in>set_pmf q. R x y" | |
| 1102 | unfolding measure_pmf_zero_iff by auto | |
| 1103 | next | |
| 1104 | fix y assume y: "y \<in> set_pmf q" | |
| 1105 |     have "{x. R x y} \<in> UNIV // ?R"
 | |
| 1106 | using assms by (auto simp: quotient_def dest: equivp_symp) | |
| 1107 |     with eq have *: "measure p {x. R x y} = measure q {x. R x y}"
 | |
| 1108 | by auto | |
| 1109 |     have "measure p {x. R x y} \<noteq> 0"
 | |
| 1110 | using y assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp) | |
| 1111 | then show "\<exists>x\<in>set_pmf p. R x y" | |
| 1112 | unfolding measure_pmf_zero_iff by auto | |
| 1113 | qed | |
| 1114 | ||
| 1115 | fix x y assume "x \<in> set_pmf p" "y \<in> set_pmf q" "R x y" | |
| 1116 |   have "{y. R x y} \<in> UNIV // ?R" "{x. R x y} = {y. R x y}"
 | |
| 61808 | 1117 | using assms \<open>R x y\<close> by (auto simp: quotient_def dest: equivp_symp equivp_transp) | 
| 59681 | 1118 |   with eq show "measure p {x. R x y} = measure q {y. R x y}"
 | 
| 1119 | by auto | |
| 1120 | qed | |
| 1121 | ||
| 59664 | 1122 | bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf | 
| 1123 | proof - | |
| 1124 | 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: 
59665diff
changeset | 1125 | show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose) | 
| 59664 | 1126 | 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" | 
| 1127 | by (intro map_pmf_cong refl) | |
| 1128 | ||
| 1129 | show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf" | |
| 1130 | by (rule pmf_set_map) | |
| 1131 | ||
| 60595 | 1132 | show "(card_of (set_pmf p), natLeq) \<in> ordLeq" for p :: "'s pmf" | 
| 1133 | proof - | |
| 59664 | 1134 | have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq" | 
| 1135 | by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"]) | |
| 1136 | (auto intro: countable_set_pmf) | |
| 1137 | also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq" | |
| 1138 | by (metis Field_natLeq card_of_least natLeq_Well_order) | |
| 60595 | 1139 | finally show ?thesis . | 
| 1140 | qed | |
| 59664 | 1141 | |
| 1142 | show "\<And>R. rel_pmf R = | |
| 1143 |          (BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf fst))\<inverse>\<inverse> OO
 | |
| 1144 |          BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf snd)"
 | |
| 1145 | by (auto simp add: fun_eq_iff BNF_Def.Grp_def OO_def rel_pmf.simps) | |
| 1146 | ||
| 60595 | 1147 | show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)" | 
| 1148 | for R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool" | |
| 1149 | proof - | |
| 1150 |     { fix p q r
 | |
| 1151 | assume pq: "rel_pmf R p q" | |
| 1152 | and qr:"rel_pmf S q r" | |
| 1153 | from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" | |
| 1154 | and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto | |
| 1155 | from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z" | |
| 1156 | 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: 
61424diff
changeset | 1157 | |
| 60595 | 1158 |       def pr \<equiv> "bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\<lambda>yz. return_pmf (fst xy, snd yz)))"
 | 
| 1159 |       have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
 | |
| 1160 | 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: 
61424diff
changeset | 1161 | |
| 60595 | 1162 | have "rel_pmf (R OO S) p r" | 
| 1163 | proof (rule rel_pmf.intros) | |
| 1164 | fix x z assume "(x, z) \<in> pr" | |
| 1165 | then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr" | |
| 1166 | by (auto simp: q pr_welldefined pr_def split_beta) | |
| 1167 | with pq qr show "(R OO S) x z" | |
| 1168 | by blast | |
| 1169 | next | |
| 1170 |         have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {yz. fst yz = y}))"
 | |
| 1171 | by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_pmf_comp) | |
| 1172 | then show "map_pmf snd pr = r" | |
| 1173 | unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) (auto simp: eq_commute) | |
| 1174 | qed (simp add: pr_def map_bind_pmf split_beta map_pmf_def[symmetric] p map_pmf_comp) | |
| 1175 | } | |
| 1176 | then show ?thesis | |
| 1177 | by(auto simp add: le_fun_def) | |
| 1178 | qed | |
| 59664 | 1179 | qed (fact natLeq_card_order natLeq_cinfinite)+ | 
| 1180 | ||
| 61634 | 1181 | lemma map_pmf_idI: "(\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_pmf f p = p" | 
| 1182 | by(simp cong: pmf.map_cong) | |
| 1183 | ||
| 59665 | 1184 | lemma rel_pmf_conj[simp]: | 
| 1185 | "rel_pmf (\<lambda>x y. P \<and> Q x y) x y \<longleftrightarrow> P \<and> rel_pmf Q x y" | |
| 1186 | "rel_pmf (\<lambda>x y. Q x y \<and> P) x y \<longleftrightarrow> P \<and> rel_pmf Q x y" | |
| 1187 | using set_pmf_not_empty by (fastforce simp: pmf.in_rel subset_eq)+ | |
| 1188 | ||
| 1189 | lemma rel_pmf_top[simp]: "rel_pmf top = top" | |
| 1190 | by (auto simp: pmf.in_rel[abs_def] fun_eq_iff map_fst_pair_pmf map_snd_pair_pmf | |
| 1191 | intro: exI[of _ "pair_pmf x y" for x y]) | |
| 1192 | ||
| 59664 | 1193 | lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)" | 
| 1194 | proof safe | |
| 1195 | fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M" | |
| 1196 | then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b" | |
| 1197 | and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq" | |
| 1198 | by (force elim: rel_pmf.cases) | |
| 1199 |   moreover have "set_pmf (return_pmf x) = {x}"
 | |
| 59665 | 1200 | by simp | 
| 61808 | 1201 | with \<open>a \<in> M\<close> have "(x, a) \<in> pq" | 
| 59665 | 1202 | by (force simp: eq) | 
| 59664 | 1203 | with * show "R x a" | 
| 1204 | by auto | |
| 1205 | qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"] | |
| 59665 | 1206 | simp: map_fst_pair_pmf map_snd_pair_pmf) | 
| 59664 | 1207 | |
| 1208 | lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)" | |
| 1209 | by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1) | |
| 1210 | ||
| 1211 | lemma rel_return_pmf[simp]: "rel_pmf R (return_pmf x1) (return_pmf x2) = R x1 x2" | |
| 1212 | unfolding rel_pmf_return_pmf2 set_return_pmf by simp | |
| 1213 | ||
| 1214 | lemma rel_pmf_False[simp]: "rel_pmf (\<lambda>x y. False) x y = False" | |
| 1215 | unfolding pmf.in_rel fun_eq_iff using set_pmf_not_empty by fastforce | |
| 1216 | ||
| 1217 | lemma rel_pmf_rel_prod: | |
| 1218 | "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'" | |
| 1219 | proof safe | |
| 1220 | assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')" | |
| 1221 | 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" | |
| 1222 | and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'" | |
| 1223 | by (force elim: rel_pmf.cases) | |
| 1224 | show "rel_pmf R A B" | |
| 1225 | proof (rule rel_pmf.intros) | |
| 1226 | let ?f = "\<lambda>(a, b). (fst a, fst b)" | |
| 1227 | have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd" | |
| 1228 | by auto | |
| 1229 | ||
| 1230 | show "map_pmf fst (map_pmf ?f pq) = A" | |
| 1231 | by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf) | |
| 1232 | show "map_pmf snd (map_pmf ?f pq) = B" | |
| 1233 | by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf) | |
| 1234 | ||
| 1235 | fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)" | |
| 1236 | then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq" | |
| 59665 | 1237 | by auto | 
| 59664 | 1238 | from pq[OF this] show "R a b" .. | 
| 1239 | qed | |
| 1240 | show "rel_pmf S A' B'" | |
| 1241 | proof (rule rel_pmf.intros) | |
| 1242 | let ?f = "\<lambda>(a, b). (snd a, snd b)" | |
| 1243 | have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd" | |
| 1244 | by auto | |
| 1245 | ||
| 1246 | show "map_pmf fst (map_pmf ?f pq) = A'" | |
| 1247 | by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf) | |
| 1248 | show "map_pmf snd (map_pmf ?f pq) = B'" | |
| 1249 | by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf) | |
| 1250 | ||
| 1251 | fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)" | |
| 1252 | then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq" | |
| 59665 | 1253 | by auto | 
| 59664 | 1254 | from pq[OF this] show "S c d" .. | 
| 1255 | qed | |
| 1256 | next | |
| 1257 | assume "rel_pmf R A B" "rel_pmf S A' B'" | |
| 1258 | then obtain Rpq Spq | |
| 1259 | where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b" | |
| 1260 | "map_pmf fst Rpq = A" "map_pmf snd Rpq = B" | |
| 1261 | and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b" | |
| 1262 | "map_pmf fst Spq = A'" "map_pmf snd Spq = B'" | |
| 1263 | by (force elim: rel_pmf.cases) | |
| 1264 | ||
| 1265 | let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))" | |
| 1266 | let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)" | |
| 1267 | 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))" | |
| 1268 | by auto | |
| 1269 | ||
| 1270 | show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')" | |
| 1271 | by (rule rel_pmf.intros[where pq="?pq"]) | |
| 59665 | 1272 | (auto simp: map_snd_pair_pmf map_fst_pair_pmf map_pmf_comp Rpq Spq | 
| 59664 | 1273 | map_pair) | 
| 1274 | qed | |
| 1275 | ||
| 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: 
59665diff
changeset | 1276 | lemma rel_pmf_reflI: | 
| 59664 | 1277 | assumes "\<And>x. x \<in> set_pmf p \<Longrightarrow> P x x" | 
| 1278 | shows "rel_pmf P p p" | |
| 59665 | 1279 | by (rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. (x, x)) p"]) | 
| 1280 | (auto simp add: pmf.map_comp o_def assms) | |
| 59664 | 1281 | |
| 61634 | 1282 | lemma rel_pmf_bij_betw: | 
| 1283 | assumes f: "bij_betw f (set_pmf p) (set_pmf q)" | |
| 1284 | and eq: "\<And>x. x \<in> set_pmf p \<Longrightarrow> pmf p x = pmf q (f x)" | |
| 1285 | shows "rel_pmf (\<lambda>x y. f x = y) p q" | |
| 1286 | proof(rule rel_pmf.intros) | |
| 1287 | let ?pq = "map_pmf (\<lambda>x. (x, f x)) p" | |
| 1288 | show "map_pmf fst ?pq = p" by(simp add: pmf.map_comp o_def) | |
| 1289 | ||
| 1290 | have "map_pmf f p = q" | |
| 1291 | proof(rule pmf_eqI) | |
| 1292 | fix i | |
| 1293 | show "pmf (map_pmf f p) i = pmf q i" | |
| 1294 | proof(cases "i \<in> set_pmf q") | |
| 1295 | case True | |
| 1296 | with f obtain j where "i = f j" "j \<in> set_pmf p" | |
| 1297 | by(auto simp add: bij_betw_def image_iff) | |
| 1298 | thus ?thesis using f by(simp add: bij_betw_def pmf_map_inj eq) | |
| 1299 | next | |
| 1300 | case False thus ?thesis | |
| 1301 | by(subst pmf_map_outside)(auto simp add: set_pmf_iff eq[symmetric]) | |
| 1302 | qed | |
| 1303 | qed | |
| 1304 | then show "map_pmf snd ?pq = q" by(simp add: pmf.map_comp o_def) | |
| 1305 | qed auto | |
| 1306 | ||
| 59664 | 1307 | context | 
| 1308 | begin | |
| 1309 | ||
| 1310 | interpretation pmf_as_measure . | |
| 1311 | ||
| 1312 | definition "join_pmf M = bind_pmf M (\<lambda>x. x)" | |
| 1313 | ||
| 1314 | lemma bind_eq_join_pmf: "bind_pmf M f = join_pmf (map_pmf f M)" | |
| 1315 | unfolding join_pmf_def bind_map_pmf .. | |
| 1316 | ||
| 1317 | lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id" | |
| 1318 | by (simp add: join_pmf_def id_def) | |
| 1319 | ||
| 1320 | lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)" | |
| 1321 | unfolding join_pmf_def pmf_bind .. | |
| 1322 | ||
| 1323 | lemma ereal_pmf_join: "ereal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)" | |
| 1324 | unfolding join_pmf_def ereal_pmf_bind .. | |
| 1325 | ||
| 59665 | 1326 | lemma set_pmf_join_pmf[simp]: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)" | 
| 1327 | by (simp add: join_pmf_def) | |
| 59664 | 1328 | |
| 1329 | lemma join_return_pmf: "join_pmf (return_pmf M) = M" | |
| 1330 | by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq) | |
| 1331 | ||
| 1332 | lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)" | |
| 1333 | by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf) | |
| 1334 | ||
| 1335 | lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A" | |
| 1336 | by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') | |
| 1337 | ||
| 1338 | end | |
| 1339 | ||
| 1340 | lemma rel_pmf_joinI: | |
| 1341 | assumes "rel_pmf (rel_pmf P) p q" | |
| 1342 | shows "rel_pmf P (join_pmf p) (join_pmf q)" | |
| 1343 | proof - | |
| 1344 | from assms obtain pq where p: "p = map_pmf fst pq" | |
| 1345 | and q: "q = map_pmf snd pq" | |
| 1346 | and P: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> rel_pmf P x y" | |
| 1347 | 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: 
59665diff
changeset | 1348 | from P obtain PQ | 
| 59664 | 1349 | 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" | 
| 1350 | and x: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf fst (PQ x y) = x" | |
| 1351 | and y: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf snd (PQ x y) = y" | |
| 1352 | by(metis rel_pmf.simps) | |
| 1353 | ||
| 1354 | let ?r = "bind_pmf pq (\<lambda>(x, y). PQ x y)" | |
| 59665 | 1355 | have "\<And>a b. (a, b) \<in> set_pmf ?r \<Longrightarrow> P a b" by (auto intro: PQ) | 
| 59664 | 1356 | moreover have "map_pmf fst ?r = join_pmf p" "map_pmf snd ?r = join_pmf q" | 
| 1357 | by (simp_all add: p q x y join_pmf_def map_bind_pmf bind_map_pmf split_def cong: bind_pmf_cong) | |
| 1358 | ultimately show ?thesis .. | |
| 1359 | qed | |
| 1360 | ||
| 1361 | lemma rel_pmf_bindI: | |
| 1362 | assumes pq: "rel_pmf R p q" | |
| 1363 | and fg: "\<And>x y. R x y \<Longrightarrow> rel_pmf P (f x) (g y)" | |
| 1364 | shows "rel_pmf P (bind_pmf p f) (bind_pmf q g)" | |
| 1365 | unfolding bind_eq_join_pmf | |
| 1366 | by (rule rel_pmf_joinI) | |
| 1367 | (auto simp add: pmf.rel_map intro: pmf.rel_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp, OF _ pq] fg) | |
| 1368 | ||
| 61808 | 1369 | text \<open> | 
| 59664 | 1370 |   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: 
59665diff
changeset | 1371 | 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: 
59665diff
changeset | 1372 | Theoretical Computer Science 12(1):19--37, 1980, | 
| 59664 | 1373 |   @{url "http://dx.doi.org/10.1016/0304-3975(80)90003-1"}
 | 
| 61808 | 1374 | \<close> | 
| 59664 | 1375 | |
| 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: 
59665diff
changeset | 1376 | lemma | 
| 59664 | 1377 | assumes *: "rel_pmf R p q" | 
| 1378 | and refl: "reflp R" and trans: "transp R" | |
| 1379 |   shows measure_Ici: "measure p {y. R x y} \<le> measure q {y. R x y}" (is ?thesis1)
 | |
| 1380 |   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)
 | |
| 1381 | proof - | |
| 1382 | from * obtain pq | |
| 1383 | where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" | |
| 1384 | and p: "p = map_pmf fst pq" | |
| 1385 | and q: "q = map_pmf snd pq" | |
| 1386 | by cases auto | |
| 1387 | show ?thesis1 ?thesis2 unfolding p q map_pmf_rep_eq using refl trans | |
| 1388 | by(auto 4 3 simp add: measure_distr reflpD AE_measure_pmf_iff intro!: measure_pmf.finite_measure_mono_AE dest!: pq elim: transpE) | |
| 1389 | qed | |
| 1390 | ||
| 1391 | lemma rel_pmf_inf: | |
| 1392 | fixes p q :: "'a pmf" | |
| 1393 | assumes 1: "rel_pmf R p q" | |
| 1394 | assumes 2: "rel_pmf R q p" | |
| 1395 | and refl: "reflp R" and trans: "transp R" | |
| 1396 | shows "rel_pmf (inf R R\<inverse>\<inverse>) p q" | |
| 59681 | 1397 | proof (subst rel_pmf_iff_equivp, safe) | 
| 1398 | show "equivp (inf R R\<inverse>\<inverse>)" | |
| 1399 | 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: 
61424diff
changeset | 1400 | |
| 59681 | 1401 |   fix C assume "C \<in> UNIV // {(x, y). inf R R\<inverse>\<inverse> x y}"
 | 
| 1402 |   then obtain x where C: "C = {y. R x y \<and> R y x}"
 | |
| 1403 | by (auto elim: quotientE) | |
| 1404 | ||
| 59670 | 1405 | let ?R = "\<lambda>x y. R x y \<and> R y x" | 
| 1406 |   let ?\<mu>R = "\<lambda>y. measure q {x. ?R x y}"
 | |
| 59681 | 1407 |   have "measure p {y. ?R x y} = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})"
 | 
| 1408 | by(auto intro!: arg_cong[where f="measure p"]) | |
| 1409 |   also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}"
 | |
| 1410 | by (rule measure_pmf.finite_measure_Diff) auto | |
| 1411 |   also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}"
 | |
| 1412 | using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi) | |
| 1413 |   also have "measure p {y. R x y} = measure q {y. R x y}"
 | |
| 1414 | using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici) | |
| 1415 |   also have "measure q {y. R x y} - measure q {y. R x y \<and> \<not> R y x} =
 | |
| 1416 |     measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})"
 | |
| 1417 | by(rule measure_pmf.finite_measure_Diff[symmetric]) auto | |
| 1418 | also have "\<dots> = ?\<mu>R x" | |
| 1419 | by(auto intro!: arg_cong[where f="measure q"]) | |
| 1420 | finally show "measure p C = measure q C" | |
| 1421 | by (simp add: C conj_commute) | |
| 59664 | 1422 | qed | 
| 1423 | ||
| 1424 | lemma rel_pmf_antisym: | |
| 1425 | fixes p q :: "'a pmf" | |
| 1426 | assumes 1: "rel_pmf R p q" | |
| 1427 | assumes 2: "rel_pmf R q p" | |
| 1428 | and refl: "reflp R" and trans: "transp R" and antisym: "antisymP R" | |
| 1429 | shows "p = q" | |
| 1430 | proof - | |
| 1431 | from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf) | |
| 1432 | also have "inf R R\<inverse>\<inverse> = op =" | |
| 59665 | 1433 | using refl antisym by (auto intro!: ext simp add: reflpD dest: antisymD) | 
| 59664 | 1434 | finally show ?thesis unfolding pmf.rel_eq . | 
| 1435 | qed | |
| 1436 | ||
| 1437 | lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)" | |
| 1438 | by(blast intro: reflpI rel_pmf_reflI reflpD) | |
| 1439 | ||
| 1440 | lemma antisymP_rel_pmf: | |
| 1441 | "\<lbrakk> reflp R; transp R; antisymP R \<rbrakk> | |
| 1442 | \<Longrightarrow> antisymP (rel_pmf R)" | |
| 1443 | by(rule antisymI)(blast intro: rel_pmf_antisym) | |
| 1444 | ||
| 1445 | lemma transp_rel_pmf: | |
| 1446 | assumes "transp R" | |
| 1447 | shows "transp (rel_pmf R)" | |
| 1448 | proof (rule transpI) | |
| 1449 | fix x y z | |
| 1450 | assume "rel_pmf R x y" and "rel_pmf R y z" | |
| 1451 | hence "rel_pmf (R OO R) x z" by (simp add: pmf.rel_compp relcompp.relcompI) | |
| 1452 | thus "rel_pmf R x z" | |
| 1453 | using assms by (metis (no_types) pmf.rel_mono rev_predicate2D transp_relcompp_less_eq) | |
| 1454 | qed | |
| 1455 | ||
| 1456 | subsection \<open> Distributions \<close> | |
| 1457 | ||
| 59000 | 1458 | context | 
| 1459 | begin | |
| 1460 | ||
| 1461 | interpretation pmf_as_function . | |
| 1462 | ||
| 59093 | 1463 | subsubsection \<open> Bernoulli Distribution \<close> | 
| 1464 | ||
| 59000 | 1465 | lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is | 
| 1466 | "\<lambda>p b. ((\<lambda>p. if b then p else 1 - p) \<circ> min 1 \<circ> max 0) p" | |
| 1467 |   by (auto simp: nn_integral_count_space_finite[where A="{False, True}"] UNIV_bool
 | |
| 1468 | split: split_max split_min) | |
| 1469 | ||
| 1470 | lemma pmf_bernoulli_True[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) True = p" | |
| 1471 | by transfer simp | |
| 1472 | ||
| 1473 | lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p" | |
| 1474 | by transfer simp | |
| 1475 | ||
| 1476 | lemma set_pmf_bernoulli: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV" | |
| 1477 | by (auto simp add: set_pmf_iff UNIV_bool) | |
| 1478 | ||
| 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: 
59665diff
changeset | 1479 | lemma nn_integral_bernoulli_pmf[simp]: | 
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 1480 | 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: 
59000diff
changeset | 1481 | 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: 
59000diff
changeset | 1482 | by (subst nn_integral_measure_pmf_support[of UNIV]) | 
| 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 1483 | (auto simp: UNIV_bool field_simps) | 
| 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 1484 | |
| 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: 
59665diff
changeset | 1485 | lemma integral_bernoulli_pmf[simp]: | 
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 1486 | assumes [simp]: "0 \<le> p" "p \<le> 1" | 
| 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 1487 | 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: 
59000diff
changeset | 1488 | 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: 
59000diff
changeset | 1489 | |
| 59525 | 1490 | lemma pmf_bernoulli_half [simp]: "pmf (bernoulli_pmf (1 / 2)) x = 1 / 2" | 
| 1491 | by(cases x) simp_all | |
| 1492 | ||
| 1493 | lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV" | |
| 1494 | by(rule measure_eqI)(simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure nn_integral_count_space_finite sets_uniform_count_measure) | |
| 1495 | ||
| 59093 | 1496 | subsubsection \<open> Geometric Distribution \<close> | 
| 1497 | ||
| 60602 | 1498 | context | 
| 1499 | fixes p :: real assumes p[arith]: "0 < p" "p \<le> 1" | |
| 1500 | begin | |
| 1501 | ||
| 1502 | lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p" | |
| 59000 | 1503 | proof | 
| 60602 | 1504 | have "(\<Sum>i. ereal (p * (1 - p) ^ i)) = ereal (p * (1 / (1 - (1 - p))))" | 
| 1505 | by (intro sums_suminf_ereal sums_mult geometric_sums) auto | |
| 1506 | then show "(\<integral>\<^sup>+ x. ereal ((1 - p)^x * p) \<partial>count_space UNIV) = 1" | |
| 59000 | 1507 | by (simp add: nn_integral_count_space_nat field_simps) | 
| 1508 | qed simp | |
| 1509 | ||
| 60602 | 1510 | lemma pmf_geometric[simp]: "pmf geometric_pmf n = (1 - p)^n * p" | 
| 59000 | 1511 | by transfer rule | 
| 1512 | ||
| 60602 | 1513 | end | 
| 1514 | ||
| 1515 | 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: 
61424diff
changeset | 1516 | by (auto simp: set_pmf_iff) | 
| 59000 | 1517 | |
| 59093 | 1518 | subsubsection \<open> Uniform Multiset Distribution \<close> | 
| 1519 | ||
| 59000 | 1520 | context | 
| 1521 |   fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
 | |
| 1522 | begin | |
| 1523 | ||
| 1524 | lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M" | |
| 1525 | proof | |
| 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: 
59665diff
changeset | 1526 | show "(\<integral>\<^sup>+ x. ereal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1" | 
| 59000 | 1527 | using M_not_empty | 
| 1528 | by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size | |
| 1529 | setsum_divide_distrib[symmetric]) | |
| 1530 | (auto simp: size_multiset_overloaded_eq intro!: setsum.cong) | |
| 1531 | qed simp | |
| 1532 | ||
| 1533 | lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M" | |
| 1534 | by transfer rule | |
| 1535 | ||
| 60495 | 1536 | lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_mset M" | 
| 59000 | 1537 | by (auto simp: set_pmf_iff) | 
| 1538 | ||
| 1539 | end | |
| 1540 | ||
| 59093 | 1541 | subsubsection \<open> Uniform Distribution \<close> | 
| 1542 | ||
| 59000 | 1543 | context | 
| 1544 |   fixes S :: "'a set" assumes S_not_empty: "S \<noteq> {}" and S_finite: "finite S"
 | |
| 1545 | begin | |
| 1546 | ||
| 1547 | lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S" | |
| 1548 | proof | |
| 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: 
59665diff
changeset | 1549 | show "(\<integral>\<^sup>+ x. ereal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1" | 
| 59000 | 1550 | using S_not_empty S_finite by (subst nn_integral_count_space'[of S]) auto | 
| 1551 | qed simp | |
| 1552 | ||
| 1553 | lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S" | |
| 1554 | by transfer rule | |
| 1555 | ||
| 1556 | lemma set_pmf_of_set[simp]: "set_pmf pmf_of_set = S" | |
| 1557 | using S_finite S_not_empty by (auto simp: set_pmf_iff) | |
| 1558 | ||
| 61634 | 1559 | 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: 
59000diff
changeset | 1560 | 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: 
59000diff
changeset | 1561 | |
| 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: 
61424diff
changeset | 1562 | lemma nn_integral_pmf_of_set': | 
| 60068 | 1563 | "(\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0) \<Longrightarrow> nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S" | 
| 1564 | apply(subst nn_integral_measure_pmf_finite, simp_all add: S_finite) | |
| 1565 | apply(simp add: setsum_ereal_left_distrib[symmetric]) | |
| 1566 | apply(subst ereal_divide', simp add: S_not_empty S_finite) | |
| 1567 | apply(simp add: ereal_times_divide_eq one_ereal_def[symmetric]) | |
| 1568 | done | |
| 1569 | ||
| 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: 
61424diff
changeset | 1570 | lemma nn_integral_pmf_of_set: | 
| 60068 | 1571 | "nn_integral (measure_pmf pmf_of_set) f = setsum (max 0 \<circ> f) S / card S" | 
| 1572 | apply(subst nn_integral_max_0[symmetric]) | |
| 1573 | apply(subst nn_integral_pmf_of_set') | |
| 1574 | apply simp_all | |
| 1575 | done | |
| 1576 | ||
| 1577 | lemma integral_pmf_of_set: | |
| 1578 | "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S" | |
| 1579 | apply(simp add: real_lebesgue_integral_def integrable_measure_pmf_finite nn_integral_pmf_of_set S_finite) | |
| 1580 | apply(subst real_of_ereal_minus') | |
| 1581 | apply(simp add: ereal_max_0 S_finite del: ereal_max) | |
| 1582 | apply(simp add: ereal_max_0 S_finite S_not_empty del: ereal_max) | |
| 1583 | apply(simp add: field_simps S_finite S_not_empty) | |
| 1584 | apply(subst setsum.distrib[symmetric]) | |
| 1585 | apply(rule setsum.cong; simp_all) | |
| 1586 | done | |
| 1587 | ||
| 61634 | 1588 | lemma emeasure_pmf_of_set: | 
| 1589 | "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S" | |
| 1590 | apply(subst nn_integral_indicator[symmetric], simp) | |
| 1591 | apply(subst nn_integral_pmf_of_set) | |
| 1592 | apply(simp add: o_def max_def ereal_indicator[symmetric] S_not_empty S_finite real_of_nat_indicator[symmetric] of_nat_setsum[symmetric] setsum_indicator_eq_card del: of_nat_setsum) | |
| 1593 | done | |
| 1594 | ||
| 59000 | 1595 | end | 
| 1596 | ||
| 60068 | 1597 | lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
 | 
| 1598 | by(rule pmf_eqI)(simp add: indicator_def) | |
| 1599 | ||
| 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: 
61424diff
changeset | 1600 | lemma map_pmf_of_set_inj: | 
| 60068 | 1601 | assumes f: "inj_on f A" | 
| 1602 |   and [simp]: "A \<noteq> {}" "finite A"
 | |
| 1603 | shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs") | |
| 1604 | proof(rule pmf_eqI) | |
| 1605 | fix i | |
| 1606 | show "pmf ?lhs i = pmf ?rhs i" | |
| 1607 | proof(cases "i \<in> f ` A") | |
| 1608 | case True | |
| 1609 | then obtain i' where "i = f i'" "i' \<in> A" by auto | |
| 1610 | thus ?thesis using f by(simp add: card_image pmf_map_inj) | |
| 1611 | next | |
| 1612 | case False | |
| 1613 | hence "pmf ?lhs i = 0" by(simp add: pmf_eq_0_set_pmf set_map_pmf) | |
| 1614 | moreover have "pmf ?rhs i = 0" using False by simp | |
| 1615 | ultimately show ?thesis by simp | |
| 1616 | qed | |
| 1617 | qed | |
| 1618 | ||
| 1619 | lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV" | |
| 1620 | by(rule pmf_eqI) simp_all | |
| 1621 | ||
| 61634 | 1622 | |
| 1623 | ||
| 1624 | lemma measure_pmf_of_set: | |
| 1625 |   assumes "S \<noteq> {}" "finite S"
 | |
| 1626 | shows "measure (measure_pmf (pmf_of_set S)) A = card (S \<inter> A) / card S" | |
| 1627 | using emeasure_pmf_of_set[OF assms, of A] | |
| 1628 | unfolding measure_pmf.emeasure_eq_measure by simp | |
| 1629 | ||
| 59093 | 1630 | subsubsection \<open> Poisson Distribution \<close> | 
| 1631 | ||
| 1632 | context | |
| 1633 | fixes rate :: real assumes rate_pos: "0 < rate" | |
| 1634 | begin | |
| 1635 | ||
| 1636 | 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: 
59667diff
changeset | 1637 | proof (* by Manuel Eberl *) | 
| 59093 | 1638 | have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp | 
| 59557 | 1639 | by (simp add: field_simps divide_inverse [symmetric]) | 
| 59093 | 1640 | have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x * exp (-rate) \<partial>count_space UNIV) = | 
| 1641 | exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)" | |
| 1642 | by (simp add: field_simps nn_integral_cmult[symmetric]) | |
| 1643 | also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)" | |
| 1644 | by (simp_all add: nn_integral_count_space_nat suminf_ereal summable suminf_ereal_finite) | |
| 1645 | also have "... = exp rate" unfolding exp_def | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 1646 | by (simp add: field_simps divide_inverse [symmetric]) | 
| 59093 | 1647 | also have "ereal (exp (-rate)) * ereal (exp rate) = 1" | 
| 1648 | by (simp add: mult_exp_exp) | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 1649 | finally show "(\<integral>\<^sup>+ x. ereal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" . | 
| 59093 | 1650 | qed (simp add: rate_pos[THEN less_imp_le]) | 
| 1651 | ||
| 1652 | lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)" | |
| 1653 | by transfer rule | |
| 1654 | ||
| 1655 | lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV" | |
| 1656 | using rate_pos by (auto simp: set_pmf_iff) | |
| 1657 | ||
| 59000 | 1658 | end | 
| 1659 | ||
| 59093 | 1660 | subsubsection \<open> Binomial Distribution \<close> | 
| 1661 | ||
| 1662 | context | |
| 1663 | fixes n :: nat and p :: real assumes p_nonneg: "0 \<le> p" and p_le_1: "p \<le> 1" | |
| 1664 | begin | |
| 1665 | ||
| 1666 | lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)" | |
| 1667 | proof | |
| 1668 | have "(\<integral>\<^sup>+k. ereal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) = | |
| 1669 | ereal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))" | |
| 1670 | using p_le_1 p_nonneg by (subst nn_integral_count_space') auto | |
| 1671 | 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: 
61424diff
changeset | 1672 | by (subst binomial_ring) (simp add: atLeast0AtMost) | 
| 59093 | 1673 | finally show "(\<integral>\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1" | 
| 1674 | by simp | |
| 1675 | qed (insert p_nonneg p_le_1, simp) | |
| 1676 | ||
| 1677 | lemma pmf_binomial[simp]: "pmf binomial_pmf k = (n choose k) * p^k * (1 - p)^(n - k)" | |
| 1678 | by transfer rule | |
| 1679 | ||
| 1680 | lemma set_pmf_binomial_eq: "set_pmf binomial_pmf = (if p = 0 then {0} else if p = 1 then {n} else {.. n})"
 | |
| 1681 | using p_nonneg p_le_1 unfolding set_eq_iff set_pmf_iff pmf_binomial by (auto simp: set_pmf_iff) | |
| 1682 | ||
| 1683 | end | |
| 1684 | ||
| 1685 | end | |
| 1686 | ||
| 1687 | lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
 | |
| 1688 | by (simp add: set_pmf_binomial_eq) | |
| 1689 | ||
| 1690 | lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
 | |
| 1691 | by (simp add: set_pmf_binomial_eq) | |
| 1692 | ||
| 1693 | lemma set_pmf_binomial[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (binomial_pmf n p) = {..n}"
 | |
| 1694 | by (simp add: set_pmf_binomial_eq) | |
| 1695 | ||
| 61634 | 1696 | context begin interpretation lifting_syntax . | 
| 1697 | ||
| 1698 | lemma bind_pmf_parametric [transfer_rule]: | |
| 1699 | "(rel_pmf A ===> (A ===> rel_pmf B) ===> rel_pmf B) bind_pmf bind_pmf" | |
| 1700 | by(blast intro: rel_pmf_bindI dest: rel_funD) | |
| 1701 | ||
| 1702 | lemma return_pmf_parametric [transfer_rule]: "(A ===> rel_pmf A) return_pmf return_pmf" | |
| 1703 | by(rule rel_funI) simp | |
| 1704 | ||
| 59000 | 1705 | end | 
| 61634 | 1706 | |
| 1707 | end |