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