author | hoelzl |
Fri, 27 Aug 2010 16:23:51 +0200 | |
changeset 39091 | 11314c196e11 |
parent 39090 | a2d38b8b693e |
child 39092 | 98de40859858 |
permissions | -rw-r--r-- |
35582 | 1 |
theory Probability_Space |
39083 | 2 |
imports Lebesgue_Integration Radon_Nikodym |
35582 | 3 |
begin |
4 |
||
38656 | 5 |
lemma (in measure_space) measure_inter_full_set: |
6 |
assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>" |
|
7 |
assumes T: "\<mu> T = \<mu> (space M)" |
|
8 |
shows "\<mu> (S \<inter> T) = \<mu> S" |
|
9 |
proof (rule antisym) |
|
10 |
show " \<mu> (S \<inter> T) \<le> \<mu> S" |
|
11 |
using assms by (auto intro!: measure_mono) |
|
12 |
||
13 |
show "\<mu> S \<le> \<mu> (S \<inter> T)" |
|
14 |
proof (rule ccontr) |
|
15 |
assume contr: "\<not> ?thesis" |
|
16 |
have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))" |
|
17 |
unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"]) |
|
18 |
also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)" |
|
19 |
using assms by (auto intro!: measure_subadditive) |
|
20 |
also have "\<dots> < \<mu> (T - S) + \<mu> S" |
|
21 |
by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto) |
|
22 |
also have "\<dots> = \<mu> (T \<union> S)" |
|
23 |
using assms by (subst measure_additive) auto |
|
24 |
also have "\<dots> \<le> \<mu> (space M)" |
|
25 |
using assms sets_into_space by (auto intro!: measure_mono) |
|
26 |
finally show False .. |
|
27 |
qed |
|
28 |
qed |
|
29 |
||
30 |
lemma (in finite_measure) finite_measure_inter_full_set: |
|
31 |
assumes "S \<in> sets M" "T \<in> sets M" |
|
32 |
assumes T: "\<mu> T = \<mu> (space M)" |
|
33 |
shows "\<mu> (S \<inter> T) = \<mu> S" |
|
34 |
using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms |
|
35 |
by auto |
|
36 |
||
35582 | 37 |
locale prob_space = measure_space + |
38656 | 38 |
assumes measure_space_1: "\<mu> (space M) = 1" |
39 |
||
40 |
sublocale prob_space < finite_measure |
|
41 |
proof |
|
42 |
from measure_space_1 show "\<mu> (space M) \<noteq> \<omega>" by simp |
|
43 |
qed |
|
44 |
||
45 |
context prob_space |
|
35582 | 46 |
begin |
47 |
||
48 |
abbreviation "events \<equiv> sets M" |
|
38656 | 49 |
abbreviation "prob \<equiv> \<lambda>A. real (\<mu> A)" |
35582 | 50 |
abbreviation "prob_preserving \<equiv> measure_preserving" |
51 |
abbreviation "random_variable \<equiv> \<lambda> s X. X \<in> measurable M s" |
|
52 |
abbreviation "expectation \<equiv> integral" |
|
53 |
||
54 |
definition |
|
55 |
"indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B" |
|
56 |
||
57 |
definition |
|
58 |
"indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)" |
|
59 |
||
60 |
definition |
|
38656 | 61 |
"distribution X = (\<lambda>s. \<mu> ((X -` s) \<inter> (space M)))" |
35582 | 62 |
|
36624 | 63 |
abbreviation |
64 |
"joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))" |
|
35582 | 65 |
|
38656 | 66 |
lemma prob_space: "prob (space M) = 1" |
67 |
unfolding measure_space_1 by simp |
|
35582 | 68 |
|
38656 | 69 |
lemma measure_le_1[simp, intro]: |
70 |
assumes "A \<in> events" shows "\<mu> A \<le> 1" |
|
71 |
proof - |
|
72 |
have "\<mu> A \<le> \<mu> (space M)" |
|
73 |
using assms sets_into_space by(auto intro!: measure_mono) |
|
74 |
also note measure_space_1 |
|
75 |
finally show ?thesis . |
|
76 |
qed |
|
35582 | 77 |
|
38656 | 78 |
lemma measure_finite[simp, intro]: |
79 |
assumes "A \<in> events" shows "\<mu> A \<noteq> \<omega>" |
|
80 |
using measure_le_1[OF assms] by auto |
|
35582 | 81 |
|
82 |
lemma prob_compl: |
|
38656 | 83 |
assumes "A \<in> events" |
84 |
shows "prob (space M - A) = 1 - prob A" |
|
85 |
using `A \<in> events`[THEN sets_into_space] `A \<in> events` measure_space_1 |
|
86 |
by (subst real_finite_measure_Diff) auto |
|
35582 | 87 |
|
88 |
lemma indep_space: |
|
89 |
assumes "s \<in> events" |
|
90 |
shows "indep (space M) s" |
|
38656 | 91 |
using assms prob_space by (simp add: indep_def) |
35582 | 92 |
|
38656 | 93 |
lemma prob_space_increasing: "increasing M prob" |
94 |
by (auto intro!: real_measure_mono simp: increasing_def) |
|
35582 | 95 |
|
96 |
lemma prob_zero_union: |
|
97 |
assumes "s \<in> events" "t \<in> events" "prob t = 0" |
|
98 |
shows "prob (s \<union> t) = prob s" |
|
38656 | 99 |
using assms |
35582 | 100 |
proof - |
101 |
have "prob (s \<union> t) \<le> prob s" |
|
38656 | 102 |
using real_finite_measure_subadditive[of s t] assms by auto |
35582 | 103 |
moreover have "prob (s \<union> t) \<ge> prob s" |
38656 | 104 |
using assms by (blast intro: real_measure_mono) |
35582 | 105 |
ultimately show ?thesis by simp |
106 |
qed |
|
107 |
||
108 |
lemma prob_eq_compl: |
|
109 |
assumes "s \<in> events" "t \<in> events" |
|
110 |
assumes "prob (space M - s) = prob (space M - t)" |
|
111 |
shows "prob s = prob t" |
|
38656 | 112 |
using assms prob_compl by auto |
35582 | 113 |
|
114 |
lemma prob_one_inter: |
|
115 |
assumes events:"s \<in> events" "t \<in> events" |
|
116 |
assumes "prob t = 1" |
|
117 |
shows "prob (s \<inter> t) = prob s" |
|
118 |
proof - |
|
38656 | 119 |
have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)" |
120 |
using events assms prob_compl[of "t"] by (auto intro!: prob_zero_union) |
|
121 |
also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)" |
|
122 |
by blast |
|
123 |
finally show "prob (s \<inter> t) = prob s" |
|
124 |
using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s]) |
|
35582 | 125 |
qed |
126 |
||
127 |
lemma prob_eq_bigunion_image: |
|
128 |
assumes "range f \<subseteq> events" "range g \<subseteq> events" |
|
129 |
assumes "disjoint_family f" "disjoint_family g" |
|
130 |
assumes "\<And> n :: nat. prob (f n) = prob (g n)" |
|
131 |
shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))" |
|
132 |
using assms |
|
133 |
proof - |
|
38656 | 134 |
have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))" |
135 |
by (rule real_finite_measure_UNION[OF assms(1,3)]) |
|
136 |
have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))" |
|
137 |
by (rule real_finite_measure_UNION[OF assms(2,4)]) |
|
138 |
show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp |
|
35582 | 139 |
qed |
140 |
||
141 |
lemma prob_countably_zero: |
|
142 |
assumes "range c \<subseteq> events" |
|
143 |
assumes "\<And> i. prob (c i) = 0" |
|
38656 | 144 |
shows "prob (\<Union> i :: nat. c i) = 0" |
145 |
proof (rule antisym) |
|
146 |
show "prob (\<Union> i :: nat. c i) \<le> 0" |
|
147 |
using real_finite_measurable_countably_subadditive[OF assms(1)] |
|
148 |
by (simp add: assms(2) suminf_zero summable_zero) |
|
149 |
show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pinfreal_nonneg) |
|
35582 | 150 |
qed |
151 |
||
152 |
lemma indep_sym: |
|
153 |
"indep a b \<Longrightarrow> indep b a" |
|
154 |
unfolding indep_def using Int_commute[of a b] by auto |
|
155 |
||
156 |
lemma indep_refl: |
|
157 |
assumes "a \<in> events" |
|
158 |
shows "indep a a = (prob a = 0) \<or> (prob a = 1)" |
|
159 |
using assms unfolding indep_def by auto |
|
160 |
||
161 |
lemma prob_equiprobable_finite_unions: |
|
38656 | 162 |
assumes "s \<in> events" |
163 |
assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events" |
|
35582 | 164 |
assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})" |
38656 | 165 |
shows "prob s = real (card s) * prob {SOME x. x \<in> s}" |
35582 | 166 |
proof (cases "s = {}") |
38656 | 167 |
case False hence "\<exists> x. x \<in> s" by blast |
35582 | 168 |
from someI_ex[OF this] assms |
169 |
have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast |
|
170 |
have "prob s = (\<Sum> x \<in> s. prob {x})" |
|
38656 | 171 |
using real_finite_measure_finite_singelton[OF s_finite] by simp |
35582 | 172 |
also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto |
38656 | 173 |
also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}" |
174 |
using setsum_constant assms by (simp add: real_eq_of_nat) |
|
35582 | 175 |
finally show ?thesis by simp |
38656 | 176 |
qed simp |
35582 | 177 |
|
178 |
lemma prob_real_sum_image_fn: |
|
179 |
assumes "e \<in> events" |
|
180 |
assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events" |
|
181 |
assumes "finite s" |
|
38656 | 182 |
assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}" |
183 |
assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)" |
|
35582 | 184 |
shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))" |
185 |
proof - |
|
38656 | 186 |
have e: "e = (\<Union> i \<in> s. e \<inter> f i)" |
187 |
using `e \<in> events` sets_into_space upper by blast |
|
188 |
hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp |
|
189 |
also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))" |
|
190 |
proof (rule real_finite_measure_finite_Union) |
|
191 |
show "finite s" by fact |
|
192 |
show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact |
|
193 |
show "disjoint_family_on (\<lambda>i. e \<inter> f i) s" |
|
194 |
using disjoint by (auto simp: disjoint_family_on_def) |
|
195 |
qed |
|
196 |
finally show ?thesis . |
|
35582 | 197 |
qed |
198 |
||
199 |
lemma distribution_prob_space: |
|
39089 | 200 |
assumes S: "sigma_algebra S" "random_variable S X" |
38656 | 201 |
shows "prob_space S (distribution X)" |
35582 | 202 |
proof - |
39089 | 203 |
interpret S: measure_space S "distribution X" |
204 |
using measure_space_vimage[OF S(2,1)] unfolding distribution_def . |
|
38656 | 205 |
show ?thesis |
206 |
proof |
|
207 |
have "X -` space S \<inter> space M = space M" |
|
208 |
using `random_variable S X` by (auto simp: measurable_def) |
|
39089 | 209 |
then show "distribution X (space S) = 1" |
210 |
using measure_space_1 by (simp add: distribution_def) |
|
35582 | 211 |
qed |
212 |
qed |
|
213 |
||
214 |
lemma distribution_lebesgue_thm1: |
|
215 |
assumes "random_variable s X" |
|
216 |
assumes "A \<in> sets s" |
|
38656 | 217 |
shows "real (distribution X A) = expectation (indicator (X -` A \<inter> space M))" |
35582 | 218 |
unfolding distribution_def |
219 |
using assms unfolding measurable_def |
|
38656 | 220 |
using integral_indicator by auto |
35582 | 221 |
|
222 |
lemma distribution_lebesgue_thm2: |
|
38656 | 223 |
assumes "sigma_algebra S" "random_variable S X" and "A \<in> sets S" |
224 |
shows "distribution X A = |
|
225 |
measure_space.positive_integral S (distribution X) (indicator A)" |
|
226 |
(is "_ = measure_space.positive_integral _ ?D _") |
|
35582 | 227 |
proof - |
38656 | 228 |
interpret S: prob_space S "distribution X" using assms(1,2) by (rule distribution_prob_space) |
35582 | 229 |
|
230 |
show ?thesis |
|
38656 | 231 |
using S.positive_integral_indicator(1) |
35582 | 232 |
using assms unfolding distribution_def by auto |
233 |
qed |
|
234 |
||
235 |
lemma finite_expectation1: |
|
38656 | 236 |
assumes "finite (X`space M)" and rv: "random_variable borel_space X" |
35582 | 237 |
shows "expectation X = (\<Sum> r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" |
38656 | 238 |
proof (rule integral_on_finite(2)[OF assms(2,1)]) |
239 |
fix x have "X -` {x} \<inter> space M \<in> sets M" |
|
240 |
using rv unfolding measurable_def by auto |
|
241 |
thus "\<mu> (X -` {x} \<inter> space M) \<noteq> \<omega>" using finite_measure by simp |
|
242 |
qed |
|
35582 | 243 |
|
244 |
lemma finite_expectation: |
|
38656 | 245 |
assumes "finite (space M)" "random_variable borel_space X" |
246 |
shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))" |
|
247 |
using assms unfolding distribution_def using finite_expectation1 by auto |
|
248 |
||
35582 | 249 |
lemma prob_x_eq_1_imp_prob_y_eq_0: |
250 |
assumes "{x} \<in> events" |
|
38656 | 251 |
assumes "prob {x} = 1" |
35582 | 252 |
assumes "{y} \<in> events" |
253 |
assumes "y \<noteq> x" |
|
254 |
shows "prob {y} = 0" |
|
255 |
using prob_one_inter[of "{y}" "{x}"] assms by auto |
|
256 |
||
38656 | 257 |
lemma distribution_empty[simp]: "distribution X {} = 0" |
258 |
unfolding distribution_def by simp |
|
259 |
||
260 |
lemma distribution_space[simp]: "distribution X (X ` space M) = 1" |
|
261 |
proof - |
|
262 |
have "X -` X ` space M \<inter> space M = space M" by auto |
|
263 |
thus ?thesis unfolding distribution_def by (simp add: measure_space_1) |
|
264 |
qed |
|
265 |
||
266 |
lemma distribution_one: |
|
267 |
assumes "random_variable M X" and "A \<in> events" |
|
268 |
shows "distribution X A \<le> 1" |
|
269 |
proof - |
|
270 |
have "distribution X A \<le> \<mu> (space M)" unfolding distribution_def |
|
271 |
using assms[unfolded measurable_def] by (auto intro!: measure_mono) |
|
272 |
thus ?thesis by (simp add: measure_space_1) |
|
273 |
qed |
|
274 |
||
275 |
lemma distribution_finite: |
|
276 |
assumes "random_variable M X" and "A \<in> events" |
|
277 |
shows "distribution X A \<noteq> \<omega>" |
|
278 |
using distribution_one[OF assms] by auto |
|
279 |
||
35582 | 280 |
lemma distribution_x_eq_1_imp_distribution_y_eq_0: |
281 |
assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X" |
|
38656 | 282 |
(is "random_variable ?S X") |
283 |
assumes "distribution X {x} = 1" |
|
35582 | 284 |
assumes "y \<noteq> x" |
285 |
shows "distribution X {y} = 0" |
|
286 |
proof - |
|
38656 | 287 |
have "sigma_algebra ?S" by (rule sigma_algebra_Pow) |
288 |
from distribution_prob_space[OF this X] |
|
289 |
interpret S: prob_space ?S "distribution X" by simp |
|
290 |
||
291 |
have x: "{x} \<in> sets ?S" |
|
292 |
proof (rule ccontr) |
|
293 |
assume "{x} \<notin> sets ?S" |
|
35582 | 294 |
hence "X -` {x} \<inter> space M = {}" by auto |
38656 | 295 |
thus "False" using assms unfolding distribution_def by auto |
296 |
qed |
|
297 |
||
298 |
have [simp]: "{y} \<inter> {x} = {}" "{x} - {y} = {x}" using `y \<noteq> x` by auto |
|
299 |
||
300 |
show ?thesis |
|
301 |
proof cases |
|
302 |
assume "{y} \<in> sets ?S" |
|
303 |
with `{x} \<in> sets ?S` assms show "distribution X {y} = 0" |
|
304 |
using S.measure_inter_full_set[of "{y}" "{x}"] |
|
305 |
by simp |
|
306 |
next |
|
307 |
assume "{y} \<notin> sets ?S" |
|
35582 | 308 |
hence "X -` {y} \<inter> space M = {}" by auto |
38656 | 309 |
thus "distribution X {y} = 0" unfolding distribution_def by auto |
310 |
qed |
|
35582 | 311 |
qed |
312 |
||
313 |
end |
|
314 |
||
35977 | 315 |
locale finite_prob_space = prob_space + finite_measure_space |
316 |
||
36624 | 317 |
lemma finite_prob_space_eq: |
38656 | 318 |
"finite_prob_space M \<mu> \<longleftrightarrow> finite_measure_space M \<mu> \<and> \<mu> (space M) = 1" |
36624 | 319 |
unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def |
320 |
by auto |
|
321 |
||
322 |
lemma (in prob_space) not_empty: "space M \<noteq> {}" |
|
323 |
using prob_space empty_measure by auto |
|
324 |
||
38656 | 325 |
lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1" |
326 |
using measure_space_1 sum_over_space by simp |
|
36624 | 327 |
|
328 |
lemma (in finite_prob_space) positive_distribution: "0 \<le> distribution X x" |
|
38656 | 329 |
unfolding distribution_def by simp |
36624 | 330 |
|
331 |
lemma (in finite_prob_space) joint_distribution_restriction_fst: |
|
332 |
"joint_distribution X Y A \<le> distribution X (fst ` A)" |
|
333 |
unfolding distribution_def |
|
334 |
proof (safe intro!: measure_mono) |
|
335 |
fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" |
|
336 |
show "x \<in> X -` fst ` A" |
|
337 |
by (auto intro!: image_eqI[OF _ *]) |
|
338 |
qed (simp_all add: sets_eq_Pow) |
|
339 |
||
340 |
lemma (in finite_prob_space) joint_distribution_restriction_snd: |
|
341 |
"joint_distribution X Y A \<le> distribution Y (snd ` A)" |
|
342 |
unfolding distribution_def |
|
343 |
proof (safe intro!: measure_mono) |
|
344 |
fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A" |
|
345 |
show "x \<in> Y -` snd ` A" |
|
346 |
by (auto intro!: image_eqI[OF _ *]) |
|
347 |
qed (simp_all add: sets_eq_Pow) |
|
348 |
||
349 |
lemma (in finite_prob_space) distribution_order: |
|
350 |
shows "0 \<le> distribution X x'" |
|
351 |
and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')" |
|
352 |
and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}" |
|
353 |
and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}" |
|
354 |
and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}" |
|
355 |
and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}" |
|
356 |
and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0" |
|
357 |
and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0" |
|
358 |
using positive_distribution[of X x'] |
|
359 |
positive_distribution[of "\<lambda>x. (X x, Y x)" "{(x, y)}"] |
|
360 |
joint_distribution_restriction_fst[of X Y "{(x, y)}"] |
|
361 |
joint_distribution_restriction_snd[of X Y "{(x, y)}"] |
|
362 |
by auto |
|
363 |
||
364 |
lemma (in finite_prob_space) finite_product_measure_space: |
|
35977 | 365 |
assumes "finite s1" "finite s2" |
38656 | 366 |
shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)" |
367 |
(is "finite_measure_space ?M ?D") |
|
35977 | 368 |
proof (rule finite_Pow_additivity_sufficient) |
38656 | 369 |
show "positive ?D" |
370 |
unfolding positive_def using assms sets_eq_Pow |
|
36624 | 371 |
by (simp add: distribution_def) |
35977 | 372 |
|
38656 | 373 |
show "additive ?M ?D" unfolding additive_def |
35977 | 374 |
proof safe |
375 |
fix x y |
|
376 |
have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto |
|
377 |
have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto |
|
378 |
assume "x \<inter> y = {}" |
|
38656 | 379 |
hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}" |
380 |
by auto |
|
35977 | 381 |
from additive[unfolded additive_def, rule_format, OF A B] this |
38656 | 382 |
finite_measure[OF A] finite_measure[OF B] |
383 |
show "?D (x \<union> y) = ?D x + ?D y" |
|
36624 | 384 |
apply (simp add: distribution_def) |
35977 | 385 |
apply (subst Int_Un_distrib2) |
38656 | 386 |
by (auto simp: real_of_pinfreal_add) |
35977 | 387 |
qed |
388 |
||
389 |
show "finite (space ?M)" |
|
390 |
using assms by auto |
|
391 |
||
392 |
show "sets ?M = Pow (space ?M)" |
|
393 |
by simp |
|
38656 | 394 |
|
395 |
{ fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>" |
|
396 |
unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } |
|
35977 | 397 |
qed |
398 |
||
36624 | 399 |
lemma (in finite_prob_space) finite_product_measure_space_of_images: |
35977 | 400 |
shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M, |
38656 | 401 |
sets = Pow (X ` space M \<times> Y ` space M) \<rparr> |
402 |
(joint_distribution X Y)" |
|
36624 | 403 |
using finite_space by (auto intro!: finite_product_measure_space) |
404 |
||
405 |
lemma (in finite_prob_space) finite_measure_space: |
|
38656 | 406 |
shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" |
407 |
(is "finite_measure_space ?S _") |
|
36624 | 408 |
proof (rule finite_Pow_additivity_sufficient, simp_all) |
409 |
show "finite (X ` space M)" using finite_space by simp |
|
410 |
||
38656 | 411 |
show "positive (distribution X)" |
412 |
unfolding distribution_def positive_def using sets_eq_Pow by auto |
|
36624 | 413 |
|
414 |
show "additive ?S (distribution X)" unfolding additive_def distribution_def |
|
415 |
proof (simp, safe) |
|
416 |
fix x y |
|
417 |
have x: "(X -` x) \<inter> space M \<in> sets M" |
|
418 |
and y: "(X -` y) \<inter> space M \<in> sets M" using sets_eq_Pow by auto |
|
419 |
assume "x \<inter> y = {}" |
|
38656 | 420 |
hence "X -` x \<inter> space M \<inter> (X -` y \<inter> space M) = {}" by auto |
36624 | 421 |
from additive[unfolded additive_def, rule_format, OF x y] this |
38656 | 422 |
finite_measure[OF x] finite_measure[OF y] |
423 |
have "\<mu> (((X -` x) \<union> (X -` y)) \<inter> space M) = |
|
424 |
\<mu> ((X -` x) \<inter> space M) + \<mu> ((X -` y) \<inter> space M)" |
|
425 |
by (subst Int_Un_distrib2) auto |
|
426 |
thus "\<mu> ((X -` x \<union> X -` y) \<inter> space M) = \<mu> (X -` x \<inter> space M) + \<mu> (X -` y \<inter> space M)" |
|
36624 | 427 |
by auto |
428 |
qed |
|
38656 | 429 |
|
430 |
{ fix x assume "x \<in> X ` space M" thus "distribution X {x} \<noteq> \<omega>" |
|
431 |
unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } |
|
36624 | 432 |
qed |
433 |
||
434 |
lemma (in finite_prob_space) finite_prob_space_of_images: |
|
38656 | 435 |
"finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" |
436 |
by (simp add: finite_prob_space_eq finite_measure_space) |
|
36624 | 437 |
|
438 |
lemma (in finite_prob_space) finite_product_prob_space_of_images: |
|
38656 | 439 |
"finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr> |
440 |
(joint_distribution X Y)" |
|
441 |
(is "finite_prob_space ?S _") |
|
442 |
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) |
|
36624 | 443 |
have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto |
444 |
thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1" |
|
38656 | 445 |
by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) |
36624 | 446 |
qed |
35977 | 447 |
|
39083 | 448 |
lemma (in prob_space) prob_space_subalgebra: |
449 |
assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" |
|
450 |
shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry |
|
451 |
||
452 |
lemma (in measure_space) measure_space_subalgebra: |
|
453 |
assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" |
|
454 |
shows "measure_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry |
|
455 |
||
456 |
lemma pinfreal_0_less_mult_iff[simp]: |
|
457 |
fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y" |
|
458 |
by (cases x, cases y) (auto simp: zero_less_mult_iff) |
|
459 |
||
460 |
lemma (in sigma_algebra) simple_function_subalgebra: |
|
461 |
assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f" |
|
462 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)" |
|
463 |
shows "simple_function f" |
|
464 |
using assms |
|
465 |
unfolding simple_function_def |
|
466 |
unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)] |
|
467 |
by auto |
|
468 |
||
469 |
lemma (in measure_space) simple_integral_subalgebra[simp]: |
|
470 |
assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>" |
|
471 |
shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral" |
|
472 |
unfolding simple_integral_def_raw |
|
473 |
unfolding measure_space.simple_integral_def_raw[OF assms] by simp |
|
474 |
||
475 |
lemma (in sigma_algebra) borel_measurable_subalgebra: |
|
476 |
assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)" |
|
477 |
shows "f \<in> borel_measurable M" |
|
478 |
using assms unfolding measurable_def by auto |
|
479 |
||
480 |
lemma (in measure_space) positive_integral_subalgebra[simp]: |
|
481 |
assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)" |
|
482 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)" |
|
483 |
shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f" |
|
484 |
proof - |
|
485 |
note msN = measure_space_subalgebra[OF N_subalgebra] |
|
486 |
then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> . |
|
487 |
||
488 |
from N.borel_measurable_implies_simple_function_sequence[OF borel] |
|
489 |
obtain fs where Nsf: "\<And>i. N.simple_function (fs i)" and "fs \<up> f" by blast |
|
490 |
then have sf: "\<And>i. simple_function (fs i)" |
|
491 |
using simple_function_subalgebra[OF _ N_subalgebra] by blast |
|
492 |
||
493 |
from positive_integral_isoton_simple[OF `fs \<up> f` sf] N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf] |
|
494 |
show ?thesis unfolding simple_integral_subalgebra[OF msN] isoton_def by simp |
|
495 |
qed |
|
39084 | 496 |
|
39085 | 497 |
section "Conditional Expectation and Probability" |
498 |
||
499 |
lemma (in prob_space) conditional_expectation_exists: |
|
39083 | 500 |
fixes X :: "'a \<Rightarrow> pinfreal" |
501 |
assumes borel: "X \<in> borel_measurable M" |
|
502 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" |
|
503 |
shows "\<exists>Y\<in>borel_measurable (M\<lparr> sets := N \<rparr>). \<forall>C\<in>N. |
|
504 |
positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)" |
|
505 |
proof - |
|
506 |
interpret P: prob_space "M\<lparr> sets := N \<rparr>" \<mu> |
|
507 |
using prob_space_subalgebra[OF N_subalgebra] . |
|
508 |
||
509 |
let "?f A" = "\<lambda>x. X x * indicator A x" |
|
510 |
let "?Q A" = "positive_integral (?f A)" |
|
511 |
||
512 |
from measure_space_density[OF borel] |
|
513 |
have Q: "measure_space (M\<lparr> sets := N \<rparr>) ?Q" |
|
514 |
by (rule measure_space.measure_space_subalgebra[OF _ N_subalgebra]) |
|
515 |
then interpret Q: measure_space "M\<lparr> sets := N \<rparr>" ?Q . |
|
516 |
||
517 |
have "P.absolutely_continuous ?Q" |
|
518 |
unfolding P.absolutely_continuous_def |
|
519 |
proof (safe, simp) |
|
520 |
fix A assume "A \<in> N" "\<mu> A = 0" |
|
521 |
moreover then have f_borel: "?f A \<in> borel_measurable M" |
|
522 |
using borel N_subalgebra by (auto intro: borel_measurable_indicator) |
|
523 |
moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A" |
|
524 |
by (auto simp: indicator_def) |
|
525 |
moreover have "\<mu> \<dots> \<le> \<mu> A" |
|
526 |
using `A \<in> N` N_subalgebra f_borel |
|
527 |
by (auto intro!: measure_mono Int[of _ A] measurable_sets) |
|
528 |
ultimately show "?Q A = 0" |
|
529 |
by (simp add: positive_integral_0_iff) |
|
530 |
qed |
|
531 |
from P.Radon_Nikodym[OF Q this] |
|
532 |
obtain Y where Y: "Y \<in> borel_measurable (M\<lparr>sets := N\<rparr>)" |
|
533 |
"\<And>A. A \<in> sets (M\<lparr>sets:=N\<rparr>) \<Longrightarrow> ?Q A = P.positive_integral (\<lambda>x. Y x * indicator A x)" |
|
534 |
by blast |
|
39084 | 535 |
with N_subalgebra show ?thesis |
536 |
by (auto intro!: bexI[OF _ Y(1)]) |
|
39083 | 537 |
qed |
538 |
||
39085 | 539 |
definition (in prob_space) |
540 |
"conditional_expectation N X = (SOME Y. Y\<in>borel_measurable (M\<lparr>sets:=N\<rparr>) |
|
541 |
\<and> (\<forall>C\<in>N. positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)))" |
|
542 |
||
543 |
abbreviation (in prob_space) |
|
544 |
"conditional_probabiltiy N A \<equiv> conditional_expectation N (indicator A)" |
|
545 |
||
546 |
lemma (in prob_space) |
|
547 |
fixes X :: "'a \<Rightarrow> pinfreal" |
|
548 |
assumes borel: "X \<in> borel_measurable M" |
|
549 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" |
|
550 |
shows borel_measurable_conditional_expectation: |
|
551 |
"conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)" |
|
552 |
and conditional_expectation: "\<And>C. C \<in> N \<Longrightarrow> |
|
553 |
positive_integral (\<lambda>x. conditional_expectation N X x * indicator C x) = |
|
554 |
positive_integral (\<lambda>x. X x * indicator C x)" |
|
555 |
(is "\<And>C. C \<in> N \<Longrightarrow> ?eq C") |
|
556 |
proof - |
|
557 |
note CE = conditional_expectation_exists[OF assms, unfolded Bex_def] |
|
558 |
then show "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)" |
|
559 |
unfolding conditional_expectation_def by (rule someI2_ex) blast |
|
560 |
||
561 |
from CE show "\<And>C. C\<in>N \<Longrightarrow> ?eq C" |
|
562 |
unfolding conditional_expectation_def by (rule someI2_ex) blast |
|
563 |
qed |
|
564 |
||
39091 | 565 |
lemma (in sigma_algebra) factorize_measurable_function: |
566 |
fixes Z :: "'a \<Rightarrow> pinfreal" and Y :: "'a \<Rightarrow> 'c" |
|
567 |
assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M" |
|
568 |
shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) |
|
569 |
\<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))" |
|
570 |
proof safe |
|
571 |
interpret M': sigma_algebra M' by fact |
|
572 |
have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto |
|
573 |
from M'.sigma_algebra_vimage[OF this] |
|
574 |
interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . |
|
575 |
||
576 |
{ fix g :: "'c \<Rightarrow> pinfreal" assume "g \<in> borel_measurable M'" |
|
577 |
with M'.measurable_vimage_algebra[OF Y] |
|
578 |
have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
579 |
by (rule measurable_comp) |
|
580 |
moreover assume "\<forall>x\<in>space M. Z x = g (Y x)" |
|
581 |
then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow> |
|
582 |
g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
583 |
by (auto intro!: measurable_cong) |
|
584 |
ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
585 |
by simp } |
|
586 |
||
587 |
assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" |
|
588 |
from va.borel_measurable_implies_simple_function_sequence[OF this] |
|
589 |
obtain f where f: "\<And>i. va.simple_function (f i)" and "f \<up> Z" by blast |
|
590 |
||
591 |
have "\<forall>i. \<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" |
|
592 |
proof |
|
593 |
fix i |
|
594 |
from f[of i] have "finite (f i`space M)" and B_ex: |
|
595 |
"\<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" |
|
596 |
unfolding va.simple_function_def by auto |
|
597 |
from B_ex[THEN bchoice] guess B .. note B = this |
|
598 |
||
599 |
let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x" |
|
600 |
||
601 |
show "\<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" |
|
602 |
proof (intro exI[of _ ?g] conjI ballI) |
|
603 |
show "M'.simple_function ?g" using B by auto |
|
604 |
||
605 |
fix x assume "x \<in> space M" |
|
606 |
then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pinfreal)" |
|
607 |
unfolding indicator_def using B by auto |
|
608 |
then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i] |
|
609 |
by (subst va.simple_function_indicator_representation) auto |
|
610 |
qed |
|
611 |
qed |
|
612 |
from choice[OF this] guess g .. note g = this |
|
613 |
||
614 |
show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)" |
|
615 |
proof (intro ballI bexI) |
|
616 |
show "(SUP i. g i) \<in> borel_measurable M'" |
|
617 |
using g by (auto intro: M'.borel_measurable_simple_function) |
|
618 |
fix x assume "x \<in> space M" |
|
619 |
have "Z x = (SUP i. f i) x" using `f \<up> Z` unfolding isoton_def by simp |
|
620 |
also have "\<dots> = (SUP i. g i) (Y x)" unfolding SUPR_fun_expand |
|
621 |
using g `x \<in> space M` by simp |
|
622 |
finally show "Z x = (SUP i. g i) (Y x)" . |
|
623 |
qed |
|
624 |
qed |
|
39090
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
hoelzl
parents:
39089
diff
changeset
|
625 |
|
35582 | 626 |
end |