| author | wenzelm |
| Fri, 20 May 2011 20:44:03 +0200 | |
| changeset 42897 | 6bc8a6dcb3e0 |
| parent 42892 | a61e30bfd0bc |
| child 42902 | e8dbf90a2f3b |
| 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 |
|
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
9 |
imports Lebesgue_Integration Radon_Nikodym Finite_Product_Measure |
| 35582 | 10 |
begin |
11 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
12 |
lemma real_of_extreal_inverse[simp]: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
13 |
fixes X :: extreal |
| 40859 | 14 |
shows "real (inverse X) = 1 / real X" |
15 |
by (cases X) (auto simp: inverse_eq_divide) |
|
16 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
17 |
lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
18 |
by (cases X) auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
19 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
20 |
lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>" |
| 40859 | 21 |
by (cases X) auto |
22 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
23 |
lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)" |
| 40859 | 24 |
by (cases X) auto |
25 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
26 |
lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
27 |
by (cases X) (auto simp: one_extreal_def) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
28 |
|
| 35582 | 29 |
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:
41661
diff
changeset
|
30 |
assumes measure_space_1: "measure M (space M) = 1" |
| 38656 | 31 |
|
32 |
sublocale prob_space < finite_measure |
|
33 |
proof |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
34 |
from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp |
| 38656 | 35 |
qed |
36 |
||
| 40859 | 37 |
abbreviation (in prob_space) "events \<equiv> sets M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
38 |
abbreviation (in prob_space) "prob \<equiv> \<mu>'" |
| 40859 | 39 |
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:
41661
diff
changeset
|
40 |
abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M" |
| 35582 | 41 |
|
| 40859 | 42 |
definition (in prob_space) |
| 35582 | 43 |
"indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B" |
44 |
||
| 40859 | 45 |
definition (in prob_space) |
| 35582 | 46 |
"indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)" |
47 |
||
| 40859 | 48 |
definition (in prob_space) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
49 |
"distribution X A = \<mu>' (X -` A \<inter> space M)" |
| 35582 | 50 |
|
| 40859 | 51 |
abbreviation (in prob_space) |
| 36624 | 52 |
"joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))" |
| 35582 | 53 |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
54 |
declare (in finite_measure) positive_measure'[intro, simp] |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
55 |
|
| 39097 | 56 |
lemma (in prob_space) distribution_cong: |
57 |
assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x" |
|
58 |
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:
39198
diff
changeset
|
59 |
unfolding distribution_def fun_eq_iff |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
60 |
using assms by (auto intro!: arg_cong[where f="\<mu>'"]) |
| 39097 | 61 |
|
62 |
lemma (in prob_space) joint_distribution_cong: |
|
63 |
assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" |
|
64 |
assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" |
|
65 |
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:
39198
diff
changeset
|
66 |
unfolding distribution_def fun_eq_iff |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
67 |
using assms by (auto intro!: arg_cong[where f="\<mu>'"]) |
| 39097 | 68 |
|
| 40859 | 69 |
lemma (in prob_space) distribution_id[simp]: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
70 |
"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:
41831
diff
changeset
|
71 |
by (auto simp: distribution_def intro!: arg_cong[where f=prob]) |
| 40859 | 72 |
|
73 |
lemma (in prob_space) prob_space: "prob (space M) = 1" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
74 |
using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
75 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
76 |
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:
41831
diff
changeset
|
77 |
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:
41831
diff
changeset
|
78 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
79 |
lemma (in prob_space) distribution_positive[simp, intro]: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
80 |
"0 \<le> distribution X A" unfolding distribution_def by auto |
| 35582 | 81 |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
82 |
lemma (in prob_space) joint_distribution_remove[simp]: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
83 |
"joint_distribution X X {(x, x)} = distribution X {x}"
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
84 |
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:
41831
diff
changeset
|
85 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
86 |
lemma (in prob_space) distribution_1: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
87 |
"distribution X A \<le> 1" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
88 |
unfolding distribution_def by simp |
| 35582 | 89 |
|
| 40859 | 90 |
lemma (in prob_space) prob_compl: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
91 |
assumes A: "A \<in> events" |
| 38656 | 92 |
shows "prob (space M - A) = 1 - prob A" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
93 |
using finite_measure_compl[OF A] by (simp add: prob_space) |
| 35582 | 94 |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
95 |
lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
96 |
by (simp add: indep_def 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:
41831
diff
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:
41831
diff
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:
41831
diff
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:
41831
diff
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:
41831
diff
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:
41831
diff
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:
41831
diff
changeset
|
154 |
qed simp |
| 35582 | 155 |
|
| 40859 | 156 |
lemma (in prob_space) indep_sym: |
| 35582 | 157 |
"indep a b \<Longrightarrow> indep b a" |
158 |
unfolding indep_def using Int_commute[of a b] by auto |
|
159 |
||
| 40859 | 160 |
lemma (in prob_space) indep_refl: |
| 35582 | 161 |
assumes "a \<in> events" |
162 |
shows "indep a a = (prob a = 0) \<or> (prob a = 1)" |
|
163 |
using assms unfolding indep_def by auto |
|
164 |
||
| 40859 | 165 |
lemma (in prob_space) prob_equiprobable_finite_unions: |
| 38656 | 166 |
assumes "s \<in> events" |
167 |
assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
|
|
| 35582 | 168 |
assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
|
| 38656 | 169 |
shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
|
| 35582 | 170 |
proof (cases "s = {}")
|
| 38656 | 171 |
case False hence "\<exists> x. x \<in> s" by blast |
| 35582 | 172 |
from someI_ex[OF this] assms |
173 |
have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
|
|
174 |
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:
41831
diff
changeset
|
175 |
using finite_measure_finite_singleton[OF s_finite] by simp |
| 35582 | 176 |
also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
|
| 38656 | 177 |
also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
|
178 |
using setsum_constant assms by (simp add: real_eq_of_nat) |
|
| 35582 | 179 |
finally show ?thesis by simp |
| 38656 | 180 |
qed simp |
| 35582 | 181 |
|
| 40859 | 182 |
lemma (in prob_space) prob_real_sum_image_fn: |
| 35582 | 183 |
assumes "e \<in> events" |
184 |
assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events" |
|
185 |
assumes "finite s" |
|
| 38656 | 186 |
assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
|
187 |
assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)" |
|
| 35582 | 188 |
shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))" |
189 |
proof - |
|
| 38656 | 190 |
have e: "e = (\<Union> i \<in> s. e \<inter> f i)" |
191 |
using `e \<in> events` sets_into_space upper by blast |
|
192 |
hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp |
|
193 |
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:
41831
diff
changeset
|
194 |
proof (rule finite_measure_finite_Union) |
| 38656 | 195 |
show "finite s" by fact |
196 |
show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact |
|
197 |
show "disjoint_family_on (\<lambda>i. e \<inter> f i) s" |
|
198 |
using disjoint by (auto simp: disjoint_family_on_def) |
|
199 |
qed |
|
200 |
finally show ?thesis . |
|
| 35582 | 201 |
qed |
202 |
||
| 42199 | 203 |
lemma (in prob_space) prob_space_vimage: |
204 |
assumes S: "sigma_algebra S" |
|
205 |
assumes T: "T \<in> measure_preserving M S" |
|
206 |
shows "prob_space S" |
|
| 35582 | 207 |
proof - |
| 42199 | 208 |
interpret S: measure_space S |
209 |
using S and T by (rule measure_space_vimage) |
|
| 38656 | 210 |
show ?thesis |
| 42199 | 211 |
proof |
212 |
from T[THEN measure_preservingD2] |
|
213 |
have "T -` space S \<inter> space M = space M" |
|
214 |
by (auto simp: measurable_def) |
|
215 |
with T[THEN measure_preservingD, of "space S", symmetric] |
|
216 |
show "measure S (space S) = 1" |
|
217 |
using measure_space_1 by simp |
|
| 35582 | 218 |
qed |
219 |
qed |
|
220 |
||
| 42199 | 221 |
lemma (in prob_space) distribution_prob_space: |
222 |
assumes X: "random_variable S X" |
|
223 |
shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)" (is "prob_space ?S") |
|
224 |
proof (rule prob_space_vimage) |
|
225 |
show "X \<in> measure_preserving M ?S" |
|
226 |
using X |
|
227 |
unfolding measure_preserving_def distribution_def_raw |
|
228 |
by (auto simp: finite_measure_eq measurable_sets) |
|
229 |
show "sigma_algebra ?S" using X by simp |
|
230 |
qed |
|
231 |
||
| 40859 | 232 |
lemma (in prob_space) AE_distribution: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
233 |
assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x" |
| 40859 | 234 |
shows "AE x. Q (X x)" |
235 |
proof - |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
236 |
interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space) |
| 40859 | 237 |
obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
|
238 |
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:
41831
diff
changeset
|
239 |
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:
41831
diff
changeset
|
240 |
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:
41831
diff
changeset
|
241 |
(auto simp: finite_measure_eq distribution_def measurable_sets) |
| 40859 | 242 |
qed |
243 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
244 |
lemma (in prob_space) distribution_eq_integral: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
245 |
"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:
41831
diff
changeset
|
246 |
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:
41831
diff
changeset
|
247 |
by (auto simp: measurable_sets distribution_def) |
| 35582 | 248 |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
249 |
lemma (in prob_space) distribution_eq_translated_integral: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
250 |
assumes "random_variable S X" "A \<in> sets S" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
251 |
shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)" |
| 35582 | 252 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
253 |
interpret S: prob_space "S\<lparr>measure := extreal \<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:
41661
diff
changeset
|
254 |
using assms(1) by (rule distribution_prob_space) |
| 35582 | 255 |
show ?thesis |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
256 |
using S.positive_integral_indicator(1)[of A] assms by simp |
| 35582 | 257 |
qed |
258 |
||
| 40859 | 259 |
lemma (in prob_space) finite_expectation1: |
260 |
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:
41831
diff
changeset
|
261 |
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:
41831
diff
changeset
|
262 |
proof (subst integral_on_finite) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
263 |
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:
41831
diff
changeset
|
264 |
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:
41831
diff
changeset
|
265 |
"\<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:
41831
diff
changeset
|
266 |
using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto |
| 38656 | 267 |
qed |
| 35582 | 268 |
|
| 40859 | 269 |
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:
41661
diff
changeset
|
270 |
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:
41831
diff
changeset
|
271 |
shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
|
| 38656 | 272 |
using assms unfolding distribution_def using finite_expectation1 by auto |
273 |
||
| 40859 | 274 |
lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0: |
| 35582 | 275 |
assumes "{x} \<in> events"
|
| 38656 | 276 |
assumes "prob {x} = 1"
|
| 35582 | 277 |
assumes "{y} \<in> events"
|
278 |
assumes "y \<noteq> x" |
|
279 |
shows "prob {y} = 0"
|
|
280 |
using prob_one_inter[of "{y}" "{x}"] assms by auto
|
|
281 |
||
| 40859 | 282 |
lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
|
| 38656 | 283 |
unfolding distribution_def by simp |
284 |
||
| 40859 | 285 |
lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1" |
| 38656 | 286 |
proof - |
287 |
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:
41831
diff
changeset
|
288 |
thus ?thesis unfolding distribution_def by (simp add: prob_space) |
| 38656 | 289 |
qed |
290 |
||
| 40859 | 291 |
lemma (in prob_space) distribution_one: |
292 |
assumes "random_variable M' X" and "A \<in> sets M'" |
|
| 38656 | 293 |
shows "distribution X A \<le> 1" |
294 |
proof - |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
295 |
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:
41831
diff
changeset
|
296 |
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:
41831
diff
changeset
|
297 |
thus ?thesis by (simp add: prob_space) |
| 38656 | 298 |
qed |
299 |
||
| 40859 | 300 |
lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0: |
| 35582 | 301 |
assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X" |
| 38656 | 302 |
(is "random_variable ?S X") |
303 |
assumes "distribution X {x} = 1"
|
|
| 35582 | 304 |
assumes "y \<noteq> x" |
305 |
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:
41661
diff
changeset
|
306 |
proof cases |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
307 |
{ 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:
41661
diff
changeset
|
308 |
proof cases |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
309 |
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:
41661
diff
changeset
|
310 |
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:
41661
diff
changeset
|
311 |
next |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
312 |
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:
41661
diff
changeset
|
313 |
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:
41661
diff
changeset
|
314 |
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:
41661
diff
changeset
|
315 |
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:
41661
diff
changeset
|
316 |
"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:
41661
diff
changeset
|
317 |
using `y \<noteq> x` by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
318 |
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:
41831
diff
changeset
|
319 |
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:
41661
diff
changeset
|
320 |
next |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
321 |
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:
41661
diff
changeset
|
322 |
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:
41661
diff
changeset
|
323 |
thus "distribution X {y} = 0" unfolding distribution_def by auto
|
| 35582 | 324 |
qed |
325 |
||
| 40859 | 326 |
lemma (in prob_space) joint_distribution_Times_le_fst: |
327 |
assumes X: "random_variable MX X" and Y: "random_variable MY Y" |
|
328 |
and A: "A \<in> sets MX" and B: "B \<in> sets MY" |
|
329 |
shows "joint_distribution X Y (A \<times> B) \<le> distribution X A" |
|
330 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
331 |
proof (intro finite_measure_mono) |
| 40859 | 332 |
show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force |
333 |
show "X -` A \<inter> space M \<in> events" |
|
334 |
using X A unfolding measurable_def by simp |
|
335 |
have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M = |
|
336 |
(X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto |
|
337 |
qed |
|
338 |
||
339 |
lemma (in prob_space) joint_distribution_commute: |
|
340 |
"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:
41831
diff
changeset
|
341 |
unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) |
| 40859 | 342 |
|
343 |
lemma (in prob_space) joint_distribution_Times_le_snd: |
|
344 |
assumes X: "random_variable MX X" and Y: "random_variable MY Y" |
|
345 |
and A: "A \<in> sets MX" and B: "B \<in> sets MY" |
|
346 |
shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B" |
|
347 |
using assms |
|
348 |
by (subst joint_distribution_commute) |
|
349 |
(simp add: swap_product joint_distribution_Times_le_fst) |
|
350 |
||
351 |
lemma (in prob_space) random_variable_pairI: |
|
352 |
assumes "random_variable MX X" |
|
353 |
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:
41661
diff
changeset
|
354 |
shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))" |
| 40859 | 355 |
proof |
356 |
interpret MX: sigma_algebra MX using assms by simp |
|
357 |
interpret MY: sigma_algebra MY using assms by simp |
|
358 |
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:
41661
diff
changeset
|
359 |
show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default |
| 40859 | 360 |
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:
41661
diff
changeset
|
361 |
show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)" |
| 41095 | 362 |
unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) |
| 40859 | 363 |
qed |
364 |
||
365 |
lemma (in prob_space) joint_distribution_commute_singleton: |
|
366 |
"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:
41831
diff
changeset
|
367 |
unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) |
| 40859 | 368 |
|
369 |
lemma (in prob_space) joint_distribution_assoc_singleton: |
|
370 |
"joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
|
|
371 |
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:
41831
diff
changeset
|
372 |
unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) |
| 40859 | 373 |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
374 |
locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2 |
| 40859 | 375 |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
376 |
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:
41661
diff
changeset
|
377 |
|
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
378 |
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:
41661
diff
changeset
|
379 |
by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure) |
| 40859 | 380 |
|
381 |
lemma countably_additiveI[case_names countably]: |
|
382 |
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:
41831
diff
changeset
|
383 |
(\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" |
| 40859 | 384 |
shows "countably_additive M \<mu>" |
385 |
using assms unfolding countably_additive_def by auto |
|
386 |
||
387 |
lemma (in prob_space) joint_distribution_prob_space: |
|
388 |
assumes "random_variable MX X" "random_variable MY Y" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
389 |
shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<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:
41661
diff
changeset
|
390 |
using random_variable_pairI[OF assms] by (rule distribution_prob_space) |
| 40859 | 391 |
|
392 |
section "Probability spaces on finite sets" |
|
| 35582 | 393 |
|
| 35977 | 394 |
locale finite_prob_space = prob_space + finite_measure_space |
395 |
||
| 40859 | 396 |
abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'" |
397 |
||
398 |
lemma (in prob_space) finite_random_variableD: |
|
399 |
assumes "finite_random_variable M' X" shows "random_variable M' X" |
|
400 |
proof - |
|
401 |
interpret M': finite_sigma_algebra M' using assms by simp |
|
402 |
then show "random_variable M' X" using assms by simp default |
|
403 |
qed |
|
404 |
||
405 |
lemma (in prob_space) distribution_finite_prob_space: |
|
406 |
assumes "finite_random_variable MX X" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
407 |
shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)" |
| 40859 | 408 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
409 |
interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" |
| 40859 | 410 |
using assms[THEN finite_random_variableD] by (rule distribution_prob_space) |
411 |
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:
41661
diff
changeset
|
412 |
using assms by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
413 |
show ?thesis by default (simp_all add: MX.finite_space) |
| 40859 | 414 |
qed |
415 |
||
416 |
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:
41661
diff
changeset
|
417 |
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:
41661
diff
changeset
|
418 |
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:
41661
diff
changeset
|
419 |
(is "finite_random_variable ?X _") |
| 40859 | 420 |
proof (intro conjI) |
421 |
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:
41661
diff
changeset
|
422 |
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:
41661
diff
changeset
|
423 |
show "finite_sigma_algebra ?X" |
| 40859 | 424 |
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:
41661
diff
changeset
|
425 |
show "X \<in> measurable M ?X" |
| 40859 | 426 |
proof (unfold measurable_def, clarsimp) |
427 |
fix A assume A: "A \<subseteq> X`space M" |
|
428 |
then have "finite A" by (rule finite_subset) simp |
|
429 |
then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
|
|
430 |
unfolding vimage_UN UN_extend_simps |
|
431 |
apply (rule finite_UN) |
|
432 |
using A assms unfolding simple_function_def by auto |
|
433 |
then show "X -` A \<inter> space M \<in> events" by simp |
|
434 |
qed |
|
435 |
qed |
|
436 |
||
437 |
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:
41661
diff
changeset
|
438 |
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:
41661
diff
changeset
|
439 |
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:
41661
diff
changeset
|
440 |
using simple_function_imp_finite_random_variable[OF assms, of ext] |
| 40859 | 441 |
by (auto dest!: finite_random_variableD) |
442 |
||
443 |
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:
41831
diff
changeset
|
444 |
"simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
|
| 40859 | 445 |
unfolding distribution_def prob_space[symmetric] |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
446 |
by (subst finite_measure_finite_Union[symmetric]) |
| 40859 | 447 |
(auto simp add: disjoint_family_on_def simple_function_def |
448 |
intro!: arg_cong[where f=prob]) |
|
449 |
||
450 |
lemma (in prob_space) finite_random_variable_pairI: |
|
451 |
assumes "finite_random_variable MX X" |
|
452 |
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:
41661
diff
changeset
|
453 |
shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))" |
| 40859 | 454 |
proof |
455 |
interpret MX: finite_sigma_algebra MX using assms by simp |
|
456 |
interpret MY: finite_sigma_algebra MY using assms by simp |
|
457 |
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:
41661
diff
changeset
|
458 |
show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default |
| 40859 | 459 |
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:
41661
diff
changeset
|
460 |
show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)" |
| 41095 | 461 |
unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) |
| 40859 | 462 |
qed |
463 |
||
464 |
lemma (in prob_space) finite_random_variable_imp_sets: |
|
465 |
"finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
|
|
466 |
unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp |
|
467 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
468 |
lemma (in prob_space) finite_random_variable_measurable: |
| 40859 | 469 |
assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events" |
470 |
proof - |
|
471 |
interpret X: finite_sigma_algebra MX using X by simp |
|
472 |
from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and |
|
473 |
"X \<in> space M \<rightarrow> space MX" |
|
474 |
by (auto simp: measurable_def) |
|
475 |
then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M" |
|
476 |
by auto |
|
477 |
show "X -` A \<inter> space M \<in> events" |
|
478 |
unfolding * by (intro vimage) auto |
|
479 |
qed |
|
480 |
||
481 |
lemma (in prob_space) joint_distribution_finite_Times_le_fst: |
|
482 |
assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" |
|
483 |
shows "joint_distribution X Y (A \<times> B) \<le> distribution X A" |
|
484 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
485 |
proof (intro finite_measure_mono) |
| 40859 | 486 |
show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force |
487 |
show "X -` A \<inter> space M \<in> events" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
488 |
using finite_random_variable_measurable[OF X] . |
| 40859 | 489 |
have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M = |
490 |
(X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto |
|
491 |
qed |
|
492 |
||
493 |
lemma (in prob_space) joint_distribution_finite_Times_le_snd: |
|
494 |
assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" |
|
495 |
shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B" |
|
496 |
using assms |
|
497 |
by (subst joint_distribution_commute) |
|
498 |
(simp add: swap_product joint_distribution_finite_Times_le_fst) |
|
499 |
||
500 |
lemma (in prob_space) finite_distribution_order: |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
501 |
fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
|
| 40859 | 502 |
assumes "finite_random_variable MX X" "finite_random_variable MY Y" |
503 |
shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
|
|
504 |
and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
|
|
505 |
and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
|
|
506 |
and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
|
|
507 |
and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
|
|
508 |
and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
|
|
509 |
using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
|
|
510 |
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:
41831
diff
changeset
|
511 |
by (auto intro: antisym) |
| 40859 | 512 |
|
513 |
lemma (in prob_space) setsum_joint_distribution: |
|
514 |
assumes X: "finite_random_variable MX X" |
|
515 |
assumes Y: "random_variable MY Y" "B \<in> sets MY" |
|
516 |
shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
|
|
517 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
518 |
proof (subst finite_measure_finite_Union[symmetric]) |
| 40859 | 519 |
interpret MX: finite_sigma_algebra MX using X by auto |
520 |
show "finite (space MX)" using MX.finite_space . |
|
521 |
let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
|
|
522 |
{ fix i assume "i \<in> space MX"
|
|
523 |
moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
|
|
524 |
ultimately show "?d i \<in> events" |
|
525 |
using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y |
|
526 |
using MX.sets_eq_Pow by auto } |
|
527 |
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:
41831
diff
changeset
|
528 |
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:
41831
diff
changeset
|
529 |
using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>']) |
| 40859 | 530 |
qed |
531 |
||
532 |
lemma (in prob_space) setsum_joint_distribution_singleton: |
|
533 |
assumes X: "finite_random_variable MX X" |
|
534 |
assumes Y: "finite_random_variable MY Y" "b \<in> space MY" |
|
535 |
shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
|
|
536 |
using setsum_joint_distribution[OF X |
|
537 |
finite_random_variableD[OF Y(1)] |
|
538 |
finite_random_variable_imp_sets[OF Y]] by simp |
|
539 |
||
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
540 |
locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 |
| 40859 | 541 |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
542 |
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:
41661
diff
changeset
|
543 |
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:
41661
diff
changeset
|
544 |
sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default |
| 40859 | 545 |
|
|
42859
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
546 |
locale product_finite_prob_space = |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
547 |
fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
548 |
and I :: "'i set" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
549 |
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:
42858
diff
changeset
|
550 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
551 |
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:
42858
diff
changeset
|
552 |
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:
42858
diff
changeset
|
553 |
sublocale product_finite_prob_space \<subseteq> prob_space "Pi\<^isub>M I M" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
554 |
proof |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
555 |
show "\<mu> (space P) = 1" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
556 |
using measure_times[OF M.top] M.measure_space_1 |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
557 |
by (simp add: setprod_1 space_product_algebra) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
558 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
559 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
560 |
lemma funset_eq_UN_fun_upd_I: |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
561 |
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:
42858
diff
changeset
|
562 |
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:
42858
diff
changeset
|
563 |
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:
42858
diff
changeset
|
564 |
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:
42858
diff
changeset
|
565 |
proof safe |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
566 |
fix f assume f: "f \<in> F (insert a A)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
567 |
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:
42858
diff
changeset
|
568 |
proof (rule UN_I[of "f(a := d)"]) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
569 |
show "f(a := d) \<in> F A" using *[OF f] . |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
570 |
show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
571 |
proof (rule image_eqI[of _ _ "f a"]) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
572 |
show "f a \<in> G (f(a := d))" using **[OF f] . |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
573 |
qed simp |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
574 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
575 |
next |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
576 |
fix f x assume "f \<in> F A" "x \<in> G f" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
577 |
from ***[OF this] show "f(a := x) \<in> F (insert a A)" . |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
578 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
579 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
580 |
lemma extensional_funcset_insert_eq[simp]: |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
581 |
assumes "a \<notin> A" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
582 |
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:
42858
diff
changeset
|
583 |
apply (rule funset_eq_UN_fun_upd_I) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
584 |
using assms |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
585 |
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:
42858
diff
changeset
|
586 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
587 |
lemma finite_extensional_funcset[simp, intro]: |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
588 |
assumes "finite A" "finite B" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
589 |
shows "finite (extensional A \<inter> (A \<rightarrow> B))" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
590 |
using assms by induct (auto simp: extensional_funcset_insert_eq) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
591 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
592 |
lemma finite_PiE[simp, intro]: |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
593 |
assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
594 |
shows "finite (Pi\<^isub>E A B)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
595 |
proof - |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
596 |
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:
42858
diff
changeset
|
597 |
show ?thesis |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
598 |
using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
599 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
600 |
|
| 42892 | 601 |
lemma (in product_finite_prob_space) singleton_eq_product: |
602 |
assumes x: "x \<in> space P" shows "{x} = (\<Pi>\<^isub>E i\<in>I. {x i})"
|
|
603 |
proof (safe intro!: ext[of _ x]) |
|
604 |
fix y i assume *: "y \<in> (\<Pi> i\<in>I. {x i})" "y \<in> extensional I"
|
|
605 |
with x show "y i = x i" |
|
606 |
by (cases "i \<in> I") (auto simp: extensional_def) |
|
607 |
qed (insert x, auto) |
|
608 |
||
|
42859
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
609 |
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:
42858
diff
changeset
|
610 |
proof |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
611 |
show "finite (space P)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
612 |
using finite_index M.finite_space by auto |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
613 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
614 |
{ fix x assume "x \<in> space P"
|
| 42892 | 615 |
with this[THEN singleton_eq_product] |
616 |
have "{x} \<in> sets P"
|
|
|
42859
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
617 |
by (auto intro!: in_P) } |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
618 |
note x_in_P = this |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
619 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
620 |
have "Pow (space P) \<subseteq> sets P" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
621 |
proof |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
622 |
fix X assume "X \<in> Pow (space P)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
623 |
moreover then have "finite X" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
624 |
using `finite (space P)` by (blast intro: finite_subset) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
625 |
ultimately have "(\<Union>x\<in>X. {x}) \<in> sets P"
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
626 |
by (intro finite_UN x_in_P) auto |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
627 |
then show "X \<in> sets P" by simp |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
628 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
629 |
with space_closed show [simp]: "sets P = Pow (space P)" .. |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
630 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
631 |
{ fix x assume "x \<in> space P"
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
632 |
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:
42858
diff
changeset
|
633 |
then show "\<mu> {x} \<noteq> \<infinity>"
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
634 |
using measure_space_1 by auto } |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
635 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
636 |
|
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
637 |
lemma (in product_finite_prob_space) measure_finite_times: |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
638 |
"(\<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:
42858
diff
changeset
|
639 |
by (rule measure_times) simp |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
640 |
|
| 42892 | 641 |
lemma (in product_finite_prob_space) measure_singleton_times: |
642 |
assumes x: "x \<in> space P" shows "\<mu> {x} = (\<Prod>i\<in>I. M.\<mu> i {x i})"
|
|
643 |
unfolding singleton_eq_product[OF x] using x |
|
644 |
by (intro measure_finite_times) auto |
|
645 |
||
646 |
lemma (in product_finite_prob_space) prob_finite_times: |
|
|
42859
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
647 |
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:
42858
diff
changeset
|
648 |
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:
42858
diff
changeset
|
649 |
proof - |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
650 |
have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
651 |
using X by (intro finite_measure_eq[symmetric] in_P) auto |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
652 |
also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
653 |
using measure_finite_times X by simp |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
654 |
also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))" |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
655 |
using X by (simp add: M.finite_measure_eq setprod_extreal) |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
656 |
finally show ?thesis by simp |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
657 |
qed |
|
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
658 |
|
| 42892 | 659 |
lemma (in product_finite_prob_space) prob_singleton_times: |
660 |
assumes x: "x \<in> space P" |
|
661 |
shows "prob {x} = (\<Prod>i\<in>I. M.prob i {x i})"
|
|
662 |
unfolding singleton_eq_product[OF x] using x |
|
663 |
by (intro prob_finite_times) auto |
|
664 |
||
665 |
lemma (in product_finite_prob_space) prob_finite_product: |
|
666 |
"A \<subseteq> space P \<Longrightarrow> prob A = (\<Sum>x\<in>A. \<Prod>i\<in>I. M.prob i {x i})"
|
|
667 |
by (auto simp add: finite_measure_singleton prob_singleton_times |
|
668 |
simp del: space_product_algebra |
|
669 |
intro!: setsum_cong prob_singleton_times) |
|
670 |
||
| 40859 | 671 |
lemma (in prob_space) joint_distribution_finite_prob_space: |
672 |
assumes X: "finite_random_variable MX X" |
|
673 |
assumes Y: "finite_random_variable MY Y" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
674 |
shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<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:
41661
diff
changeset
|
675 |
by (intro distribution_finite_prob_space finite_random_variable_pairI X Y) |
| 40859 | 676 |
|
| 36624 | 677 |
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:
41661
diff
changeset
|
678 |
"finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1" |
| 36624 | 679 |
unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def |
680 |
by auto |
|
681 |
||
682 |
lemma (in prob_space) not_empty: "space M \<noteq> {}"
|
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
683 |
using prob_space empty_measure' by auto |
| 36624 | 684 |
|
| 38656 | 685 |
lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
|
686 |
using measure_space_1 sum_over_space by simp |
|
| 36624 | 687 |
|
688 |
lemma (in finite_prob_space) joint_distribution_restriction_fst: |
|
689 |
"joint_distribution X Y A \<le> distribution X (fst ` A)" |
|
690 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
691 |
proof (safe intro!: finite_measure_mono) |
| 36624 | 692 |
fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" |
693 |
show "x \<in> X -` fst ` A" |
|
694 |
by (auto intro!: image_eqI[OF _ *]) |
|
695 |
qed (simp_all add: sets_eq_Pow) |
|
696 |
||
697 |
lemma (in finite_prob_space) joint_distribution_restriction_snd: |
|
698 |
"joint_distribution X Y A \<le> distribution Y (snd ` A)" |
|
699 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
700 |
proof (safe intro!: finite_measure_mono) |
| 36624 | 701 |
fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" |
702 |
show "x \<in> Y -` snd ` A" |
|
703 |
by (auto intro!: image_eqI[OF _ *]) |
|
704 |
qed (simp_all add: sets_eq_Pow) |
|
705 |
||
706 |
lemma (in finite_prob_space) distribution_order: |
|
707 |
shows "0 \<le> distribution X x'" |
|
708 |
and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')" |
|
709 |
and "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"
|
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
715 |
using |
| 36624 | 716 |
joint_distribution_restriction_fst[of X Y "{(x, y)}"]
|
717 |
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:
41831
diff
changeset
|
718 |
by (auto intro: antisym) |
| 36624 | 719 |
|
| 39097 | 720 |
lemma (in finite_prob_space) distribution_mono: |
721 |
assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" |
|
722 |
shows "distribution X x \<le> distribution Y y" |
|
723 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
724 |
using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono) |
| 39097 | 725 |
|
726 |
lemma (in finite_prob_space) distribution_mono_gt_0: |
|
727 |
assumes gt_0: "0 < distribution X x" |
|
728 |
assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" |
|
729 |
shows "0 < distribution Y y" |
|
730 |
by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) |
|
731 |
||
732 |
lemma (in finite_prob_space) sum_over_space_distrib: |
|
733 |
"(\<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:
41831
diff
changeset
|
734 |
unfolding distribution_def prob_space[symmetric] using finite_space |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
735 |
by (subst finite_measure_finite_Union[symmetric]) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
736 |
(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:
41831
diff
changeset
|
737 |
intro!: arg_cong[where f=\<mu>']) |
| 39097 | 738 |
|
739 |
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:
41831
diff
changeset
|
740 |
"(\<Sum>x\<in>space M. prob {x}) = 1"
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
741 |
using prob_space finite_space |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
742 |
by (subst (asm) finite_measure_finite_singleton) auto |
| 39097 | 743 |
|
744 |
lemma (in prob_space) distribution_remove_const: |
|
745 |
shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
|
|
746 |
and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
|
|
747 |
and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
|
|
748 |
and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
|
|
749 |
and "distribution (\<lambda>x. ()) {()} = 1"
|
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
750 |
by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric]) |
| 35977 | 751 |
|
| 39097 | 752 |
lemma (in finite_prob_space) setsum_distribution_gen: |
753 |
assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
|
|
754 |
and "inj_on f (X`space M)" |
|
755 |
shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
|
|
756 |
unfolding distribution_def assms |
|
757 |
using finite_space assms |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
758 |
by (subst finite_measure_finite_Union[symmetric]) |
| 39097 | 759 |
(auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def |
760 |
intro!: arg_cong[where f=prob]) |
|
761 |
||
762 |
lemma (in finite_prob_space) setsum_distribution: |
|
763 |
"(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
|
|
764 |
"(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
|
|
765 |
"(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
|
|
766 |
"(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
|
|
767 |
"(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
|
|
768 |
by (auto intro!: inj_onI setsum_distribution_gen) |
|
769 |
||
770 |
lemma (in finite_prob_space) uniform_prob: |
|
771 |
assumes "x \<in> space M" |
|
772 |
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:
41831
diff
changeset
|
773 |
shows "prob {x} = 1 / card (space M)"
|
| 39097 | 774 |
proof - |
775 |
have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
|
|
776 |
using assms(2)[OF _ `x \<in> space M`] by blast |
|
777 |
have "1 = prob (space M)" |
|
778 |
using prob_space by auto |
|
779 |
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:
41831
diff
changeset
|
780 |
using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
|
| 39097 | 781 |
sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] |
782 |
finite_space unfolding disjoint_family_on_def prob_space[symmetric] |
|
783 |
by (auto simp add:setsum_restrict_set) |
|
784 |
also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
|
|
785 |
using prob_x by auto |
|
786 |
also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
|
|
787 |
finally have one: "1 = real (card (space M)) * prob {x}"
|
|
788 |
using real_eq_of_nat by auto |
|
789 |
hence two: "real (card (space M)) \<noteq> 0" by fastsimp |
|
790 |
from one have three: "prob {x} \<noteq> 0" by fastsimp
|
|
791 |
thus ?thesis using one two three divide_cancel_right |
|
792 |
by (auto simp:field_simps) |
|
| 39092 | 793 |
qed |
| 35977 | 794 |
|
| 39092 | 795 |
lemma (in prob_space) prob_space_subalgebra: |
| 41545 | 796 |
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:
41661
diff
changeset
|
797 |
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:
41661
diff
changeset
|
798 |
shows "prob_space N" |
| 39092 | 799 |
proof - |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
800 |
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:
41661
diff
changeset
|
801 |
by (rule measure_space_subalgebra[OF assms]) |
| 39092 | 802 |
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:
41661
diff
changeset
|
803 |
proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1) |
| 35977 | 804 |
qed |
805 |
||
| 39092 | 806 |
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:
41831
diff
changeset
|
807 |
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:
41661
diff
changeset
|
808 |
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:
41661
diff
changeset
|
809 |
(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:
41661
diff
changeset
|
810 |
proof - |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
811 |
interpret A: measure_space "restricted_space A" |
| 39092 | 812 |
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:
41661
diff
changeset
|
813 |
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:
41661
diff
changeset
|
814 |
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:
41661
diff
changeset
|
815 |
show "prob_space ?P" |
| 39092 | 816 |
proof |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
817 |
show "measure ?P (space ?P) = 1" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
818 |
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:
41831
diff
changeset
|
819 |
show "positive ?P (measure ?P)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
820 |
proof (simp add: positive_def, safe) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
821 |
show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
822 |
fix B assume "B \<in> events" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
823 |
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:
41831
diff
changeset
|
824 |
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:
41831
diff
changeset
|
825 |
qed |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
826 |
show "countably_additive ?P (measure ?P)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
827 |
proof (simp add: countably_additive_def, safe) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
828 |
fix B and F :: "nat \<Rightarrow> 'a set" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
829 |
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:
41831
diff
changeset
|
830 |
{ fix i
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
831 |
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:
41831
diff
changeset
|
832 |
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:
41831
diff
changeset
|
833 |
moreover then have "range F \<subseteq> events" by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
834 |
moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
835 |
by (simp add: mult_commute divide_extreal_def) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
836 |
moreover have "0 \<le> inverse (\<mu> A)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
837 |
using real_measure[OF `A \<in> events`] by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
838 |
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:
41831
diff
changeset
|
839 |
using measure_countably_additive[of F] F |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
840 |
by (auto simp: suminf_cmult_extreal) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
841 |
qed |
| 39092 | 842 |
qed |
843 |
qed |
|
844 |
||
845 |
lemma finite_prob_spaceI: |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
846 |
assumes "finite (space M)" "sets M = Pow(space M)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
847 |
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:
41661
diff
changeset
|
848 |
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:
41661
diff
changeset
|
849 |
shows "finite_prob_space M" |
| 39092 | 850 |
unfolding finite_prob_space_eq |
851 |
proof |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
852 |
show "finite_measure_space M" using assms |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
853 |
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:
41661
diff
changeset
|
854 |
show "measure M (space M) = 1" by fact |
| 39092 | 855 |
qed |
| 36624 | 856 |
|
857 |
lemma (in finite_prob_space) finite_measure_space: |
|
| 39097 | 858 |
fixes X :: "'a \<Rightarrow> 'x" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
859 |
shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<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:
41661
diff
changeset
|
860 |
(is "finite_measure_space ?S") |
| 39092 | 861 |
proof (rule finite_measure_spaceI, simp_all) |
| 36624 | 862 |
show "finite (X ` space M)" using finite_space by simp |
| 39097 | 863 |
next |
864 |
fix A B :: "'x set" assume "A \<inter> B = {}"
|
|
865 |
then show "distribution X (A \<union> B) = distribution X A + distribution X B" |
|
866 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
867 |
by (subst finite_measure_Union[symmetric]) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
868 |
(auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow) |
| 36624 | 869 |
qed |
870 |
||
| 39097 | 871 |
lemma (in finite_prob_space) finite_prob_space_of_images: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
872 |
"finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
873 |
by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def) |
| 39097 | 874 |
|
| 39096 | 875 |
lemma (in finite_prob_space) finite_product_measure_space: |
| 39097 | 876 |
fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y" |
| 39096 | 877 |
assumes "finite s1" "finite s2" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
878 |
shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<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:
41661
diff
changeset
|
879 |
(is "finite_measure_space ?M") |
| 39097 | 880 |
proof (rule finite_measure_spaceI, simp_all) |
881 |
show "finite (s1 \<times> s2)" |
|
| 39096 | 882 |
using assms by auto |
| 39097 | 883 |
next |
884 |
fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
|
|
885 |
then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B" |
|
886 |
unfolding distribution_def |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
887 |
by (subst finite_measure_Union[symmetric]) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
888 |
(auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow) |
| 39096 | 889 |
qed |
890 |
||
| 39097 | 891 |
lemma (in finite_prob_space) finite_product_measure_space_of_images: |
| 39096 | 892 |
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:
41661
diff
changeset
|
893 |
sets = Pow (X ` space M \<times> Y ` space M), |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
894 |
measure = extreal \<circ> joint_distribution X Y \<rparr>" |
| 39096 | 895 |
using finite_space by (auto intro!: finite_product_measure_space) |
896 |
||
| 40859 | 897 |
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:
41661
diff
changeset
|
898 |
"finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M), |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
899 |
measure = extreal \<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:
41661
diff
changeset
|
900 |
(is "finite_prob_space ?S") |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
901 |
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def) |
| 40859 | 902 |
have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto |
903 |
thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1" |
|
904 |
by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) |
|
905 |
qed |
|
906 |
||
| 39085 | 907 |
section "Conditional Expectation and Probability" |
908 |
||
909 |
lemma (in prob_space) conditional_expectation_exists: |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
910 |
fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
911 |
assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
912 |
and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
913 |
shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N. |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
914 |
(\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))" |
| 39083 | 915 |
proof - |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
916 |
note N(4)[simp] |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
917 |
interpret P: prob_space N |
| 41545 | 918 |
using prob_space_subalgebra[OF N] . |
| 39083 | 919 |
|
920 |
let "?f A" = "\<lambda>x. X x * indicator A x" |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
921 |
let "?Q A" = "integral\<^isup>P M (?f A)" |
| 39083 | 922 |
|
923 |
from measure_space_density[OF borel] |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
924 |
have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)" |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
925 |
apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"]) |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
926 |
using N by (auto intro!: P.sigma_algebra_cong) |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
927 |
then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" . |
| 39083 | 928 |
|
929 |
have "P.absolutely_continuous ?Q" |
|
930 |
unfolding P.absolutely_continuous_def |
|
| 41545 | 931 |
proof safe |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
932 |
fix A assume "A \<in> sets N" "P.\<mu> A = 0" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
933 |
then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
934 |
using borel N by (auto intro!: borel_measurable_indicator AE_not_in) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
935 |
then show "?Q A = 0" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
936 |
by (auto simp add: positive_integral_0_iff_AE) |
| 39083 | 937 |
qed |
938 |
from P.Radon_Nikodym[OF Q this] |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
939 |
obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
940 |
"\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)" |
| 39083 | 941 |
by blast |
| 41545 | 942 |
with N(2) show ?thesis |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
943 |
by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)]) |
| 39083 | 944 |
qed |
945 |
||
| 39085 | 946 |
definition (in prob_space) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
947 |
"conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
948 |
\<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))" |
| 39085 | 949 |
|
950 |
abbreviation (in prob_space) |
|
| 39092 | 951 |
"conditional_prob N A \<equiv> conditional_expectation N (indicator A)" |
| 39085 | 952 |
|
953 |
lemma (in prob_space) |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
954 |
fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
955 |
assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
956 |
and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A" |
| 39085 | 957 |
shows borel_measurable_conditional_expectation: |
| 41545 | 958 |
"conditional_expectation N X \<in> borel_measurable N" |
959 |
and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow> |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
960 |
(\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) = |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
961 |
(\<integral>\<^isup>+x. X x * indicator C x \<partial>M)" |
| 41545 | 962 |
(is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C") |
| 39085 | 963 |
proof - |
964 |
note CE = conditional_expectation_exists[OF assms, unfolded Bex_def] |
|
| 41545 | 965 |
then show "conditional_expectation N X \<in> borel_measurable N" |
| 39085 | 966 |
unfolding conditional_expectation_def by (rule someI2_ex) blast |
967 |
||
| 41545 | 968 |
from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C" |
| 39085 | 969 |
unfolding conditional_expectation_def by (rule someI2_ex) blast |
970 |
qed |
|
971 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
972 |
lemma (in sigma_algebra) factorize_measurable_function_pos: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
973 |
fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c" |
| 39091 | 974 |
assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
975 |
assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
976 |
shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
977 |
proof - |
| 39091 | 978 |
interpret M': sigma_algebra M' by fact |
979 |
have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto |
|
980 |
from M'.sigma_algebra_vimage[OF this] |
|
981 |
interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . |
|
982 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
983 |
from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this |
| 39091 | 984 |
|
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
985 |
have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" |
| 39091 | 986 |
proof |
987 |
fix i |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
988 |
from f(1)[of i] have "finite (f i`space M)" and B_ex: |
| 39091 | 989 |
"\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> 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:
41661
diff
changeset
|
990 |
unfolding simple_function_def by auto |
| 39091 | 991 |
from B_ex[THEN bchoice] guess B .. note B = this |
992 |
||
993 |
let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x" |
|
994 |
||
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
995 |
show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" |
| 39091 | 996 |
proof (intro exI[of _ ?g] conjI ballI) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
997 |
show "simple_function M' ?g" using B by auto |
| 39091 | 998 |
|
999 |
fix x assume "x \<in> space M" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1000 |
then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
|
| 39091 | 1001 |
unfolding indicator_def using B by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1002 |
then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i] |
| 39091 | 1003 |
by (subst va.simple_function_indicator_representation) auto |
1004 |
qed |
|
1005 |
qed |
|
1006 |
from choice[OF this] guess g .. note g = this |
|
1007 |
||
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1008 |
show ?thesis |
| 39091 | 1009 |
proof (intro ballI bexI) |
|
41097
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents:
41095
diff
changeset
|
1010 |
show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'" |
| 39091 | 1011 |
using g by (auto intro: M'.borel_measurable_simple_function) |
1012 |
fix x assume "x \<in> space M" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1013 |
have "max 0 (Z x) = (SUP i. f i x)" using f by simp |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1014 |
also have "\<dots> = (SUP i. g i (Y x))" |
| 39091 | 1015 |
using g `x \<in> space M` by simp |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1016 |
finally show "max 0 (Z x) = (SUP i. g i (Y x))" . |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1017 |
qed |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1018 |
qed |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1019 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1020 |
lemma extreal_0_le_iff_le_0[simp]: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1021 |
fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1022 |
by (cases rule: extreal2_cases[of a]) auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1023 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1024 |
lemma (in sigma_algebra) factorize_measurable_function: |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1025 |
fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1026 |
assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1027 |
shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1028 |
\<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1029 |
proof safe |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1030 |
interpret M': sigma_algebra M' by fact |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1031 |
have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1032 |
from M'.sigma_algebra_vimage[OF this] |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1033 |
interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1034 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1035 |
{ fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1036 |
with M'.measurable_vimage_algebra[OF Y] |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1037 |
have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1038 |
by (rule measurable_comp) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1039 |
moreover assume "\<forall>x\<in>space M. Z x = g (Y x)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1040 |
then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow> |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1041 |
g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1042 |
by (auto intro!: measurable_cong) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1043 |
ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1044 |
by simp } |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1045 |
|
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1046 |
assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1047 |
with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1048 |
"(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1049 |
by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1050 |
from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1051 |
from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1052 |
let "?g x" = "p x - n x" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1053 |
show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1054 |
proof (intro bexI ballI) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1055 |
show "?g \<in> borel_measurable M'" using p n by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1056 |
fix x assume "x \<in> space M" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1057 |
then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1058 |
using p n by auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1059 |
then show "Z x = ?g (Y x)" |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1060 |
by (auto split: split_max) |
| 39091 | 1061 |
qed |
1062 |
qed |
|
|
39090
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
hoelzl
parents:
39089
diff
changeset
|
1063 |
|
| 42860 | 1064 |
subsection "Bernoulli space" |
1065 |
||
1066 |
definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV, |
|
1067 |
measure = extreal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>" |
|
1068 |
||
1069 |
interpretation bernoulli: finite_prob_space "bernoulli_space p" for p |
|
1070 |
by (rule finite_prob_spaceI) |
|
1071 |
(auto simp: bernoulli_space_def UNIV_bool one_extreal_def setsum_Un_disjoint intro!: setsum_nonneg) |
|
1072 |
||
1073 |
lemma bernoulli_measure: |
|
1074 |
"0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)" |
|
1075 |
unfolding bernoulli.\<mu>'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong) |
|
1076 |
||
1077 |
lemma bernoulli_measure_True: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {True} = p"
|
|
1078 |
and bernoulli_measure_False: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {False} = 1 - p"
|
|
1079 |
unfolding bernoulli_measure by simp_all |
|
1080 |
||
| 35582 | 1081 |
end |