| author | wenzelm | 
| Sun, 04 Sep 2011 19:36:19 +0200 | |
| changeset 44706 | fe319b45315c | 
| parent 43920 | cedb5cb948fd | 
| child 44890 | 22f665a2e91c | 
| permissions | -rw-r--r-- | 
| 42148 | 1 | (* Title: HOL/Probability/Probability_Measure.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 5 | ||
| 42148 | 6 | header {*Probability measure*}
 | 
| 42067 | 7 | |
| 42148 | 8 | theory Probability_Measure | 
| 43556 
0d78c8d31d0d
move conditional expectation to its own theory file
 hoelzl parents: 
43340diff
changeset | 9 | imports Lebesgue_Measure | 
| 35582 | 10 | begin | 
| 11 | ||
| 12 | locale prob_space = measure_space + | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 13 | assumes measure_space_1: "measure M (space M) = 1" | 
| 38656 | 14 | |
| 15 | sublocale prob_space < finite_measure | |
| 16 | proof | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 17 | from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp | 
| 38656 | 18 | qed | 
| 19 | ||
| 40859 | 20 | abbreviation (in prob_space) "events \<equiv> sets M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 21 | abbreviation (in prob_space) "prob \<equiv> \<mu>'" | 
| 40859 | 22 | abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 23 | abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M" | 
| 35582 | 24 | |
| 40859 | 25 | definition (in prob_space) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 26 | "distribution X A = \<mu>' (X -` A \<inter> space M)" | 
| 35582 | 27 | |
| 40859 | 28 | abbreviation (in prob_space) | 
| 36624 | 29 | "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))" | 
| 35582 | 30 | |
| 43339 | 31 | lemma (in prob_space) prob_space_cong: | 
| 32 | assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M" | |
| 33 | shows "prob_space N" | |
| 34 | proof - | |
| 35 | interpret N: measure_space N by (intro measure_space_cong assms) | |
| 36 | show ?thesis by default (insert assms measure_space_1, simp) | |
| 37 | qed | |
| 38 | ||
| 39097 | 39 | lemma (in prob_space) distribution_cong: | 
| 40 | assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x" | |
| 41 | shows "distribution X = distribution Y" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 42 | unfolding distribution_def fun_eq_iff | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 43 | using assms by (auto intro!: arg_cong[where f="\<mu>'"]) | 
| 39097 | 44 | |
| 45 | lemma (in prob_space) joint_distribution_cong: | |
| 46 | assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" | |
| 47 | assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" | |
| 48 | shows "joint_distribution X Y = joint_distribution X' Y'" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 49 | unfolding distribution_def fun_eq_iff | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 50 | using assms by (auto intro!: arg_cong[where f="\<mu>'"]) | 
| 39097 | 51 | |
| 40859 | 52 | lemma (in prob_space) distribution_id[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 53 | "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 54 | by (auto simp: distribution_def intro!: arg_cong[where f=prob]) | 
| 40859 | 55 | |
| 56 | lemma (in prob_space) prob_space: "prob (space M) = 1" | |
| 43920 | 57 | using measure_space_1 unfolding \<mu>'_def by (simp add: one_ereal_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 58 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 59 | lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 60 | using bounded_measure[of A] by (simp add: prob_space) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 61 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 62 | lemma (in prob_space) distribution_positive[simp, intro]: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 63 | "0 \<le> distribution X A" unfolding distribution_def by auto | 
| 35582 | 64 | |
| 43339 | 65 | lemma (in prob_space) not_zero_less_distribution[simp]: | 
| 66 | "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0" | |
| 67 | using distribution_positive[of X A] by arith | |
| 68 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 69 | lemma (in prob_space) joint_distribution_remove[simp]: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 70 |     "joint_distribution X X {(x, x)} = distribution X {x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 71 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 72 | |
| 43339 | 73 | lemma (in prob_space) not_empty: "space M \<noteq> {}"
 | 
| 74 | using prob_space empty_measure' by auto | |
| 75 | ||
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42902diff
changeset | 76 | lemma (in prob_space) measure_le_1: "X \<in> sets M \<Longrightarrow> \<mu> X \<le> 1" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42902diff
changeset | 77 | unfolding measure_space_1[symmetric] | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42902diff
changeset | 78 | using sets_into_space | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42902diff
changeset | 79 | by (intro measure_mono) auto | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42902diff
changeset | 80 | |
| 43339 | 81 | lemma (in prob_space) AE_I_eq_1: | 
| 82 |   assumes "\<mu> {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
 | |
| 83 | shows "AE x. P x" | |
| 84 | proof (rule AE_I) | |
| 85 |   show "\<mu> (space M - {x \<in> space M. P x}) = 0"
 | |
| 86 | using assms measure_space_1 by (simp add: measure_compl) | |
| 87 | qed (insert assms, auto) | |
| 88 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 89 | lemma (in prob_space) distribution_1: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 90 | "distribution X A \<le> 1" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 91 | unfolding distribution_def by simp | 
| 35582 | 92 | |
| 40859 | 93 | lemma (in prob_space) prob_compl: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 94 | assumes A: "A \<in> events" | 
| 38656 | 95 | shows "prob (space M - A) = 1 - prob A" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 96 | using finite_measure_compl[OF A] by (simp add: prob_space) | 
| 35582 | 97 | |
| 40859 | 98 | lemma (in prob_space) prob_space_increasing: "increasing M prob" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 99 | by (auto intro!: finite_measure_mono simp: increasing_def) | 
| 35582 | 100 | |
| 40859 | 101 | lemma (in prob_space) prob_zero_union: | 
| 35582 | 102 | assumes "s \<in> events" "t \<in> events" "prob t = 0" | 
| 103 | shows "prob (s \<union> t) = prob s" | |
| 38656 | 104 | using assms | 
| 35582 | 105 | proof - | 
| 106 | have "prob (s \<union> t) \<le> prob s" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 107 | using finite_measure_subadditive[of s t] assms by auto | 
| 35582 | 108 | moreover have "prob (s \<union> t) \<ge> prob s" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 109 | using assms by (blast intro: finite_measure_mono) | 
| 35582 | 110 | ultimately show ?thesis by simp | 
| 111 | qed | |
| 112 | ||
| 40859 | 113 | lemma (in prob_space) prob_eq_compl: | 
| 35582 | 114 | assumes "s \<in> events" "t \<in> events" | 
| 115 | assumes "prob (space M - s) = prob (space M - t)" | |
| 116 | shows "prob s = prob t" | |
| 38656 | 117 | using assms prob_compl by auto | 
| 35582 | 118 | |
| 40859 | 119 | lemma (in prob_space) prob_one_inter: | 
| 35582 | 120 | assumes events:"s \<in> events" "t \<in> events" | 
| 121 | assumes "prob t = 1" | |
| 122 | shows "prob (s \<inter> t) = prob s" | |
| 123 | proof - | |
| 38656 | 124 | have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)" | 
| 125 | using events assms prob_compl[of "t"] by (auto intro!: prob_zero_union) | |
| 126 | also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)" | |
| 127 | by blast | |
| 128 | finally show "prob (s \<inter> t) = prob s" | |
| 129 | using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s]) | |
| 35582 | 130 | qed | 
| 131 | ||
| 40859 | 132 | lemma (in prob_space) prob_eq_bigunion_image: | 
| 35582 | 133 | assumes "range f \<subseteq> events" "range g \<subseteq> events" | 
| 134 | assumes "disjoint_family f" "disjoint_family g" | |
| 135 | assumes "\<And> n :: nat. prob (f n) = prob (g n)" | |
| 136 | shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))" | |
| 137 | using assms | |
| 138 | proof - | |
| 38656 | 139 | have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 140 | by (rule finite_measure_UNION[OF assms(1,3)]) | 
| 38656 | 141 | have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 142 | by (rule finite_measure_UNION[OF assms(2,4)]) | 
| 38656 | 143 | show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp | 
| 35582 | 144 | qed | 
| 145 | ||
| 40859 | 146 | lemma (in prob_space) prob_countably_zero: | 
| 35582 | 147 | assumes "range c \<subseteq> events" | 
| 148 | assumes "\<And> i. prob (c i) = 0" | |
| 38656 | 149 | shows "prob (\<Union> i :: nat. c i) = 0" | 
| 150 | proof (rule antisym) | |
| 151 | show "prob (\<Union> i :: nat. c i) \<le> 0" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 152 | using finite_measure_countably_subadditive[OF assms(1)] | 
| 38656 | 153 | by (simp add: assms(2) suminf_zero summable_zero) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 154 | qed simp | 
| 35582 | 155 | |
| 40859 | 156 | lemma (in prob_space) prob_equiprobable_finite_unions: | 
| 38656 | 157 | assumes "s \<in> events" | 
| 158 |   assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
 | |
| 35582 | 159 |   assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
 | 
| 38656 | 160 |   shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
 | 
| 35582 | 161 | proof (cases "s = {}")
 | 
| 38656 | 162 | case False hence "\<exists> x. x \<in> s" by blast | 
| 35582 | 163 | from someI_ex[OF this] assms | 
| 164 |   have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
 | |
| 165 |   have "prob s = (\<Sum> x \<in> s. prob {x})"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 166 | using finite_measure_finite_singleton[OF s_finite] by simp | 
| 35582 | 167 |   also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
 | 
| 38656 | 168 |   also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
 | 
| 169 | using setsum_constant assms by (simp add: real_eq_of_nat) | |
| 35582 | 170 | finally show ?thesis by simp | 
| 38656 | 171 | qed simp | 
| 35582 | 172 | |
| 40859 | 173 | lemma (in prob_space) prob_real_sum_image_fn: | 
| 35582 | 174 | assumes "e \<in> events" | 
| 175 | assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events" | |
| 176 | assumes "finite s" | |
| 38656 | 177 |   assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
 | 
| 178 | assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)" | |
| 35582 | 179 | shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))" | 
| 180 | proof - | |
| 38656 | 181 | have e: "e = (\<Union> i \<in> s. e \<inter> f i)" | 
| 182 | using `e \<in> events` sets_into_space upper by blast | |
| 183 | hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp | |
| 184 | also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 185 | proof (rule finite_measure_finite_Union) | 
| 38656 | 186 | show "finite s" by fact | 
| 187 | show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact | |
| 188 | show "disjoint_family_on (\<lambda>i. e \<inter> f i) s" | |
| 189 | using disjoint by (auto simp: disjoint_family_on_def) | |
| 190 | qed | |
| 191 | finally show ?thesis . | |
| 35582 | 192 | qed | 
| 193 | ||
| 42199 | 194 | lemma (in prob_space) prob_space_vimage: | 
| 195 | assumes S: "sigma_algebra S" | |
| 196 | assumes T: "T \<in> measure_preserving M S" | |
| 197 | shows "prob_space S" | |
| 35582 | 198 | proof - | 
| 42199 | 199 | interpret S: measure_space S | 
| 200 | using S and T by (rule measure_space_vimage) | |
| 38656 | 201 | show ?thesis | 
| 42199 | 202 | proof | 
| 203 | from T[THEN measure_preservingD2] | |
| 204 | have "T -` space S \<inter> space M = space M" | |
| 205 | by (auto simp: measurable_def) | |
| 206 | with T[THEN measure_preservingD, of "space S", symmetric] | |
| 207 | show "measure S (space S) = 1" | |
| 208 | using measure_space_1 by simp | |
| 35582 | 209 | qed | 
| 210 | qed | |
| 211 | ||
| 42988 | 212 | lemma prob_space_unique_Int_stable: | 
| 213 |   fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
 | |
| 214 | assumes E: "Int_stable E" "space E \<in> sets E" | |
| 215 | and M: "prob_space M" "space M = space E" "sets M = sets (sigma E)" | |
| 216 | and N: "prob_space N" "space N = space E" "sets N = sets (sigma E)" | |
| 217 | and eq: "\<And>X. X \<in> sets E \<Longrightarrow> finite_measure.\<mu>' M X = finite_measure.\<mu>' N X" | |
| 218 | assumes "X \<in> sets (sigma E)" | |
| 219 | shows "finite_measure.\<mu>' M X = finite_measure.\<mu>' N X" | |
| 220 | proof - | |
| 221 | interpret M!: prob_space M by fact | |
| 222 | interpret N!: prob_space N by fact | |
| 223 | have "measure M X = measure N X" | |
| 224 | proof (rule measure_unique_Int_stable[OF `Int_stable E`]) | |
| 225 | show "range (\<lambda>i. space M) \<subseteq> sets E" "incseq (\<lambda>i. space M)" "(\<Union>i. space M) = space E" | |
| 226 | using E M N by auto | |
| 227 | show "\<And>i. M.\<mu> (space M) \<noteq> \<infinity>" | |
| 228 | using M.measure_space_1 by simp | |
| 229 | show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = M.\<mu>\<rparr>" | |
| 230 | using E M N by (auto intro!: M.measure_space_cong) | |
| 231 | show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = N.\<mu>\<rparr>" | |
| 232 | using E M N by (auto intro!: N.measure_space_cong) | |
| 233 |     { fix X assume "X \<in> sets E"
 | |
| 234 | then have "X \<in> sets (sigma E)" | |
| 235 | by (auto simp: sets_sigma sigma_sets.Basic) | |
| 236 | with eq[OF `X \<in> sets E`] M N show "M.\<mu> X = N.\<mu> X" | |
| 237 | by (simp add: M.finite_measure_eq N.finite_measure_eq) } | |
| 238 | qed fact | |
| 239 | with `X \<in> sets (sigma E)` M N show ?thesis | |
| 240 | by (simp add: M.finite_measure_eq N.finite_measure_eq) | |
| 241 | qed | |
| 242 | ||
| 42199 | 243 | lemma (in prob_space) distribution_prob_space: | 
| 244 | assumes X: "random_variable S X" | |
| 43920 | 245 | shows "prob_space (S\<lparr>measure := ereal \<circ> distribution X\<rparr>)" (is "prob_space ?S") | 
| 42199 | 246 | proof (rule prob_space_vimage) | 
| 247 | show "X \<in> measure_preserving M ?S" | |
| 248 | using X | |
| 249 | unfolding measure_preserving_def distribution_def_raw | |
| 250 | by (auto simp: finite_measure_eq measurable_sets) | |
| 251 | show "sigma_algebra ?S" using X by simp | |
| 252 | qed | |
| 253 | ||
| 40859 | 254 | lemma (in prob_space) AE_distribution: | 
| 43920 | 255 | assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := ereal \<circ> distribution X\<rparr>. Q x" | 
| 40859 | 256 | shows "AE x. Q (X x)" | 
| 257 | proof - | |
| 43920 | 258 | interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space) | 
| 40859 | 259 |   obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
 | 
| 260 | using assms unfolding X.almost_everywhere_def by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 261 | from X[unfolded measurable_def] N show "AE x. Q (X x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 262 | by (intro AE_I'[where N="X -` N \<inter> space M"]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 263 | (auto simp: finite_measure_eq distribution_def measurable_sets) | 
| 40859 | 264 | qed | 
| 265 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 266 | lemma (in prob_space) distribution_eq_integral: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 267 | "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 268 | using finite_measure_eq[of "X -` A \<inter> space M"] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 269 | by (auto simp: measurable_sets distribution_def) | 
| 35582 | 270 | |
| 43339 | 271 | lemma (in prob_space) expectation_less: | 
| 272 | assumes [simp]: "integrable M X" | |
| 273 | assumes gt: "\<forall>x\<in>space M. X x < b" | |
| 274 | shows "expectation X < b" | |
| 275 | proof - | |
| 276 | have "expectation X < expectation (\<lambda>x. b)" | |
| 277 | using gt measure_space_1 | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
43339diff
changeset | 278 | by (intro integral_less_AE_space) auto | 
| 43339 | 279 | then show ?thesis using prob_space by simp | 
| 280 | qed | |
| 281 | ||
| 282 | lemma (in prob_space) expectation_greater: | |
| 283 | assumes [simp]: "integrable M X" | |
| 284 | assumes gt: "\<forall>x\<in>space M. a < X x" | |
| 285 | shows "a < expectation X" | |
| 286 | proof - | |
| 287 | have "expectation (\<lambda>x. a) < expectation X" | |
| 288 | using gt measure_space_1 | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
43339diff
changeset | 289 | by (intro integral_less_AE_space) auto | 
| 43339 | 290 | then show ?thesis using prob_space by simp | 
| 291 | qed | |
| 292 | ||
| 293 | lemma convex_le_Inf_differential: | |
| 294 | fixes f :: "real \<Rightarrow> real" | |
| 295 | assumes "convex_on I f" | |
| 296 | assumes "x \<in> interior I" "y \<in> I" | |
| 297 |   shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
 | |
| 298 | (is "_ \<ge> _ + Inf (?F x) * (y - x)") | |
| 299 | proof - | |
| 300 | show ?thesis | |
| 301 | proof (cases rule: linorder_cases) | |
| 302 | assume "x < y" | |
| 303 | moreover | |
| 304 | have "open (interior I)" by auto | |
| 305 | from openE[OF this `x \<in> interior I`] guess e . note e = this | |
| 306 | moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)" | |
| 307 | ultimately have "x < t" "t < y" "t \<in> ball x e" | |
| 308 | by (auto simp: mem_ball dist_real_def field_simps split: split_min) | |
| 309 | with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto | |
| 310 | ||
| 311 | have "open (interior I)" by auto | |
| 312 | from openE[OF this `x \<in> interior I`] guess e . | |
| 313 | moreover def K \<equiv> "x - e / 2" | |
| 314 | with `0 < e` have "K \<in> ball x e" "K < x" by (auto simp: mem_ball dist_real_def) | |
| 315 | ultimately have "K \<in> I" "K < x" "x \<in> I" | |
| 316 | using interior_subset[of I] `x \<in> interior I` by auto | |
| 317 | ||
| 318 | have "Inf (?F x) \<le> (f x - f y) / (x - y)" | |
| 319 | proof (rule Inf_lower2) | |
| 320 | show "(f x - f t) / (x - t) \<in> ?F x" | |
| 321 | using `t \<in> I` `x < t` by auto | |
| 322 | show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" | |
| 323 | using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y` by (rule convex_on_diff) | |
| 324 | next | |
| 325 | fix y assume "y \<in> ?F x" | |
| 326 | with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]] | |
| 327 | show "(f K - f x) / (K - x) \<le> y" by auto | |
| 328 | qed | |
| 329 | then show ?thesis | |
| 330 | using `x < y` by (simp add: field_simps) | |
| 331 | next | |
| 332 | assume "y < x" | |
| 333 | moreover | |
| 334 | have "open (interior I)" by auto | |
| 335 | from openE[OF this `x \<in> interior I`] guess e . note e = this | |
| 336 | moreover def t \<equiv> "x + e / 2" | |
| 337 | ultimately have "x < t" "t \<in> ball x e" | |
| 338 | by (auto simp: mem_ball dist_real_def field_simps) | |
| 339 | with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto | |
| 340 | ||
| 341 | have "(f x - f y) / (x - y) \<le> Inf (?F x)" | |
| 342 | proof (rule Inf_greatest) | |
| 343 | have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" | |
| 344 | using `y < x` by (auto simp: field_simps) | |
| 345 | also | |
| 346 | fix z assume "z \<in> ?F x" | |
| 347 | with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]] | |
| 348 | have "(f y - f x) / (y - x) \<le> z" by auto | |
| 349 | finally show "(f x - f y) / (x - y) \<le> z" . | |
| 350 | next | |
| 351 | have "open (interior I)" by auto | |
| 352 | from openE[OF this `x \<in> interior I`] guess e . note e = this | |
| 353 | then have "x + e / 2 \<in> ball x e" by (auto simp: mem_ball dist_real_def) | |
| 354 |       with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" by auto
 | |
| 355 |       then show "?F x \<noteq> {}" by blast
 | |
| 356 | qed | |
| 357 | then show ?thesis | |
| 358 | using `y < x` by (simp add: field_simps) | |
| 359 | qed simp | |
| 360 | qed | |
| 361 | ||
| 362 | lemma (in prob_space) jensens_inequality: | |
| 363 | fixes a b :: real | |
| 364 | assumes X: "integrable M X" "X ` space M \<subseteq> I" | |
| 365 |   assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV"
 | |
| 366 | assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q" | |
| 367 | shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" | |
| 368 | proof - | |
| 369 |   let "?F x" = "Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))"
 | |
| 370 |   from not_empty X(2) have "I \<noteq> {}" by auto
 | |
| 371 | ||
| 372 | from I have "open I" by auto | |
| 373 | ||
| 374 | note I | |
| 375 | moreover | |
| 376 |   { assume "I \<subseteq> {a <..}"
 | |
| 377 | with X have "a < expectation X" | |
| 378 | by (intro expectation_greater) auto } | |
| 379 | moreover | |
| 380 |   { assume "I \<subseteq> {..< b}"
 | |
| 381 | with X have "expectation X < b" | |
| 382 | by (intro expectation_less) auto } | |
| 383 | ultimately have "expectation X \<in> I" | |
| 384 | by (elim disjE) (auto simp: subset_eq) | |
| 385 | moreover | |
| 386 |   { fix y assume y: "y \<in> I"
 | |
| 387 | with q(2) `open I` have "Sup ((\<lambda>x. q x + ?F x * (y - x)) ` I) = q y" | |
| 388 | by (auto intro!: Sup_eq_maximum convex_le_Inf_differential image_eqI[OF _ y] simp: interior_open) } | |
| 389 | ultimately have "q (expectation X) = Sup ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)" | |
| 390 | by simp | |
| 391 | also have "\<dots> \<le> expectation (\<lambda>w. q (X w))" | |
| 392 | proof (rule Sup_least) | |
| 393 |     show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}"
 | |
| 394 |       using `I \<noteq> {}` by auto
 | |
| 395 | next | |
| 396 | fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I" | |
| 397 | then guess x .. note x = this | |
| 398 | have "q x + ?F x * (expectation X - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))" | |
| 399 | using prob_space | |
| 400 | by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X) | |
| 401 | also have "\<dots> \<le> expectation (\<lambda>w. q (X w))" | |
| 402 | using `x \<in> I` `open I` X(2) | |
| 403 | by (intro integral_mono integral_add integral_cmult integral_diff | |
| 404 | lebesgue_integral_const X q convex_le_Inf_differential) | |
| 405 | (auto simp: interior_open) | |
| 406 | finally show "k \<le> expectation (\<lambda>w. q (X w))" using x by auto | |
| 407 | qed | |
| 408 | finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" . | |
| 409 | qed | |
| 410 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 411 | lemma (in prob_space) distribution_eq_translated_integral: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 412 | assumes "random_variable S X" "A \<in> sets S" | 
| 43920 | 413 | shows "distribution X A = integral\<^isup>P (S\<lparr>measure := ereal \<circ> distribution X\<rparr>) (indicator A)" | 
| 35582 | 414 | proof - | 
| 43920 | 415 | interpret S: prob_space "S\<lparr>measure := ereal \<circ> distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 416 | using assms(1) by (rule distribution_prob_space) | 
| 35582 | 417 | show ?thesis | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 418 | using S.positive_integral_indicator(1)[of A] assms by simp | 
| 35582 | 419 | qed | 
| 420 | ||
| 40859 | 421 | lemma (in prob_space) finite_expectation1: | 
| 422 | assumes f: "finite (X`space M)" and rv: "random_variable borel X" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 423 |   shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 424 | proof (subst integral_on_finite) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 425 | show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 426 |   show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 427 |     "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 428 | using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto | 
| 38656 | 429 | qed | 
| 35582 | 430 | |
| 40859 | 431 | lemma (in prob_space) finite_expectation: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 432 | assumes "finite (X`space M)" "random_variable borel X" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 433 |   shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
 | 
| 38656 | 434 | using assms unfolding distribution_def using finite_expectation1 by auto | 
| 435 | ||
| 40859 | 436 | lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0: | 
| 35582 | 437 |   assumes "{x} \<in> events"
 | 
| 38656 | 438 |   assumes "prob {x} = 1"
 | 
| 35582 | 439 |   assumes "{y} \<in> events"
 | 
| 440 | assumes "y \<noteq> x" | |
| 441 |   shows "prob {y} = 0"
 | |
| 442 |   using prob_one_inter[of "{y}" "{x}"] assms by auto
 | |
| 443 | ||
| 40859 | 444 | lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
 | 
| 38656 | 445 | unfolding distribution_def by simp | 
| 446 | ||
| 40859 | 447 | lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1" | 
| 38656 | 448 | proof - | 
| 449 | have "X -` X ` space M \<inter> space M = space M" by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 450 | thus ?thesis unfolding distribution_def by (simp add: prob_space) | 
| 38656 | 451 | qed | 
| 452 | ||
| 40859 | 453 | lemma (in prob_space) distribution_one: | 
| 454 | assumes "random_variable M' X" and "A \<in> sets M'" | |
| 38656 | 455 | shows "distribution X A \<le> 1" | 
| 456 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 457 | have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 458 | using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 459 | thus ?thesis by (simp add: prob_space) | 
| 38656 | 460 | qed | 
| 461 | ||
| 40859 | 462 | lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0: | 
| 35582 | 463 | assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X" | 
| 38656 | 464 | (is "random_variable ?S X") | 
| 465 |   assumes "distribution X {x} = 1"
 | |
| 35582 | 466 | assumes "y \<noteq> x" | 
| 467 |   shows "distribution X {y} = 0"
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 468 | proof cases | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 469 |   { fix x have "X -` {x} \<inter> space M \<in> sets M"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 470 | proof cases | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 471 | assume "x \<in> X`space M" with X show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 472 | by (auto simp: measurable_def image_iff) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 473 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 474 |       assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 475 | then show ?thesis by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 476 | qed } note single = this | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 477 |   have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 478 |     "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 479 | using `y \<noteq> x` by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 480 | with finite_measure_inter_full_set[OF single single, of x y] assms(2) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 481 | show ?thesis by (auto simp: distribution_def prob_space) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 482 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 483 |   assume "{y} \<notin> sets ?S"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 484 |   then have "X -` {y} \<inter> space M = {}" by auto
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 485 |   thus "distribution X {y} = 0" unfolding distribution_def by auto
 | 
| 35582 | 486 | qed | 
| 487 | ||
| 40859 | 488 | lemma (in prob_space) joint_distribution_Times_le_fst: | 
| 489 | assumes X: "random_variable MX X" and Y: "random_variable MY Y" | |
| 490 | and A: "A \<in> sets MX" and B: "B \<in> sets MY" | |
| 491 | shows "joint_distribution X Y (A \<times> B) \<le> distribution X A" | |
| 492 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 493 | proof (intro finite_measure_mono) | 
| 40859 | 494 | show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force | 
| 495 | show "X -` A \<inter> space M \<in> events" | |
| 496 | using X A unfolding measurable_def by simp | |
| 497 | have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M = | |
| 498 | (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto | |
| 499 | qed | |
| 500 | ||
| 501 | lemma (in prob_space) joint_distribution_commute: | |
| 502 | "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 503 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) | 
| 40859 | 504 | |
| 505 | lemma (in prob_space) joint_distribution_Times_le_snd: | |
| 506 | assumes X: "random_variable MX X" and Y: "random_variable MY Y" | |
| 507 | and A: "A \<in> sets MX" and B: "B \<in> sets MY" | |
| 508 | shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B" | |
| 509 | using assms | |
| 510 | by (subst joint_distribution_commute) | |
| 511 | (simp add: swap_product joint_distribution_Times_le_fst) | |
| 512 | ||
| 513 | lemma (in prob_space) random_variable_pairI: | |
| 514 | assumes "random_variable MX X" | |
| 515 | assumes "random_variable MY Y" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 516 | shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))" | 
| 40859 | 517 | proof | 
| 518 | interpret MX: sigma_algebra MX using assms by simp | |
| 519 | interpret MY: sigma_algebra MY using assms by simp | |
| 520 | interpret P: pair_sigma_algebra MX MY by default | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 521 | show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default | 
| 40859 | 522 | have sa: "sigma_algebra M" by default | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 523 | show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)" | 
| 41095 | 524 | unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) | 
| 40859 | 525 | qed | 
| 526 | ||
| 527 | lemma (in prob_space) joint_distribution_commute_singleton: | |
| 528 |   "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 529 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) | 
| 40859 | 530 | |
| 531 | lemma (in prob_space) joint_distribution_assoc_singleton: | |
| 532 |   "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
 | |
| 533 |    joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 534 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) | 
| 40859 | 535 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 536 | locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2 | 
| 40859 | 537 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 538 | sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 539 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 540 | sublocale pair_prob_space \<subseteq> P: prob_space P | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 541 | by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure) | 
| 40859 | 542 | |
| 543 | lemma countably_additiveI[case_names countably]: | |
| 544 | assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 545 | (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" | 
| 40859 | 546 | shows "countably_additive M \<mu>" | 
| 547 | using assms unfolding countably_additive_def by auto | |
| 548 | ||
| 549 | lemma (in prob_space) joint_distribution_prob_space: | |
| 550 | assumes "random_variable MX X" "random_variable MY Y" | |
| 43920 | 551 | shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 552 | using random_variable_pairI[OF assms] by (rule distribution_prob_space) | 
| 40859 | 553 | |
| 42988 | 554 | |
| 555 | locale finite_product_prob_space = | |
| 556 |   fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
 | |
| 557 | and I :: "'i set" | |
| 558 | assumes prob_space: "\<And>i. prob_space (M i)" and finite_index: "finite I" | |
| 559 | ||
| 560 | sublocale finite_product_prob_space \<subseteq> M: prob_space "M i" for i | |
| 561 | by (rule prob_space) | |
| 562 | ||
| 563 | sublocale finite_product_prob_space \<subseteq> finite_product_sigma_finite M I | |
| 564 | by default (rule finite_index) | |
| 565 | ||
| 566 | sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^isub>M i\<in>I. M i" | |
| 567 | proof qed (simp add: measure_times M.measure_space_1 setprod_1) | |
| 568 | ||
| 569 | lemma (in finite_product_prob_space) prob_times: | |
| 570 | assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)" | |
| 571 | shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))" | |
| 572 | proof - | |
| 43920 | 573 | have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)" | 
| 42988 | 574 | using X by (intro finite_measure_eq[symmetric] in_P) auto | 
| 575 | also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))" | |
| 576 | using measure_times X by simp | |
| 43920 | 577 | also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))" | 
| 578 | using X by (simp add: M.finite_measure_eq setprod_ereal) | |
| 42988 | 579 | finally show ?thesis by simp | 
| 580 | qed | |
| 581 | ||
| 582 | lemma (in prob_space) random_variable_restrict: | |
| 583 | assumes I: "finite I" | |
| 584 | assumes X: "\<And>i. i \<in> I \<Longrightarrow> random_variable (N i) (X i)" | |
| 585 | shows "random_variable (Pi\<^isub>M I N) (\<lambda>x. \<lambda>i\<in>I. X i x)" | |
| 586 | proof | |
| 587 |   { fix i assume "i \<in> I"
 | |
| 588 | with X interpret N: sigma_algebra "N i" by simp | |
| 589 | have "sets (N i) \<subseteq> Pow (space (N i))" by (rule N.space_closed) } | |
| 590 | note N_closed = this | |
| 591 | then show "sigma_algebra (Pi\<^isub>M I N)" | |
| 592 | by (simp add: product_algebra_def) | |
| 593 | (intro sigma_algebra_sigma product_algebra_generator_sets_into_space) | |
| 594 | show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)" | |
| 595 | using X by (intro measurable_restrict[OF I N_closed]) auto | |
| 596 | qed | |
| 597 | ||
| 40859 | 598 | section "Probability spaces on finite sets" | 
| 35582 | 599 | |
| 35977 | 600 | locale finite_prob_space = prob_space + finite_measure_space | 
| 601 | ||
| 40859 | 602 | abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'" | 
| 603 | ||
| 604 | lemma (in prob_space) finite_random_variableD: | |
| 605 | assumes "finite_random_variable M' X" shows "random_variable M' X" | |
| 606 | proof - | |
| 607 | interpret M': finite_sigma_algebra M' using assms by simp | |
| 608 | then show "random_variable M' X" using assms by simp default | |
| 609 | qed | |
| 610 | ||
| 611 | lemma (in prob_space) distribution_finite_prob_space: | |
| 612 | assumes "finite_random_variable MX X" | |
| 43920 | 613 | shows "finite_prob_space (MX\<lparr>measure := ereal \<circ> distribution X\<rparr>)" | 
| 40859 | 614 | proof - | 
| 43920 | 615 | interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>" | 
| 40859 | 616 | using assms[THEN finite_random_variableD] by (rule distribution_prob_space) | 
| 617 | interpret MX: finite_sigma_algebra MX | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 618 | using assms by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 619 | show ?thesis by default (simp_all add: MX.finite_space) | 
| 40859 | 620 | qed | 
| 621 | ||
| 622 | lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 623 | assumes "simple_function M X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 624 | shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 625 | (is "finite_random_variable ?X _") | 
| 40859 | 626 | proof (intro conjI) | 
| 627 | have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 628 | interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 629 | show "finite_sigma_algebra ?X" | 
| 40859 | 630 | by default auto | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 631 | show "X \<in> measurable M ?X" | 
| 40859 | 632 | proof (unfold measurable_def, clarsimp) | 
| 633 | fix A assume A: "A \<subseteq> X`space M" | |
| 634 | then have "finite A" by (rule finite_subset) simp | |
| 635 |     then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
 | |
| 636 | unfolding vimage_UN UN_extend_simps | |
| 637 | apply (rule finite_UN) | |
| 638 | using A assms unfolding simple_function_def by auto | |
| 639 | then show "X -` A \<inter> space M \<in> events" by simp | |
| 640 | qed | |
| 641 | qed | |
| 642 | ||
| 643 | lemma (in prob_space) simple_function_imp_random_variable[simp, intro]: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 644 | assumes "simple_function M X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 645 | shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 646 | using simple_function_imp_finite_random_variable[OF assms, of ext] | 
| 40859 | 647 | by (auto dest!: finite_random_variableD) | 
| 648 | ||
| 649 | lemma (in prob_space) sum_over_space_real_distribution: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 650 |   "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
 | 
| 40859 | 651 | unfolding distribution_def prob_space[symmetric] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 652 | by (subst finite_measure_finite_Union[symmetric]) | 
| 40859 | 653 | (auto simp add: disjoint_family_on_def simple_function_def | 
| 654 | intro!: arg_cong[where f=prob]) | |
| 655 | ||
| 656 | lemma (in prob_space) finite_random_variable_pairI: | |
| 657 | assumes "finite_random_variable MX X" | |
| 658 | assumes "finite_random_variable MY Y" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 659 | shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))" | 
| 40859 | 660 | proof | 
| 661 | interpret MX: finite_sigma_algebra MX using assms by simp | |
| 662 | interpret MY: finite_sigma_algebra MY using assms by simp | |
| 663 | interpret P: pair_finite_sigma_algebra MX MY by default | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 664 | show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default | 
| 40859 | 665 | have sa: "sigma_algebra M" by default | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 666 | show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)" | 
| 41095 | 667 | unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) | 
| 40859 | 668 | qed | 
| 669 | ||
| 670 | lemma (in prob_space) finite_random_variable_imp_sets: | |
| 671 |   "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
 | |
| 672 | unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp | |
| 673 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 674 | lemma (in prob_space) finite_random_variable_measurable: | 
| 40859 | 675 | assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events" | 
| 676 | proof - | |
| 677 | interpret X: finite_sigma_algebra MX using X by simp | |
| 678 | from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and | |
| 679 | "X \<in> space M \<rightarrow> space MX" | |
| 680 | by (auto simp: measurable_def) | |
| 681 | then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M" | |
| 682 | by auto | |
| 683 | show "X -` A \<inter> space M \<in> events" | |
| 684 | unfolding * by (intro vimage) auto | |
| 685 | qed | |
| 686 | ||
| 687 | lemma (in prob_space) joint_distribution_finite_Times_le_fst: | |
| 688 | assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" | |
| 689 | shows "joint_distribution X Y (A \<times> B) \<le> distribution X A" | |
| 690 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 691 | proof (intro finite_measure_mono) | 
| 40859 | 692 | show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force | 
| 693 | show "X -` A \<inter> space M \<in> events" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 694 | using finite_random_variable_measurable[OF X] . | 
| 40859 | 695 | have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M = | 
| 696 | (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto | |
| 697 | qed | |
| 698 | ||
| 699 | lemma (in prob_space) joint_distribution_finite_Times_le_snd: | |
| 700 | assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" | |
| 701 | shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B" | |
| 702 | using assms | |
| 703 | by (subst joint_distribution_commute) | |
| 704 | (simp add: swap_product joint_distribution_finite_Times_le_fst) | |
| 705 | ||
| 706 | lemma (in prob_space) finite_distribution_order: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 707 |   fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
 | 
| 40859 | 708 | assumes "finite_random_variable MX X" "finite_random_variable MY Y" | 
| 709 |   shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
 | |
| 710 |     and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
 | |
| 711 |     and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
 | |
| 712 |     and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
 | |
| 713 |     and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
 | |
| 714 |     and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
 | |
| 715 |   using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
 | |
| 716 |   using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 717 | by (auto intro: antisym) | 
| 40859 | 718 | |
| 719 | lemma (in prob_space) setsum_joint_distribution: | |
| 720 | assumes X: "finite_random_variable MX X" | |
| 721 | assumes Y: "random_variable MY Y" "B \<in> sets MY" | |
| 722 |   shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
 | |
| 723 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 724 | proof (subst finite_measure_finite_Union[symmetric]) | 
| 40859 | 725 | interpret MX: finite_sigma_algebra MX using X by auto | 
| 726 | show "finite (space MX)" using MX.finite_space . | |
| 727 |   let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
 | |
| 728 |   { fix i assume "i \<in> space MX"
 | |
| 729 |     moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
 | |
| 730 | ultimately show "?d i \<in> events" | |
| 731 | using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y | |
| 732 | using MX.sets_eq_Pow by auto } | |
| 733 | show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 734 | show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 735 | using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>']) | 
| 40859 | 736 | qed | 
| 737 | ||
| 738 | lemma (in prob_space) setsum_joint_distribution_singleton: | |
| 739 | assumes X: "finite_random_variable MX X" | |
| 740 | assumes Y: "finite_random_variable MY Y" "b \<in> space MY" | |
| 741 |   shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
 | |
| 742 | using setsum_joint_distribution[OF X | |
| 743 | finite_random_variableD[OF Y(1)] | |
| 744 | finite_random_variable_imp_sets[OF Y]] by simp | |
| 745 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 746 | locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 | 
| 40859 | 747 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 748 | sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 749 | sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2 by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 750 | sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default | 
| 40859 | 751 | |
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 752 | locale product_finite_prob_space = | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 753 |   fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 754 | and I :: "'i set" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 755 | assumes finite_space: "\<And>i. finite_prob_space (M i)" and finite_index: "finite I" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 756 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 757 | sublocale product_finite_prob_space \<subseteq> M!: finite_prob_space "M i" using finite_space . | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 758 | sublocale product_finite_prob_space \<subseteq> finite_product_sigma_finite M I by default (rule finite_index) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 759 | sublocale product_finite_prob_space \<subseteq> prob_space "Pi\<^isub>M I M" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 760 | proof | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 761 | show "\<mu> (space P) = 1" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 762 | using measure_times[OF M.top] M.measure_space_1 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 763 | by (simp add: setprod_1 space_product_algebra) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 764 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 765 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 766 | lemma funset_eq_UN_fun_upd_I: | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 767 | assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 768 | and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 769 | and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 770 | shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 771 | proof safe | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 772 | fix f assume f: "f \<in> F (insert a A)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 773 | show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 774 | proof (rule UN_I[of "f(a := d)"]) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 775 | show "f(a := d) \<in> F A" using *[OF f] . | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 776 | show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 777 | proof (rule image_eqI[of _ _ "f a"]) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 778 | show "f a \<in> G (f(a := d))" using **[OF f] . | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 779 | qed simp | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 780 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 781 | next | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 782 | fix f x assume "f \<in> F A" "x \<in> G f" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 783 | from ***[OF this] show "f(a := x) \<in> F (insert a A)" . | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 784 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 785 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 786 | lemma extensional_funcset_insert_eq[simp]: | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 787 | assumes "a \<notin> A" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 788 | shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 789 | apply (rule funset_eq_UN_fun_upd_I) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 790 | using assms | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 791 | by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 792 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 793 | lemma finite_extensional_funcset[simp, intro]: | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 794 | assumes "finite A" "finite B" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 795 | shows "finite (extensional A \<inter> (A \<rightarrow> B))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 796 | using assms by induct (auto simp: extensional_funcset_insert_eq) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 797 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 798 | lemma finite_PiE[simp, intro]: | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 799 | assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 800 | shows "finite (Pi\<^isub>E A B)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 801 | proof - | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 802 | have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 803 | show ?thesis | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 804 | using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 805 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 806 | |
| 42892 | 807 | lemma (in product_finite_prob_space) singleton_eq_product: | 
| 808 |   assumes x: "x \<in> space P" shows "{x} = (\<Pi>\<^isub>E i\<in>I. {x i})"
 | |
| 809 | proof (safe intro!: ext[of _ x]) | |
| 810 |   fix y i assume *: "y \<in> (\<Pi> i\<in>I. {x i})" "y \<in> extensional I"
 | |
| 811 | with x show "y i = x i" | |
| 812 | by (cases "i \<in> I") (auto simp: extensional_def) | |
| 813 | qed (insert x, auto) | |
| 814 | ||
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 815 | sublocale product_finite_prob_space \<subseteq> finite_prob_space "Pi\<^isub>M I M" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 816 | proof | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 817 | show "finite (space P)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 818 | using finite_index M.finite_space by auto | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 819 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 820 |   { fix x assume "x \<in> space P"
 | 
| 42892 | 821 | with this[THEN singleton_eq_product] | 
| 822 |     have "{x} \<in> sets P"
 | |
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 823 | by (auto intro!: in_P) } | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 824 | note x_in_P = this | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 825 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 826 | have "Pow (space P) \<subseteq> sets P" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 827 | proof | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 828 | fix X assume "X \<in> Pow (space P)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 829 | moreover then have "finite X" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 830 | using `finite (space P)` by (blast intro: finite_subset) | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 831 |     ultimately have "(\<Union>x\<in>X. {x}) \<in> sets P"
 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 832 | by (intro finite_UN x_in_P) auto | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 833 | then show "X \<in> sets P" by simp | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 834 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 835 | with space_closed show [simp]: "sets P = Pow (space P)" .. | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 836 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 837 |   { fix x assume "x \<in> space P"
 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 838 |     from this top have "\<mu> {x} \<le> \<mu> (space P)" by (intro measure_mono) auto
 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 839 |     then show "\<mu> {x} \<noteq> \<infinity>"
 | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 840 | using measure_space_1 by auto } | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 841 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 842 | |
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 843 | lemma (in product_finite_prob_space) measure_finite_times: | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 844 | "(\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)) \<Longrightarrow> \<mu> (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.\<mu> i (X i))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 845 | by (rule measure_times) simp | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 846 | |
| 42892 | 847 | lemma (in product_finite_prob_space) measure_singleton_times: | 
| 848 |   assumes x: "x \<in> space P" shows "\<mu> {x} = (\<Prod>i\<in>I. M.\<mu> i {x i})"
 | |
| 849 | unfolding singleton_eq_product[OF x] using x | |
| 850 | by (intro measure_finite_times) auto | |
| 851 | ||
| 852 | lemma (in product_finite_prob_space) prob_finite_times: | |
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 853 | assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 854 | shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 855 | proof - | 
| 43920 | 856 | have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)" | 
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 857 | using X by (intro finite_measure_eq[symmetric] in_P) auto | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 858 | also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))" | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 859 | using measure_finite_times X by simp | 
| 43920 | 860 | also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))" | 
| 861 | using X by (simp add: M.finite_measure_eq setprod_ereal) | |
| 42859 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 862 | finally show ?thesis by simp | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 863 | qed | 
| 
d9dfc733f25c
add product of probability spaces with finite cardinality
 hoelzl parents: 
42858diff
changeset | 864 | |
| 42892 | 865 | lemma (in product_finite_prob_space) prob_singleton_times: | 
| 866 | assumes x: "x \<in> space P" | |
| 867 |   shows "prob {x} = (\<Prod>i\<in>I. M.prob i {x i})"
 | |
| 868 | unfolding singleton_eq_product[OF x] using x | |
| 869 | by (intro prob_finite_times) auto | |
| 870 | ||
| 871 | lemma (in product_finite_prob_space) prob_finite_product: | |
| 872 |   "A \<subseteq> space P \<Longrightarrow> prob A = (\<Sum>x\<in>A. \<Prod>i\<in>I. M.prob i {x i})"
 | |
| 873 | by (auto simp add: finite_measure_singleton prob_singleton_times | |
| 874 | simp del: space_product_algebra | |
| 875 | intro!: setsum_cong prob_singleton_times) | |
| 876 | ||
| 40859 | 877 | lemma (in prob_space) joint_distribution_finite_prob_space: | 
| 878 | assumes X: "finite_random_variable MX X" | |
| 879 | assumes Y: "finite_random_variable MY Y" | |
| 43920 | 880 | shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 881 | by (intro distribution_finite_prob_space finite_random_variable_pairI X Y) | 
| 40859 | 882 | |
| 36624 | 883 | lemma finite_prob_space_eq: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 884 | "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1" | 
| 36624 | 885 | unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def | 
| 886 | by auto | |
| 887 | ||
| 38656 | 888 | lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
 | 
| 889 | using measure_space_1 sum_over_space by simp | |
| 36624 | 890 | |
| 891 | lemma (in finite_prob_space) joint_distribution_restriction_fst: | |
| 892 | "joint_distribution X Y A \<le> distribution X (fst ` A)" | |
| 893 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 894 | proof (safe intro!: finite_measure_mono) | 
| 36624 | 895 | fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" | 
| 896 | show "x \<in> X -` fst ` A" | |
| 897 | by (auto intro!: image_eqI[OF _ *]) | |
| 898 | qed (simp_all add: sets_eq_Pow) | |
| 899 | ||
| 900 | lemma (in finite_prob_space) joint_distribution_restriction_snd: | |
| 901 | "joint_distribution X Y A \<le> distribution Y (snd ` A)" | |
| 902 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 903 | proof (safe intro!: finite_measure_mono) | 
| 36624 | 904 | fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" | 
| 905 | show "x \<in> Y -` snd ` A" | |
| 906 | by (auto intro!: image_eqI[OF _ *]) | |
| 907 | qed (simp_all add: sets_eq_Pow) | |
| 908 | ||
| 909 | lemma (in finite_prob_space) distribution_order: | |
| 910 | shows "0 \<le> distribution X x'" | |
| 911 | and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')" | |
| 912 |   and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
 | |
| 913 |   and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
 | |
| 914 |   and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
 | |
| 915 |   and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
 | |
| 916 |   and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
 | |
| 917 |   and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 918 | using | 
| 36624 | 919 |     joint_distribution_restriction_fst[of X Y "{(x, y)}"]
 | 
| 920 |     joint_distribution_restriction_snd[of X Y "{(x, y)}"]
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 921 | by (auto intro: antisym) | 
| 36624 | 922 | |
| 39097 | 923 | lemma (in finite_prob_space) distribution_mono: | 
| 924 | assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" | |
| 925 | shows "distribution X x \<le> distribution Y y" | |
| 926 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 927 | using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono) | 
| 39097 | 928 | |
| 929 | lemma (in finite_prob_space) distribution_mono_gt_0: | |
| 930 | assumes gt_0: "0 < distribution X x" | |
| 931 | assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" | |
| 932 | shows "0 < distribution Y y" | |
| 933 | by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) | |
| 934 | ||
| 935 | lemma (in finite_prob_space) sum_over_space_distrib: | |
| 936 |   "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 937 | unfolding distribution_def prob_space[symmetric] using finite_space | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 938 | by (subst finite_measure_finite_Union[symmetric]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 939 | (auto simp add: disjoint_family_on_def sets_eq_Pow | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 940 | intro!: arg_cong[where f=\<mu>']) | 
| 39097 | 941 | |
| 942 | lemma (in finite_prob_space) finite_sum_over_space_eq_1: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 943 |   "(\<Sum>x\<in>space M. prob {x}) = 1"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 944 | using prob_space finite_space | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 945 | by (subst (asm) finite_measure_finite_singleton) auto | 
| 39097 | 946 | |
| 947 | lemma (in prob_space) distribution_remove_const: | |
| 948 |   shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
 | |
| 949 |   and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
 | |
| 950 |   and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
 | |
| 951 |   and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
 | |
| 952 |   and "distribution (\<lambda>x. ()) {()} = 1"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 953 | by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric]) | 
| 35977 | 954 | |
| 39097 | 955 | lemma (in finite_prob_space) setsum_distribution_gen: | 
| 956 |   assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
 | |
| 957 | and "inj_on f (X`space M)" | |
| 958 |   shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
 | |
| 959 | unfolding distribution_def assms | |
| 960 | using finite_space assms | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 961 | by (subst finite_measure_finite_Union[symmetric]) | 
| 39097 | 962 | (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def | 
| 963 | intro!: arg_cong[where f=prob]) | |
| 964 | ||
| 965 | lemma (in finite_prob_space) setsum_distribution: | |
| 966 |   "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
 | |
| 967 |   "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
 | |
| 968 |   "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
 | |
| 969 |   "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
 | |
| 970 |   "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
 | |
| 971 | by (auto intro!: inj_onI setsum_distribution_gen) | |
| 972 | ||
| 973 | lemma (in finite_prob_space) uniform_prob: | |
| 974 | assumes "x \<in> space M" | |
| 975 |   assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 976 |   shows "prob {x} = 1 / card (space M)"
 | 
| 39097 | 977 | proof - | 
| 978 |   have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
 | |
| 979 | using assms(2)[OF _ `x \<in> space M`] by blast | |
| 980 | have "1 = prob (space M)" | |
| 981 | using prob_space by auto | |
| 982 |   also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 983 |     using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
 | 
| 39097 | 984 | sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] | 
| 985 | finite_space unfolding disjoint_family_on_def prob_space[symmetric] | |
| 986 | by (auto simp add:setsum_restrict_set) | |
| 987 |   also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
 | |
| 988 | using prob_x by auto | |
| 989 |   also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
 | |
| 990 |   finally have one: "1 = real (card (space M)) * prob {x}"
 | |
| 991 | using real_eq_of_nat by auto | |
| 43339 | 992 | hence two: "real (card (space M)) \<noteq> 0" by fastsimp | 
| 39097 | 993 |   from one have three: "prob {x} \<noteq> 0" by fastsimp
 | 
| 994 | thus ?thesis using one two three divide_cancel_right | |
| 995 | by (auto simp:field_simps) | |
| 39092 | 996 | qed | 
| 35977 | 997 | |
| 39092 | 998 | lemma (in prob_space) prob_space_subalgebra: | 
| 41545 | 999 | assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1000 | and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1001 | shows "prob_space N" | 
| 39092 | 1002 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1003 | interpret N: measure_space N | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1004 | by (rule measure_space_subalgebra[OF assms]) | 
| 39092 | 1005 | show ?thesis | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1006 | proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1) | 
| 35977 | 1007 | qed | 
| 1008 | ||
| 39092 | 1009 | lemma (in prob_space) prob_space_of_restricted_space: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1010 | assumes "\<mu> A \<noteq> 0" "A \<in> sets M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1011 | shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1012 | (is "prob_space ?P") | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1013 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1014 | interpret A: measure_space "restricted_space A" | 
| 39092 | 1015 | using `A \<in> sets M` by (rule restricted_measure_space) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1016 | interpret A': sigma_algebra ?P | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1017 | by (rule A.sigma_algebra_cong) auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1018 | show "prob_space ?P" | 
| 39092 | 1019 | proof | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1020 | show "measure ?P (space ?P) = 1" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1021 | using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1022 | show "positive ?P (measure ?P)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1023 | proof (simp add: positive_def, safe) | 
| 43920 | 1024 | show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_ereal_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1025 | fix B assume "B \<in> events" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1026 | with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M` | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1027 | show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1028 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1029 | show "countably_additive ?P (measure ?P)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1030 | proof (simp add: countably_additive_def, safe) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1031 | fix B and F :: "nat \<Rightarrow> 'a set" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1032 | assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1033 |       { fix i
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1034 | from F have "F i \<in> op \<inter> A ` events" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1035 | with `A \<in> events` have "F i \<in> events" by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1036 | moreover then have "range F \<subseteq> events" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1037 | moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" | 
| 43920 | 1038 | by (simp add: mult_commute divide_ereal_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1039 | moreover have "0 \<le> inverse (\<mu> A)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1040 | using real_measure[OF `A \<in> events`] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1041 | ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1042 | using measure_countably_additive[of F] F | 
| 43920 | 1043 | by (auto simp: suminf_cmult_ereal) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1044 | qed | 
| 39092 | 1045 | qed | 
| 1046 | qed | |
| 1047 | ||
| 1048 | lemma finite_prob_spaceI: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1049 | assumes "finite (space M)" "sets M = Pow(space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1050 |     and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1051 |     and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1052 | shows "finite_prob_space M" | 
| 39092 | 1053 | unfolding finite_prob_space_eq | 
| 1054 | proof | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1055 | show "finite_measure_space M" using assms | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1056 | by (auto intro!: finite_measure_spaceI) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1057 | show "measure M (space M) = 1" by fact | 
| 39092 | 1058 | qed | 
| 36624 | 1059 | |
| 1060 | lemma (in finite_prob_space) finite_measure_space: | |
| 39097 | 1061 | fixes X :: "'a \<Rightarrow> 'x" | 
| 43920 | 1062 | shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1063 | (is "finite_measure_space ?S") | 
| 39092 | 1064 | proof (rule finite_measure_spaceI, simp_all) | 
| 36624 | 1065 | show "finite (X ` space M)" using finite_space by simp | 
| 39097 | 1066 | next | 
| 1067 |   fix A B :: "'x set" assume "A \<inter> B = {}"
 | |
| 1068 | then show "distribution X (A \<union> B) = distribution X A + distribution X B" | |
| 1069 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1070 | by (subst finite_measure_Union[symmetric]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1071 | (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow) | 
| 36624 | 1072 | qed | 
| 1073 | ||
| 39097 | 1074 | lemma (in finite_prob_space) finite_prob_space_of_images: | 
| 43920 | 1075 | "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X \<rparr>" | 
| 1076 | by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def) | |
| 39097 | 1077 | |
| 39096 | 1078 | lemma (in finite_prob_space) finite_product_measure_space: | 
| 39097 | 1079 | fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y" | 
| 39096 | 1080 | assumes "finite s1" "finite s2" | 
| 43920 | 1081 | shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = ereal \<circ> joint_distribution X Y\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1082 | (is "finite_measure_space ?M") | 
| 39097 | 1083 | proof (rule finite_measure_spaceI, simp_all) | 
| 1084 | show "finite (s1 \<times> s2)" | |
| 39096 | 1085 | using assms by auto | 
| 39097 | 1086 | next | 
| 1087 |   fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
 | |
| 1088 | then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B" | |
| 1089 | unfolding distribution_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1090 | by (subst finite_measure_Union[symmetric]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1091 | (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow) | 
| 39096 | 1092 | qed | 
| 1093 | ||
| 39097 | 1094 | lemma (in finite_prob_space) finite_product_measure_space_of_images: | 
| 39096 | 1095 | shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M, | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1096 | sets = Pow (X ` space M \<times> Y ` space M), | 
| 43920 | 1097 | measure = ereal \<circ> joint_distribution X Y \<rparr>" | 
| 39096 | 1098 | using finite_space by (auto intro!: finite_product_measure_space) | 
| 1099 | ||
| 40859 | 1100 | lemma (in finite_prob_space) finite_product_prob_space_of_images: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1101 | "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M), | 
| 43920 | 1102 | measure = ereal \<circ> joint_distribution X Y \<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1103 | (is "finite_prob_space ?S") | 
| 43920 | 1104 | proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def) | 
| 40859 | 1105 | have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto | 
| 1106 | thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1" | |
| 1107 | by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) | |
| 1108 | qed | |
| 1109 | ||
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1110 | subsection "Borel Measure on {0 ..< 1}"
 | 
| 42902 | 1111 | |
| 1112 | definition pborel :: "real measure_space" where | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1113 |   "pborel = lborel.restricted_space {0 ..< 1}"
 | 
| 42902 | 1114 | |
| 1115 | lemma space_pborel[simp]: | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1116 |   "space pborel = {0 ..< 1}"
 | 
| 42902 | 1117 | unfolding pborel_def by auto | 
| 1118 | ||
| 1119 | lemma sets_pborel: | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1120 |   "A \<in> sets pborel \<longleftrightarrow> A \<in> sets borel \<and> A \<subseteq> { 0 ..< 1}"
 | 
| 42902 | 1121 | unfolding pborel_def by auto | 
| 1122 | ||
| 1123 | lemma in_pborel[intro, simp]: | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1124 |   "A \<subseteq> {0 ..< 1} \<Longrightarrow> A \<in> sets borel \<Longrightarrow> A \<in> sets pborel"
 | 
| 42902 | 1125 | unfolding pborel_def by auto | 
| 1126 | ||
| 1127 | interpretation pborel: measure_space pborel | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1128 |   using lborel.restricted_measure_space[of "{0 ..< 1}"]
 | 
| 42902 | 1129 | by (simp add: pborel_def) | 
| 1130 | ||
| 1131 | interpretation pborel: prob_space pborel | |
| 43920 | 1132 | by default (simp add: one_ereal_def pborel_def) | 
| 42902 | 1133 | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1134 | lemma pborel_prob: "pborel.prob A = (if A \<in> sets borel \<and> A \<subseteq> {0 ..< 1} then real (lborel.\<mu> A) else 0)"
 | 
| 42902 | 1135 | unfolding pborel.\<mu>'_def by (auto simp: pborel_def) | 
| 1136 | ||
| 1137 | lemma pborel_singelton[simp]: "pborel.prob {a} = 0"
 | |
| 1138 | by (auto simp: pborel_prob) | |
| 1139 | ||
| 1140 | lemma | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1141 |   shows pborel_atLeastAtMost[simp]: "pborel.\<mu>' {a .. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
 | 
| 42902 | 1142 |     and pborel_atLeastLessThan[simp]: "pborel.\<mu>' {a ..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
 | 
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1143 |     and pborel_greaterThanAtMost[simp]: "pborel.\<mu>' {a <.. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
 | 
| 42902 | 1144 |     and pborel_greaterThanLessThan[simp]: "pborel.\<mu>' {a <..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
 | 
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1145 | unfolding pborel_prob | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1146 | by (auto simp: atLeastAtMost_subseteq_atLeastLessThan_iff | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1147 | greaterThanAtMost_subseteq_atLeastLessThan_iff greaterThanLessThan_subseteq_atLeastLessThan_iff) | 
| 42902 | 1148 | |
| 1149 | lemma pborel_lebesgue_measure: | |
| 1150 | "A \<in> sets pborel \<Longrightarrow> pborel.prob A = real (measure lebesgue A)" | |
| 1151 | by (simp add: sets_pborel pborel_prob) | |
| 1152 | ||
| 1153 | lemma pborel_alt: | |
| 1154 | "pborel = sigma \<lparr> | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1155 |     space = {0..<1},
 | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1156 |     sets = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1}),
 | 
| 42902 | 1157 | measure = measure lborel \<rparr>" (is "_ = ?R") | 
| 1158 | proof - | |
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1159 |   have *: "{0..<1::real} \<in> sets borel" by auto
 | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1160 |   have **: "op \<inter> {0..<1::real} ` range (\<lambda>(x, y). {x..<y}) = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1})"
 | 
| 42902 | 1161 | unfolding image_image by (intro arg_cong[where f=range]) auto | 
| 43658 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1162 |   have "pborel = algebra.restricted_space (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real}),
 | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1163 |     measure = measure pborel\<rparr>) {0 ..< 1}"
 | 
| 
0d96ec6ec33b
the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
 hoelzl parents: 
43556diff
changeset | 1164 | by (simp add: sigma_def lebesgue_def pborel_def borel_eq_atLeastLessThan lborel_def) | 
| 42902 | 1165 | also have "\<dots> = ?R" | 
| 1166 | by (subst restricted_sigma) | |
| 1167 | (simp_all add: sets_sigma sigma_sets.Basic ** pborel_def image_eqI[of _ _ "(0,1)"]) | |
| 1168 | finally show ?thesis . | |
| 1169 | qed | |
| 1170 | ||
| 42860 | 1171 | subsection "Bernoulli space" | 
| 1172 | ||
| 1173 | definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV, | |
| 43920 | 1174 | measure = ereal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>" | 
| 42860 | 1175 | |
| 1176 | interpretation bernoulli: finite_prob_space "bernoulli_space p" for p | |
| 1177 | by (rule finite_prob_spaceI) | |
| 43920 | 1178 | (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg) | 
| 42860 | 1179 | |
| 1180 | lemma bernoulli_measure: | |
| 1181 | "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)" | |
| 1182 | unfolding bernoulli.\<mu>'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong) | |
| 1183 | ||
| 1184 | lemma bernoulli_measure_True: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {True} = p"
 | |
| 1185 |   and bernoulli_measure_False: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {False} = 1 - p"
 | |
| 1186 | unfolding bernoulli_measure by simp_all | |
| 1187 | ||
| 35582 | 1188 | end |