author | wenzelm |
Fri, 18 Oct 2024 14:20:09 +0200 | |
changeset 81182 | fc5066122e68 |
parent 80768 | c7723cc15de8 |
child 81545 | 6f8a56a6b391 |
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 |
||
61808 | 6 |
section \<open>Probability measure\<close> |
42067 | 7 |
|
42148 | 8 |
theory Probability_Measure |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
64272
diff
changeset
|
9 |
imports "HOL-Analysis.Analysis" |
35582 | 10 |
begin |
11 |
||
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
12 |
locale prob_space = finite_measure + |
47694 | 13 |
assumes emeasure_space_1: "emeasure M (space M) = 1" |
38656 | 14 |
|
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
15 |
lemma prob_spaceI[Pure.intro!]: |
47694 | 16 |
assumes *: "emeasure M (space M) = 1" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
17 |
shows "prob_space M" |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
18 |
by (simp add: assms finite_measureI prob_space_axioms.intro prob_space_def) |
38656 | 19 |
|
59425 | 20 |
lemma prob_space_imp_sigma_finite: "prob_space M \<Longrightarrow> sigma_finite_measure M" |
21 |
unfolding prob_space_def finite_measure_def by simp |
|
22 |
||
40859 | 23 |
abbreviation (in prob_space) "events \<equiv> sets M" |
47694 | 24 |
abbreviation (in prob_space) "prob \<equiv> measure M" |
25 |
abbreviation (in prob_space) "random_variable M' X \<equiv> X \<in> measurable M M'" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
26 |
abbreviation (in prob_space) "expectation \<equiv> integral\<^sup>L M" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
27 |
abbreviation (in prob_space) "variance X \<equiv> integral\<^sup>L M (\<lambda>x. (X x - expectation X)\<^sup>2)" |
35582 | 28 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
29 |
lemma (in prob_space) finite_measure [simp]: "finite_measure M" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
30 |
by unfold_locales |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
31 |
|
47694 | 32 |
lemma (in prob_space) prob_space_distr: |
33 |
assumes f: "f \<in> measurable M M'" shows "prob_space (distr M M' f)" |
|
34 |
proof (rule prob_spaceI) |
|
35 |
have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space) |
|
36 |
with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1" |
|
37 |
by (auto simp: emeasure_distr emeasure_space_1) |
|
43339 | 38 |
qed |
39 |
||
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
40 |
lemma prob_space_distrD: |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
41 |
assumes f: "f \<in> measurable M N" and M: "prob_space (distr M N f)" shows "prob_space M" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
42 |
proof |
61605 | 43 |
interpret M: prob_space "distr M N f" by fact |
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
44 |
have "f -` space N \<inter> space M = space M" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
45 |
using f[THEN measurable_space] by auto |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
46 |
then show "emeasure M (space M) = 1" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
47 |
using M.emeasure_space_1 by (simp add: emeasure_distr[OF f]) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
48 |
qed |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
49 |
|
40859 | 50 |
lemma (in prob_space) prob_space: "prob (space M) = 1" |
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
51 |
by (simp add: emeasure_space_1 measure_eq_emeasure_eq_ennreal) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
52 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
53 |
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
|
54 |
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
|
55 |
|
47694 | 56 |
lemma (in prob_space) not_empty: "space M \<noteq> {}" |
57 |
using prob_space by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
58 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
59 |
lemma (in prob_space) emeasure_eq_1_AE: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
60 |
"S \<in> sets M \<Longrightarrow> AE x in M. x \<in> S \<Longrightarrow> emeasure M S = 1" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
61 |
by (subst emeasure_eq_AE[where B="space M"]) (auto simp: emeasure_space_1) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
62 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
63 |
lemma (in prob_space) emeasure_le_1: "emeasure M S \<le> 1" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
64 |
unfolding ennreal_1[symmetric] emeasure_eq_measure by (subst ennreal_le_iff) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
65 |
|
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63040
diff
changeset
|
66 |
lemma (in prob_space) emeasure_ge_1_iff: "emeasure M A \<ge> 1 \<longleftrightarrow> emeasure M A = 1" |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63040
diff
changeset
|
67 |
by (rule iffI, intro antisym emeasure_le_1) simp_all |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63040
diff
changeset
|
68 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
69 |
lemma (in prob_space) AE_iff_emeasure_eq_1: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
70 |
assumes [measurable]: "Measurable.pred M P" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
71 |
shows "(AE x in M. P x) \<longleftrightarrow> emeasure M {x\<in>space M. P x} = 1" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
72 |
proof - |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
73 |
have *: "{x \<in> space M. \<not> P x} = space M - {x\<in>space M. P x}" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
74 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
75 |
show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
76 |
by (auto simp add: ennreal_minus_eq_0 * emeasure_compl emeasure_space_1 AE_iff_measurable[OF _ refl] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
77 |
intro: antisym emeasure_le_1) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
78 |
qed |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
79 |
|
47694 | 80 |
lemma (in prob_space) measure_le_1: "emeasure M X \<le> 1" |
81 |
using emeasure_space[of M X] by (simp add: emeasure_space_1) |
|
42950
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents:
42902
diff
changeset
|
82 |
|
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63040
diff
changeset
|
83 |
lemma (in prob_space) measure_ge_1_iff: "measure M A \<ge> 1 \<longleftrightarrow> measure M A = 1" |
63626
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents:
63099
diff
changeset
|
84 |
by (auto intro!: antisym) |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63040
diff
changeset
|
85 |
|
43339 | 86 |
lemma (in prob_space) AE_I_eq_1: |
47694 | 87 |
assumes "emeasure M {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M" |
88 |
shows "AE x in M. P x" |
|
43339 | 89 |
proof (rule AE_I) |
47694 | 90 |
show "emeasure M (space M - {x \<in> space M. P x}) = 0" |
91 |
using assms emeasure_space_1 by (simp add: emeasure_compl) |
|
43339 | 92 |
qed (insert assms, auto) |
93 |
||
59000 | 94 |
lemma prob_space_restrict_space: |
95 |
"S \<in> sets M \<Longrightarrow> emeasure M S = 1 \<Longrightarrow> prob_space (restrict_space M S)" |
|
96 |
by (intro prob_spaceI) |
|
97 |
(simp add: emeasure_restrict_space space_restrict_space) |
|
98 |
||
40859 | 99 |
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
|
100 |
assumes A: "A \<in> events" |
38656 | 101 |
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
|
102 |
using finite_measure_compl[OF A] by (simp add: prob_space) |
35582 | 103 |
|
47694 | 104 |
lemma (in prob_space) AE_in_set_eq_1: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
105 |
assumes A[measurable]: "A \<in> events" shows "(AE x in M. x \<in> A) \<longleftrightarrow> prob A = 1" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
106 |
proof - |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
107 |
have *: "{x\<in>space M. x \<in> A} = A" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
108 |
using A[THEN sets.sets_into_space] by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
109 |
show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
110 |
by (subst AE_iff_emeasure_eq_1) (auto simp: emeasure_eq_measure *) |
47694 | 111 |
qed |
112 |
||
113 |
lemma (in prob_space) AE_False: "(AE x in M. False) \<longleftrightarrow> False" |
|
114 |
proof |
|
115 |
assume "AE x in M. False" |
|
116 |
then have "AE x in M. x \<in> {}" by simp |
|
117 |
then show False |
|
118 |
by (subst (asm) AE_in_set_eq_1) auto |
|
119 |
qed simp |
|
120 |
||
121 |
lemma (in prob_space) AE_prob_1: |
|
122 |
assumes "prob A = 1" shows "AE x in M. x \<in> A" |
|
123 |
proof - |
|
61808 | 124 |
from \<open>prob A = 1\<close> have "A \<in> events" |
47694 | 125 |
by (metis measure_notin_sets zero_neq_one) |
126 |
with AE_in_set_eq_1 assms show ?thesis by simp |
|
127 |
qed |
|
128 |
||
50098 | 129 |
lemma (in prob_space) AE_const[simp]: "(AE x in M. P) \<longleftrightarrow> P" |
130 |
by (cases P) (auto simp: AE_False) |
|
131 |
||
59000 | 132 |
lemma (in prob_space) ae_filter_bot: "ae_filter M \<noteq> bot" |
133 |
by (simp add: trivial_limit_def) |
|
134 |
||
50098 | 135 |
lemma (in prob_space) AE_contr: |
136 |
assumes ae: "AE \<omega> in M. P \<omega>" "AE \<omega> in M. \<not> P \<omega>" |
|
137 |
shows False |
|
138 |
proof - |
|
139 |
from ae have "AE \<omega> in M. False" by eventually_elim auto |
|
140 |
then show False by auto |
|
141 |
qed |
|
142 |
||
57025 | 143 |
lemma (in prob_space) integral_ge_const: |
144 |
fixes c :: real |
|
145 |
shows "integrable M f \<Longrightarrow> (AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>x. f x \<partial>M)" |
|
146 |
using integral_mono_AE[of M "\<lambda>x. c" f] prob_space by simp |
|
147 |
||
148 |
lemma (in prob_space) integral_le_const: |
|
149 |
fixes c :: real |
|
150 |
shows "integrable M f \<Longrightarrow> (AE x in M. f x \<le> c) \<Longrightarrow> (\<integral>x. f x \<partial>M) \<le> c" |
|
151 |
using integral_mono_AE[of M f "\<lambda>x. c"] prob_space by simp |
|
152 |
||
153 |
lemma (in prob_space) nn_integral_ge_const: |
|
154 |
"(AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>\<^sup>+x. f x \<partial>M)" |
|
155 |
using nn_integral_mono_AE[of "\<lambda>x. c" f M] emeasure_space_1 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
156 |
by (simp split: if_split_asm) |
57025 | 157 |
|
43339 | 158 |
lemma (in prob_space) expectation_less: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56166
diff
changeset
|
159 |
fixes X :: "_ \<Rightarrow> real" |
43339 | 160 |
assumes [simp]: "integrable M X" |
49786 | 161 |
assumes gt: "AE x in M. X x < b" |
43339 | 162 |
shows "expectation X < b" |
163 |
proof - |
|
164 |
have "expectation X < expectation (\<lambda>x. b)" |
|
47694 | 165 |
using gt emeasure_space_1 |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
43339
diff
changeset
|
166 |
by (intro integral_less_AE_space) auto |
43339 | 167 |
then show ?thesis using prob_space by simp |
168 |
qed |
|
169 |
||
170 |
lemma (in prob_space) expectation_greater: |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56166
diff
changeset
|
171 |
fixes X :: "_ \<Rightarrow> real" |
43339 | 172 |
assumes [simp]: "integrable M X" |
49786 | 173 |
assumes gt: "AE x in M. a < X x" |
43339 | 174 |
shows "a < expectation X" |
175 |
proof - |
|
176 |
have "expectation (\<lambda>x. a) < expectation X" |
|
47694 | 177 |
using gt emeasure_space_1 |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
43339
diff
changeset
|
178 |
by (intro integral_less_AE_space) auto |
43339 | 179 |
then show ?thesis using prob_space by simp |
180 |
qed |
|
181 |
||
182 |
lemma (in prob_space) jensens_inequality: |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56166
diff
changeset
|
183 |
fixes q :: "real \<Rightarrow> real" |
49786 | 184 |
assumes X: "integrable M X" "AE x in M. X x \<in> I" |
43339 | 185 |
assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV" |
186 |
assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q" |
|
187 |
shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" |
|
188 |
proof - |
|
46731 | 189 |
let ?F = "\<lambda>x. Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))" |
49786 | 190 |
from X(2) AE_False have "I \<noteq> {}" by auto |
43339 | 191 |
|
192 |
from I have "open I" by auto |
|
193 |
||
194 |
note I |
|
195 |
moreover |
|
196 |
{ assume "I \<subseteq> {a <..}" |
|
197 |
with X have "a < expectation X" |
|
198 |
by (intro expectation_greater) auto } |
|
199 |
moreover |
|
200 |
{ assume "I \<subseteq> {..< b}" |
|
201 |
with X have "expectation X < b" |
|
202 |
by (intro expectation_less) auto } |
|
203 |
ultimately have "expectation X \<in> I" |
|
204 |
by (elim disjE) (auto simp: subset_eq) |
|
205 |
moreover |
|
206 |
{ fix y assume y: "y \<in> I" |
|
61808 | 207 |
with q(2) \<open>open I\<close> have "Sup ((\<lambda>x. q x + ?F x * (y - x)) ` I) = q y" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62083
diff
changeset
|
208 |
by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open) } |
43339 | 209 |
ultimately have "q (expectation X) = Sup ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)" |
210 |
by simp |
|
211 |
also have "\<dots> \<le> expectation (\<lambda>w. q (X w))" |
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
50419
diff
changeset
|
212 |
proof (rule cSup_least) |
43339 | 213 |
show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}" |
61808 | 214 |
using \<open>I \<noteq> {}\<close> by auto |
43339 | 215 |
next |
216 |
fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I" |
|
74362 | 217 |
then obtain x |
218 |
where x: "k = q x + (INF t\<in>{x<..} \<inter> I. (q x - q t) / (x - t)) * (expectation X - x)" "x \<in> I" .. |
|
43339 | 219 |
have "q x + ?F x * (expectation X - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))" |
47694 | 220 |
using prob_space by (simp add: X) |
43339 | 221 |
also have "\<dots> \<le> expectation (\<lambda>w. q (X w))" |
61808 | 222 |
using \<open>x \<in> I\<close> \<open>open I\<close> X(2) |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
223 |
apply (intro integral_mono_AE Bochner_Integration.integrable_add Bochner_Integration.integrable_mult_right Bochner_Integration.integrable_diff |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56166
diff
changeset
|
224 |
integrable_const X q) |
61810 | 225 |
apply (elim eventually_mono) |
49786 | 226 |
apply (intro convex_le_Inf_differential) |
227 |
apply (auto simp: interior_open q) |
|
228 |
done |
|
43339 | 229 |
finally show "k \<le> expectation (\<lambda>w. q (X w))" using x by auto |
230 |
qed |
|
231 |
finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" . |
|
232 |
qed |
|
233 |
||
79945
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
234 |
lemma (in prob_space) finite_borel_measurable_integrable: |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
235 |
assumes "f\<in> borel_measurable M" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
236 |
and "finite (f`(space M))" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
237 |
shows "integrable M f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
238 |
proof - |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
239 |
have "simple_function M f" using assms by (simp add: simple_function_borel_measurable) |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
240 |
moreover have "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
241 |
ultimately have "Bochner_Integration.simple_bochner_integrable M f" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
242 |
using Bochner_Integration.simple_bochner_integrable.simps by blast |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
243 |
hence "has_bochner_integral M f (Bochner_Integration.simple_bochner_integral M f)" |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
244 |
using has_bochner_integral_simple_bochner_integrable by auto |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
245 |
thus ?thesis using integrable.simps by auto |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
246 |
qed |
ca004ccf2352
New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
247 |
|
61808 | 248 |
subsection \<open>Introduce binder for probability\<close> |
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
249 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
250 |
syntax |
69597 | 251 |
"_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (\<open>('\<P>'((/_ in _./ _)'))\<close>) |
80768 | 252 |
syntax_consts |
253 |
"_prob" == measure |
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
254 |
translations |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
255 |
"\<P>(x in M. P)" => "CONST measure M {x \<in> CONST space M. P}" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
256 |
|
61808 | 257 |
print_translation \<open> |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
258 |
let |
69597 | 259 |
fun to_pattern (Const (\<^const_syntax>\<open>Pair\<close>, _) $ l $ r) = |
260 |
Syntax.const \<^const_syntax>\<open>Pair\<close> :: to_pattern l @ to_pattern r |
|
261 |
| to_pattern (t as (Const (\<^syntax_const>\<open>_bound\<close>, _)) $ _) = [t] |
|
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
262 |
|
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
263 |
fun mk_pattern ((t, n) :: xs) = mk_patterns n xs |>> curry list_comb t |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
264 |
and mk_patterns 0 xs = ([], xs) |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
265 |
| mk_patterns n xs = |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
266 |
let |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
267 |
val (t, xs') = mk_pattern xs |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
268 |
val (ts, xs'') = mk_patterns (n - 1) xs' |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
269 |
in |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
270 |
(t :: ts, xs'') |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
271 |
end |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
272 |
|
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
273 |
fun unnest_tuples |
69597 | 274 |
(Const (\<^syntax_const>\<open>_pattern\<close>, _) $ |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
275 |
t1 $ |
69597 | 276 |
(t as (Const (\<^syntax_const>\<open>_pattern\<close>, _) $ _ $ _))) |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
277 |
= let |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
278 |
val (_ $ t2 $ t3) = unnest_tuples t |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
279 |
in |
69597 | 280 |
Syntax.const \<^syntax_const>\<open>_pattern\<close> $ |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
281 |
unnest_tuples t1 $ |
69597 | 282 |
(Syntax.const \<^syntax_const>\<open>_patterns\<close> $ t2 $ t3) |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
283 |
end |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
284 |
| unnest_tuples pat = pat |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
285 |
|
69597 | 286 |
fun tr' [sig_alg, Const (\<^const_syntax>\<open>Collect\<close>, _) $ t] = |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
287 |
let |
69597 | 288 |
val bound_dummyT = Const (\<^syntax_const>\<open>_bound\<close>, dummyT) |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
289 |
|
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
290 |
fun go pattern elem |
69597 | 291 |
(Const (\<^const_syntax>\<open>conj\<close>, _) $ |
292 |
(Const (\<^const_syntax>\<open>Set.member\<close>, _) $ elem' $ (Const (\<^const_syntax>\<open>space\<close>, _) $ sig_alg')) $ |
|
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
293 |
u) |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
294 |
= let |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
295 |
val _ = if sig_alg aconv sig_alg' andalso to_pattern elem' = rev elem then () else raise Match; |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
296 |
val (pat, rest) = mk_pattern (rev pattern); |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
297 |
val _ = case rest of [] => () | _ => raise Match |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
298 |
in |
69597 | 299 |
Syntax.const \<^syntax_const>\<open>_prob\<close> $ unnest_tuples pat $ sig_alg $ u |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
300 |
end |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
301 |
| go pattern elem (Abs abs) = |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
302 |
let |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
303 |
val (x as (_ $ tx), t) = Syntax_Trans.atomic_abs_tr' abs |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
304 |
in |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
305 |
go ((x, 0) :: pattern) (bound_dummyT $ tx :: elem) t |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
306 |
end |
69597 | 307 |
| go pattern elem (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ t) = |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
308 |
go |
69597 | 309 |
((Syntax.const \<^syntax_const>\<open>_pattern\<close>, 2) :: pattern) |
310 |
(Syntax.const \<^const_syntax>\<open>Pair\<close> :: elem) |
|
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
311 |
t |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
312 |
in |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
313 |
go [] [] t |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
314 |
end |
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
315 |
in |
69597 | 316 |
[(\<^const_syntax>\<open>Sigma_Algebra.measure\<close>, K tr')] |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
317 |
end |
61808 | 318 |
\<close> |
58764
ca2f59aef665
add print translation for probability notation \<P>
Andreas Lochbihler
parents:
57447
diff
changeset
|
319 |
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
320 |
definition |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
321 |
"cond_prob M P Q = \<P>(\<omega> in M. P \<omega> \<and> Q \<omega>) / \<P>(\<omega> in M. Q \<omega>)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
322 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
323 |
syntax |
69597 | 324 |
"_conditional_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (\<open>('\<P>'(_ in _. _ \<bar>/ _'))\<close>) |
80768 | 325 |
syntax_consts |
326 |
"_conditional_prob" == cond_prob |
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
327 |
translations |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
328 |
"\<P>(x in M. P \<bar> Q)" => "CONST cond_prob M (\<lambda>x. P) (\<lambda>x. Q)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
329 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
330 |
lemma (in prob_space) AE_E_prob: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
331 |
assumes ae: "AE x in M. P x" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
332 |
obtains S where "S \<subseteq> {x \<in> space M. P x}" "S \<in> events" "prob S = 1" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
333 |
proof - |
74362 | 334 |
from ae[THEN AE_E] obtain N |
335 |
where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> events" by auto |
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
336 |
then show thesis |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
337 |
by (intro that[of "space M - N"]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
338 |
(auto simp: prob_compl prob_space emeasure_eq_measure measure_nonneg) |
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
339 |
qed |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
340 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
341 |
lemma (in prob_space) prob_neg: "{x\<in>space M. P x} \<in> events \<Longrightarrow> \<P>(x in M. \<not> P x) = 1 - \<P>(x in M. P x)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
342 |
by (auto intro!: arg_cong[where f=prob] simp add: prob_compl[symmetric]) |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
343 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
344 |
lemma (in prob_space) prob_eq_AE: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
345 |
"(AE x in M. P x \<longleftrightarrow> Q x) \<Longrightarrow> {x\<in>space M. P x} \<in> events \<Longrightarrow> {x\<in>space M. Q x} \<in> events \<Longrightarrow> \<P>(x in M. P x) = \<P>(x in M. Q x)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
346 |
by (rule finite_measure_eq_AE) auto |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
347 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
348 |
lemma (in prob_space) prob_eq_0_AE: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
349 |
assumes not: "AE x in M. \<not> P x" shows "\<P>(x in M. P x) = 0" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
350 |
proof cases |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
351 |
assume "{x\<in>space M. P x} \<in> events" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
352 |
with not have "\<P>(x in M. P x) = \<P>(x in M. False)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
353 |
by (intro prob_eq_AE) auto |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
354 |
then show ?thesis by simp |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
355 |
qed (simp add: measure_notin_sets) |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
356 |
|
50098 | 357 |
lemma (in prob_space) prob_Collect_eq_0: |
358 |
"{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 0 \<longleftrightarrow> (AE x in M. \<not> P x)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
359 |
using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"] by (simp add: emeasure_eq_measure measure_nonneg) |
50098 | 360 |
|
361 |
lemma (in prob_space) prob_Collect_eq_1: |
|
362 |
"{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 1 \<longleftrightarrow> (AE x in M. P x)" |
|
363 |
using AE_in_set_eq_1[of "{x\<in>space M. P x}"] by simp |
|
364 |
||
365 |
lemma (in prob_space) prob_eq_0: |
|
366 |
"A \<in> sets M \<Longrightarrow> prob A = 0 \<longleftrightarrow> (AE x in M. x \<notin> A)" |
|
367 |
using AE_iff_measurable[OF _ refl, of M "\<lambda>x. x \<notin> A"] |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
368 |
by (auto simp add: emeasure_eq_measure Int_def[symmetric] measure_nonneg) |
50098 | 369 |
|
370 |
lemma (in prob_space) prob_eq_1: |
|
371 |
"A \<in> sets M \<Longrightarrow> prob A = 1 \<longleftrightarrow> (AE x in M. x \<in> A)" |
|
372 |
using AE_in_set_eq_1[of A] by simp |
|
373 |
||
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
374 |
lemma (in prob_space) prob_sums: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
375 |
assumes P: "\<And>n. {x\<in>space M. P n x} \<in> events" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
376 |
assumes Q: "{x\<in>space M. Q x} \<in> events" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
377 |
assumes ae: "AE x in M. (\<forall>n. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n. P n x))" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
378 |
shows "(\<lambda>n. \<P>(x in M. P n x)) sums \<P>(x in M. Q x)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
379 |
proof - |
74362 | 380 |
from ae[THEN AE_E_prob] obtain S |
381 |
where S: |
|
382 |
"S \<subseteq> {x \<in> space M. (\<forall>n. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n. P n x))}" |
|
383 |
"S \<in> events" |
|
384 |
"prob S = 1" |
|
385 |
by auto |
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
386 |
then have disj: "disjoint_family (\<lambda>n. {x\<in>space M. P n x} \<inter> S)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
387 |
by (auto simp: disjoint_family_on_def) |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
388 |
from S have ae_S: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
389 |
"AE x in M. x \<in> {x\<in>space M. Q x} \<longleftrightarrow> x \<in> (\<Union>n. {x\<in>space M. P n x} \<inter> S)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
390 |
"\<And>n. AE x in M. x \<in> {x\<in>space M. P n x} \<longleftrightarrow> x \<in> {x\<in>space M. P n x} \<inter> S" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
391 |
using ae by (auto dest!: AE_prob_1) |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
392 |
from ae_S have *: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
393 |
"\<P>(x in M. Q x) = prob (\<Union>n. {x\<in>space M. P n x} \<inter> S)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
394 |
using P Q S by (intro finite_measure_eq_AE) auto |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
395 |
from ae_S have **: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
396 |
"\<And>n. \<P>(x in M. P n x) = prob ({x\<in>space M. P n x} \<inter> S)" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
397 |
using P Q S by (intro finite_measure_eq_AE) auto |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
398 |
show ?thesis |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
399 |
unfolding * ** using S P disj |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
400 |
by (intro finite_measure_UNION) auto |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
401 |
qed |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
402 |
|
64267 | 403 |
lemma (in prob_space) prob_sum: |
59000 | 404 |
assumes [simp, intro]: "finite I" |
405 |
assumes P: "\<And>n. n \<in> I \<Longrightarrow> {x\<in>space M. P n x} \<in> events" |
|
406 |
assumes Q: "{x\<in>space M. Q x} \<in> events" |
|
407 |
assumes ae: "AE x in M. (\<forall>n\<in>I. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n\<in>I. P n x))" |
|
408 |
shows "\<P>(x in M. Q x) = (\<Sum>n\<in>I. \<P>(x in M. P n x))" |
|
409 |
proof - |
|
74362 | 410 |
from ae[THEN AE_E_prob] obtain S |
411 |
where S: |
|
412 |
"S \<subseteq> {x \<in> space M. (\<forall>n\<in>I. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n. n \<in> I \<and> P n x))}" |
|
413 |
"S \<in> events" |
|
414 |
"prob S = 1" |
|
415 |
by auto |
|
59000 | 416 |
then have disj: "disjoint_family_on (\<lambda>n. {x\<in>space M. P n x} \<inter> S) I" |
417 |
by (auto simp: disjoint_family_on_def) |
|
418 |
from S have ae_S: |
|
419 |
"AE x in M. x \<in> {x\<in>space M. Q x} \<longleftrightarrow> x \<in> (\<Union>n\<in>I. {x\<in>space M. P n x} \<inter> S)" |
|
420 |
"\<And>n. n \<in> I \<Longrightarrow> AE x in M. x \<in> {x\<in>space M. P n x} \<longleftrightarrow> x \<in> {x\<in>space M. P n x} \<inter> S" |
|
421 |
using ae by (auto dest!: AE_prob_1) |
|
422 |
from ae_S have *: |
|
423 |
"\<P>(x in M. Q x) = prob (\<Union>n\<in>I. {x\<in>space M. P n x} \<inter> S)" |
|
424 |
using P Q S by (intro finite_measure_eq_AE) (auto intro!: sets.Int) |
|
425 |
from ae_S have **: |
|
426 |
"\<And>n. n \<in> I \<Longrightarrow> \<P>(x in M. P n x) = prob ({x\<in>space M. P n x} \<inter> S)" |
|
427 |
using P Q S by (intro finite_measure_eq_AE) auto |
|
428 |
show ?thesis |
|
429 |
using S P disj |
|
430 |
by (auto simp add: * ** simp del: UN_simps intro!: finite_measure_finite_Union) |
|
431 |
qed |
|
432 |
||
54418 | 433 |
lemma (in prob_space) prob_EX_countable: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
434 |
assumes sets: "\<And>i. i \<in> I \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" and I: "countable I" |
54418 | 435 |
assumes disj: "AE x in M. \<forall>i\<in>I. \<forall>j\<in>I. P i x \<longrightarrow> P j x \<longrightarrow> i = j" |
436 |
shows "\<P>(x in M. \<exists>i\<in>I. P i x) = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)" |
|
437 |
proof - |
|
438 |
let ?N= "\<lambda>x. \<exists>!i\<in>I. P i x" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
439 |
have "ennreal (\<P>(x in M. \<exists>i\<in>I. P i x)) = \<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
440 |
unfolding ennreal_inj[OF measure_nonneg measure_nonneg] |
54418 | 441 |
proof (rule prob_eq_AE) |
442 |
show "AE x in M. (\<exists>i\<in>I. P i x) = (\<exists>i\<in>I. P i x \<and> ?N x)" |
|
443 |
using disj by eventually_elim blast |
|
444 |
qed (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets)+ |
|
445 |
also have "\<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x)) = emeasure M (\<Union>i\<in>I. {x\<in>space M. P i x \<and> ?N x})" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
446 |
unfolding emeasure_eq_measure by (auto intro!: arg_cong[where f=prob] simp: measure_nonneg) |
54418 | 447 |
also have "\<dots> = (\<integral>\<^sup>+i. emeasure M {x\<in>space M. P i x \<and> ?N x} \<partial>count_space I)" |
448 |
by (rule emeasure_UN_countable) |
|
449 |
(auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets |
|
450 |
simp: disjoint_family_on_def) |
|
451 |
also have "\<dots> = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)" |
|
452 |
unfolding emeasure_eq_measure using disj |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
453 |
by (intro nn_integral_cong ennreal_inj[THEN iffD2] prob_eq_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
454 |
(auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets measure_nonneg)+ |
54418 | 455 |
finally show ?thesis . |
456 |
qed |
|
457 |
||
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
458 |
lemma (in prob_space) cond_prob_eq_AE: |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
459 |
assumes P: "AE x in M. Q x \<longrightarrow> P x \<longleftrightarrow> P' x" "{x\<in>space M. P x} \<in> events" "{x\<in>space M. P' x} \<in> events" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
460 |
assumes Q: "AE x in M. Q x \<longleftrightarrow> Q' x" "{x\<in>space M. Q x} \<in> events" "{x\<in>space M. Q' x} \<in> events" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
461 |
shows "cond_prob M P Q = cond_prob M P' Q'" |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
462 |
using P Q |
67399 | 463 |
by (auto simp: cond_prob_def intro!: arg_cong2[where f="(/)"] prob_eq_AE sets.sets_Collect_conj) |
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
464 |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49795
diff
changeset
|
465 |
|
40859 | 466 |
lemma (in prob_space) joint_distribution_Times_le_fst: |
47694 | 467 |
"random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
468 |
\<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MX X) A" |
47694 | 469 |
by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) |
40859 | 470 |
|
471 |
lemma (in prob_space) joint_distribution_Times_le_snd: |
|
47694 | 472 |
"random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
473 |
\<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MY Y) B" |
47694 | 474 |
by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) |
40859 | 475 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
476 |
lemma (in prob_space) variance_eq: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
477 |
fixes X :: "'a \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
478 |
assumes [simp]: "integrable M X" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
479 |
assumes [simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
480 |
shows "variance X = expectation (\<lambda>x. (X x)\<^sup>2) - (expectation X)\<^sup>2" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
481 |
by (simp add: field_simps prob_space power2_diff power2_eq_square[symmetric]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
482 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
483 |
lemma (in prob_space) variance_positive: "0 \<le> variance (X::'a \<Rightarrow> real)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
484 |
by (intro integral_nonneg_AE) (auto intro!: integral_nonneg_AE) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
485 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
486 |
lemma (in prob_space) variance_mean_zero: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
487 |
"expectation X = 0 \<Longrightarrow> variance X = expectation (\<lambda>x. (X x)^2)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
488 |
by simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
489 |
|
73253
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
490 |
theorem%important (in prob_space) Chebyshev_inequality: |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
491 |
assumes [measurable]: "random_variable borel f" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
492 |
assumes "integrable M (\<lambda>x. f x ^ 2)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
493 |
defines "\<mu> \<equiv> expectation f" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
494 |
assumes "a > 0" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
495 |
shows "prob {x\<in>space M. \<bar>f x - \<mu>\<bar> \<ge> a} \<le> variance f / a\<^sup>2" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
496 |
unfolding \<mu>_def |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
497 |
proof (rule second_moment_method) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
498 |
have integrable: "integrable M f" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
499 |
using assms by (blast dest: square_integrable_imp_integrable) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
500 |
show "integrable M (\<lambda>x. (f x - expectation f)\<^sup>2)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
501 |
using assms integrable unfolding power2_eq_square ring_distribs |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
502 |
by (intro Bochner_Integration.integrable_diff) auto |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
503 |
qed (use assms in auto) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
504 |
|
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
505 |
locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2 |
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
|
506 |
|
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61424
diff
changeset
|
507 |
sublocale pair_prob_space \<subseteq> P?: prob_space "M1 \<Otimes>\<^sub>M M2" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
508 |
proof |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
509 |
show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) = 1" |
49776 | 510 |
by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure) |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
511 |
qed |
40859 | 512 |
|
47694 | 513 |
locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" + |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
514 |
fixes I :: "'i set" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
515 |
assumes prob_space: "\<And>i. prob_space (M i)" |
42988 | 516 |
|
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61424
diff
changeset
|
517 |
sublocale product_prob_space \<subseteq> M?: prob_space "M i" for i |
42988 | 518 |
by (rule prob_space) |
519 |
||
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
520 |
locale finite_product_prob_space = finite_product_sigma_finite M I + product_prob_space M I for M I |
42988 | 521 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
522 |
sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^sub>M i\<in>I. M i" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
523 |
proof |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
524 |
show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (space (\<Pi>\<^sub>M i\<in>I. M i)) = 1" |
64272 | 525 |
by (simp add: measure_times M.emeasure_space_1 prod.neutral_const space_PiM) |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
526 |
qed |
42988 | 527 |
|
528 |
lemma (in finite_product_prob_space) prob_times: |
|
529 |
assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
530 |
shows "prob (\<Pi>\<^sub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))" |
42988 | 531 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
532 |
have "ennreal (measure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)) = emeasure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)" |
47694 | 533 |
using X by (simp add: emeasure_eq_measure) |
534 |
also have "\<dots> = (\<Prod>i\<in>I. emeasure (M i) (X i))" |
|
42988 | 535 |
using measure_times X by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
536 |
also have "\<dots> = ennreal (\<Prod>i\<in>I. measure (M i) (X i))" |
64272 | 537 |
using X by (simp add: M.emeasure_eq_measure prod_ennreal measure_nonneg) |
538 |
finally show ?thesis by (simp add: measure_nonneg prod_nonneg) |
|
42859
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
539 |
qed |
d9dfc733f25c
add product of probability spaces with finite cardinality
hoelzl
parents:
42858
diff
changeset
|
540 |
|
73253
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
541 |
lemma product_prob_spaceI: |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
542 |
assumes "\<And>i. prob_space (M i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
543 |
shows "product_prob_space M" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
544 |
unfolding product_prob_space_def product_prob_space_axioms_def product_sigma_finite_def |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
545 |
proof safe |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
546 |
fix i |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
547 |
interpret prob_space "M i" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
548 |
by (rule assms) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
549 |
show "sigma_finite_measure (M i)" "prob_space (M i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
550 |
by unfold_locales |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
551 |
qed |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
552 |
|
61808 | 553 |
subsection \<open>Distributions\<close> |
42892 | 554 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
555 |
definition distributed :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> ennreal) \<Rightarrow> bool" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
556 |
where |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
557 |
"distributed M N X f \<longleftrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
558 |
distr M N X = density N f \<and> f \<in> borel_measurable N \<and> X \<in> measurable M N" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
559 |
|
47694 | 560 |
lemma |
50003 | 561 |
assumes "distributed M N X f" |
562 |
shows distributed_distr_eq_density: "distr M N X = density N f" |
|
563 |
and distributed_measurable: "X \<in> measurable M N" |
|
564 |
and distributed_borel_measurable: "f \<in> borel_measurable N" |
|
565 |
using assms by (simp_all add: distributed_def) |
|
566 |
||
567 |
lemma |
|
568 |
assumes D: "distributed M N X f" |
|
569 |
shows distributed_measurable'[measurable_dest]: |
|
570 |
"g \<in> measurable L M \<Longrightarrow> (\<lambda>x. X (g x)) \<in> measurable L N" |
|
571 |
and distributed_borel_measurable'[measurable_dest]: |
|
572 |
"h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L" |
|
573 |
using distributed_measurable[OF D] distributed_borel_measurable[OF D] |
|
574 |
by simp_all |
|
39097 | 575 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
576 |
lemma distributed_real_measurable: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
577 |
"(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow> f \<in> borel_measurable N" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
578 |
by (simp_all add: distributed_def) |
35977 | 579 |
|
59353
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
580 |
lemma distributed_real_measurable': |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
581 |
"(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
582 |
h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
583 |
using distributed_real_measurable[measurable] by simp |
50003 | 584 |
|
59353
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
585 |
lemma joint_distributed_measurable1: |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
586 |
"distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f \<Longrightarrow> h1 \<in> measurable N M \<Longrightarrow> (\<lambda>x. X (h1 x)) \<in> measurable N S" |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
587 |
by simp |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
588 |
|
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
589 |
lemma joint_distributed_measurable2: |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
590 |
"distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f \<Longrightarrow> h2 \<in> measurable N M \<Longrightarrow> (\<lambda>x. Y (h2 x)) \<in> measurable N T" |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59000
diff
changeset
|
591 |
by simp |
50003 | 592 |
|
47694 | 593 |
lemma distributed_count_space: |
594 |
assumes X: "distributed M (count_space A) X P" and a: "a \<in> A" and A: "finite A" |
|
595 |
shows "P a = emeasure M (X -` {a} \<inter> space M)" |
|
39097 | 596 |
proof - |
47694 | 597 |
have "emeasure M (X -` {a} \<inter> space M) = emeasure (distr M (count_space A) X) {a}" |
50003 | 598 |
using X a A by (simp add: emeasure_distr) |
47694 | 599 |
also have "\<dots> = emeasure (density (count_space A) P) {a}" |
600 |
using X by (simp add: distributed_distr_eq_density) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
601 |
also have "\<dots> = (\<integral>\<^sup>+x. P a * indicator {a} x \<partial>count_space A)" |
56996 | 602 |
using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: nn_integral_cong) |
47694 | 603 |
also have "\<dots> = P a" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
604 |
using X a by (subst nn_integral_cmult_indicator) (auto simp: distributed_def one_ennreal_def[symmetric] AE_count_space) |
47694 | 605 |
finally show ?thesis .. |
39092 | 606 |
qed |
35977 | 607 |
|
47694 | 608 |
lemma distributed_cong_density: |
609 |
"(AE x in N. f x = g x) \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable N \<Longrightarrow> |
|
610 |
distributed M N X f \<longleftrightarrow> distributed M N X g" |
|
611 |
by (auto simp: distributed_def intro!: density_cong) |
|
612 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
613 |
lemma (in prob_space) distributed_imp_emeasure_nonzero: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
614 |
assumes X: "distributed M MX X Px" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
615 |
shows "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
616 |
proof |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
617 |
note Px = distributed_borel_measurable[OF X] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
618 |
interpret X: prob_space "distr M MX X" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
619 |
using distributed_measurable[OF X] by (rule prob_space_distr) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
620 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
621 |
assume "emeasure MX {x \<in> space MX. Px x \<noteq> 0} = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
622 |
with Px have "AE x in MX. Px x = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
623 |
by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ennreal_iff) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
624 |
moreover |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
625 |
from X.emeasure_space_1 have "(\<integral>\<^sup>+x. Px x \<partial>MX) = 1" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
626 |
unfolding distributed_distr_eq_density[OF X] using Px |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
627 |
by (subst (asm) emeasure_density) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
628 |
(auto simp: borel_measurable_ennreal_iff intro!: integral_cong cong: nn_integral_cong) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
629 |
ultimately show False |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
630 |
by (simp add: nn_integral_cong_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
631 |
qed |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
632 |
|
47694 | 633 |
lemma subdensity: |
634 |
assumes T: "T \<in> measurable P Q" |
|
635 |
assumes f: "distributed M P X f" |
|
636 |
assumes g: "distributed M Q Y g" |
|
637 |
assumes Y: "Y = T \<circ> X" |
|
638 |
shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0" |
|
639 |
proof - |
|
640 |
have "{x\<in>space Q. g x = 0} \<in> null_sets (distr M Q (T \<circ> X))" |
|
641 |
using g Y by (auto simp: null_sets_density_iff distributed_def) |
|
642 |
also have "distr M Q (T \<circ> X) = distr (distr M P X) Q T" |
|
643 |
using T f[THEN distributed_measurable] by (rule distr_distr[symmetric]) |
|
644 |
finally have "T -` {x\<in>space Q. g x = 0} \<inter> space P \<in> null_sets (distr M P X)" |
|
645 |
using T by (subst (asm) null_sets_distr_iff) auto |
|
646 |
also have "T -` {x\<in>space Q. g x = 0} \<inter> space P = {x\<in>space P. g (T x) = 0}" |
|
647 |
using T by (auto dest: measurable_space) |
|
648 |
finally show ?thesis |
|
649 |
using f g by (auto simp add: null_sets_density_iff distributed_def) |
|
35977 | 650 |
qed |
651 |
||
47694 | 652 |
lemma subdensity_real: |
653 |
fixes g :: "'a \<Rightarrow> real" and f :: "'b \<Rightarrow> real" |
|
654 |
assumes T: "T \<in> measurable P Q" |
|
655 |
assumes f: "distributed M P X f" |
|
656 |
assumes g: "distributed M Q Y g" |
|
657 |
assumes Y: "Y = T \<circ> X" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
658 |
shows "(AE x in P. 0 \<le> g (T x)) \<Longrightarrow> (AE x in P. 0 \<le> f x) \<Longrightarrow> AE x in P. g (T x) = 0 \<longrightarrow> f x = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
659 |
using subdensity[OF T, of M X "\<lambda>x. ennreal (f x)" Y "\<lambda>x. ennreal (g x)"] assms |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
660 |
by auto |
47694 | 661 |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
662 |
lemma distributed_emeasure: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
663 |
"distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>N)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
664 |
by (auto simp: distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
665 |
|
56996 | 666 |
lemma distributed_nn_integral: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
667 |
"distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f x * g x \<partial>N) = (\<integral>\<^sup>+x. g (X x) \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
668 |
by (auto simp: distributed_distr_eq_density[symmetric] nn_integral_density[symmetric] nn_integral_distr) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
669 |
|
47694 | 670 |
lemma distributed_integral: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
671 |
"distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
672 |
(\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
673 |
supply distributed_real_measurable[measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
674 |
by (auto simp: distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
675 |
|
47694 | 676 |
lemma distributed_transform_integral: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
677 |
assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
678 |
assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x" |
47694 | 679 |
assumes Y: "Y = T \<circ> X" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P" |
680 |
shows "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. Px x * f (T x) \<partial>N)" |
|
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
|
681 |
proof - |
47694 | 682 |
have "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. f (Y x) \<partial>M)" |
683 |
by (rule distributed_integral) fact+ |
|
684 |
also have "\<dots> = (\<integral>x. f (T (X x)) \<partial>M)" |
|
685 |
using Y by simp |
|
686 |
also have "\<dots> = (\<integral>x. Px x * f (T x) \<partial>N)" |
|
687 |
using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def) |
|
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45712
diff
changeset
|
688 |
finally show ?thesis . |
39092 | 689 |
qed |
36624 | 690 |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
691 |
lemma (in prob_space) distributed_unique: |
47694 | 692 |
assumes Px: "distributed M S X Px" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
693 |
assumes Py: "distributed M S X Py" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
694 |
shows "AE x in S. Px x = Py x" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
695 |
proof - |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
696 |
interpret X: prob_space "distr M S X" |
50003 | 697 |
using Px by (intro prob_space_distr) simp |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
698 |
have "sigma_finite_measure (distr M S X)" .. |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
699 |
with sigma_finite_density_unique[of Px S Py ] Px Py |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
700 |
show ?thesis |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
701 |
by (auto simp: distributed_def) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
702 |
qed |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
703 |
|
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
704 |
lemma (in prob_space) distributed_jointI: |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
705 |
assumes "sigma_finite_measure S" "sigma_finite_measure T" |
50003 | 706 |
assumes X[measurable]: "X \<in> measurable M S" and Y[measurable]: "Y \<in> measurable M T" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
707 |
assumes [measurable]: "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)" and f: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> f x" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
708 |
assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
709 |
emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
710 |
shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
711 |
unfolding distributed_def |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
712 |
proof safe |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
713 |
interpret S: sigma_finite_measure S by fact |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
714 |
interpret T: sigma_finite_measure T by fact |
61169 | 715 |
interpret ST: pair_sigma_finite S T .. |
47694 | 716 |
|
74362 | 717 |
from ST.sigma_finite_up_in_pair_measure_generator |
718 |
obtain F :: "nat \<Rightarrow> ('b \<times> 'c) set" |
|
719 |
where F: "range F \<subseteq> {A \<times> B |A B. A \<in> sets S \<and> B \<in> sets T} \<and> incseq F \<and> |
|
720 |
\<Union> (range F) = space S \<times> space T \<and> (\<forall>i. emeasure (S \<Otimes>\<^sub>M T) (F i) \<noteq> \<infinity>)" .. |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
721 |
let ?E = "{a \<times> b |a b. a \<in> sets S \<and> b \<in> sets T}" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
722 |
let ?P = "S \<Otimes>\<^sub>M T" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
723 |
show "distr M ?P (\<lambda>x. (X x, Y x)) = density ?P f" (is "?L = ?R") |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
724 |
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]]) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
725 |
show "?E \<subseteq> Pow (space ?P)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50104
diff
changeset
|
726 |
using sets.space_closed[of S] sets.space_closed[of T] by (auto simp: space_pair_measure) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
727 |
show "sets ?L = sigma_sets (space ?P) ?E" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
728 |
by (simp add: sets_pair_measure space_pair_measure) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
729 |
then show "sets ?R = sigma_sets (space ?P) ?E" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
730 |
by simp |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
731 |
next |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
732 |
interpret L: prob_space ?L |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
733 |
by (rule prob_space_distr) (auto intro!: measurable_Pair) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
734 |
show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?L (F i) \<noteq> \<infinity>" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
735 |
using F by (auto simp: space_pair_measure) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
736 |
next |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
737 |
fix E assume "E \<in> ?E" |
50003 | 738 |
then obtain A B where E[simp]: "E = A \<times> B" |
739 |
and A[measurable]: "A \<in> sets S" and B[measurable]: "B \<in> sets T" by auto |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
740 |
have "emeasure ?L E = emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B}" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
741 |
by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
742 |
also have "\<dots> = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. (f (x, y) * indicator B y) * indicator A x \<partial>T) \<partial>S)" |
56996 | 743 |
using f by (auto simp add: eq nn_integral_multc intro!: nn_integral_cong) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
744 |
also have "\<dots> = emeasure ?R E" |
56996 | 745 |
by (auto simp add: emeasure_density T.nn_integral_fst[symmetric] |
746 |
intro!: nn_integral_cong split: split_indicator) |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
747 |
finally show "emeasure ?L E = emeasure ?R E" . |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
748 |
qed |
50003 | 749 |
qed (auto simp: f) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
750 |
|
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
751 |
lemma (in prob_space) distributed_swap: |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
752 |
assumes "sigma_finite_measure S" "sigma_finite_measure T" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
753 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
754 |
shows "distributed M (T \<Otimes>\<^sub>M S) (\<lambda>x. (Y x, X x)) (\<lambda>(x, y). Pxy (y, x))" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
755 |
proof - |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
756 |
interpret S: sigma_finite_measure S by fact |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
757 |
interpret T: sigma_finite_measure T by fact |
61169 | 758 |
interpret ST: pair_sigma_finite S T .. |
759 |
interpret TS: pair_sigma_finite T S .. |
|
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
760 |
|
50003 | 761 |
note Pxy[measurable] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
762 |
show ?thesis |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
763 |
apply (subst TS.distr_pair_swap) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
764 |
unfolding distributed_def |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
765 |
proof safe |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
766 |
let ?D = "distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
767 |
show 1: "(\<lambda>(x, y). Pxy (y, x)) \<in> borel_measurable ?D" |
50003 | 768 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
769 |
show 2: "random_variable (distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))" |
50003 | 770 |
using Pxy by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
771 |
{ fix A assume A: "A \<in> sets (T \<Otimes>\<^sub>M S)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
772 |
let ?B = "(\<lambda>(x, y). (y, x)) -` A \<inter> space (S \<Otimes>\<^sub>M T)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50104
diff
changeset
|
773 |
from sets.sets_into_space[OF A] |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
774 |
have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) = |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
775 |
emeasure M ((\<lambda>x. (X x, Y x)) -` ?B \<inter> space M)" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
776 |
by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
777 |
also have "\<dots> = (\<integral>\<^sup>+ x. Pxy x * indicator ?B x \<partial>(S \<Otimes>\<^sub>M T))" |
50003 | 778 |
using Pxy A by (intro distributed_emeasure) auto |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
779 |
finally have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
780 |
(\<integral>\<^sup>+ x. Pxy x * indicator A (snd x, fst x) \<partial>(S \<Otimes>\<^sub>M T))" |
56996 | 781 |
by (auto intro!: nn_integral_cong split: split_indicator) } |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
782 |
note * = this |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
783 |
show "distr M ?D (\<lambda>x. (Y x, X x)) = density ?D (\<lambda>(x, y). Pxy (y, x))" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
784 |
apply (intro measure_eqI) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
785 |
apply (simp_all add: emeasure_distr[OF 2] emeasure_density[OF 1]) |
56996 | 786 |
apply (subst nn_integral_distr) |
50003 | 787 |
apply (auto intro!: * simp: comp_def split_beta) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
788 |
done |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
789 |
qed |
36624 | 790 |
qed |
791 |
||
47694 | 792 |
lemma (in prob_space) distr_marginal1: |
793 |
assumes "sigma_finite_measure S" "sigma_finite_measure T" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
794 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
795 |
defines "Px \<equiv> \<lambda>x. (\<integral>\<^sup>+z. Pxy (x, z) \<partial>T)" |
47694 | 796 |
shows "distributed M S X Px" |
797 |
unfolding distributed_def |
|
798 |
proof safe |
|
799 |
interpret S: sigma_finite_measure S by fact |
|
800 |
interpret T: sigma_finite_measure T by fact |
|
61169 | 801 |
interpret ST: pair_sigma_finite S T .. |
47694 | 802 |
|
50003 | 803 |
note Pxy[measurable] |
804 |
show X: "X \<in> measurable M S" by simp |
|
47694 | 805 |
|
50003 | 806 |
show borel: "Px \<in> borel_measurable S" |
56996 | 807 |
by (auto intro!: T.nn_integral_fst simp: Px_def) |
39097 | 808 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
809 |
interpret Pxy: prob_space "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
50003 | 810 |
by (intro prob_space_distr) simp |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
811 |
|
47694 | 812 |
show "distr M S X = density S Px" |
813 |
proof (rule measure_eqI) |
|
814 |
fix A assume A: "A \<in> sets (distr M S X)" |
|
50003 | 815 |
with X measurable_space[of Y M T] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
816 |
have "emeasure (distr M S X) A = emeasure (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (A \<times> space T)" |
50003 | 817 |
by (auto simp add: emeasure_distr intro!: arg_cong[where f="emeasure M"]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
818 |
also have "\<dots> = emeasure (density (S \<Otimes>\<^sub>M T) Pxy) (A \<times> space T)" |
47694 | 819 |
using Pxy by (simp add: distributed_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
820 |
also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S" |
47694 | 821 |
using A borel Pxy |
56996 | 822 |
by (simp add: emeasure_density T.nn_integral_fst[symmetric]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
823 |
also have "\<dots> = \<integral>\<^sup>+ x. Px x * indicator A x \<partial>S" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
824 |
proof (rule nn_integral_cong) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
825 |
fix x assume "x \<in> space S" |
47694 | 826 |
moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x" |
827 |
by (auto simp: indicator_def) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
828 |
ultimately have "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) * indicator A x" |
56996 | 829 |
by (simp add: eq nn_integral_multc cong: nn_integral_cong) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
830 |
also have "(\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) = Px x" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
831 |
by (simp add: Px_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
832 |
finally show "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" . |
47694 | 833 |
qed |
834 |
finally show "emeasure (distr M S X) A = emeasure (density S Px) A" |
|
835 |
using A borel Pxy by (simp add: emeasure_density) |
|
836 |
qed simp |
|
40859 | 837 |
qed |
838 |
||
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
839 |
lemma (in prob_space) distr_marginal2: |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
840 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
841 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
842 |
shows "distributed M T Y (\<lambda>y. (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S))" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
843 |
using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
844 |
|
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
845 |
lemma (in prob_space) distributed_marginal_eq_joint1: |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
846 |
assumes T: "sigma_finite_measure T" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
847 |
assumes S: "sigma_finite_measure S" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
848 |
assumes Px: "distributed M S X Px" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
849 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
850 |
shows "AE x in S. Px x = (\<integral>\<^sup>+y. Pxy (x, y) \<partial>T)" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
851 |
using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
852 |
|
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
853 |
lemma (in prob_space) distributed_marginal_eq_joint2: |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
854 |
assumes T: "sigma_finite_measure T" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
855 |
assumes S: "sigma_finite_measure S" |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
856 |
assumes Py: "distributed M T Y Py" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
857 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
858 |
shows "AE y in T. Py y = (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S)" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
859 |
using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49786
diff
changeset
|
860 |
|
49795 | 861 |
lemma (in prob_space) distributed_joint_indep': |
862 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
|
50003 | 863 |
assumes X[measurable]: "distributed M S X Px" and Y[measurable]: "distributed M T Y Py" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
864 |
assumes indep: "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
865 |
shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)" |
49795 | 866 |
unfolding distributed_def |
867 |
proof safe |
|
868 |
interpret S: sigma_finite_measure S by fact |
|
869 |
interpret T: sigma_finite_measure T by fact |
|
61169 | 870 |
interpret ST: pair_sigma_finite S T .. |
49795 | 871 |
|
872 |
interpret X: prob_space "density S Px" |
|
873 |
unfolding distributed_distr_eq_density[OF X, symmetric] |
|
50003 | 874 |
by (rule prob_space_distr) simp |
49795 | 875 |
have sf_X: "sigma_finite_measure (density S Px)" .. |
876 |
||
877 |
interpret Y: prob_space "density T Py" |
|
878 |
unfolding distributed_distr_eq_density[OF Y, symmetric] |
|
50003 | 879 |
by (rule prob_space_distr) simp |
49795 | 880 |
have sf_Y: "sigma_finite_measure (density T Py)" .. |
881 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
882 |
show "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). Px x * Py y)" |
49795 | 883 |
unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
884 |
using distributed_borel_measurable[OF X] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
885 |
using distributed_borel_measurable[OF Y] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
886 |
by (rule pair_measure_density[OF _ _ T sf_Y]) |
49795 | 887 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
888 |
show "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" by auto |
49795 | 889 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
890 |
show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^sub>M T)" by auto |
49795 | 891 |
qed |
892 |
||
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
893 |
lemma distributed_integrable: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
894 |
"distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
895 |
integrable N (\<lambda>x. f x * g x) \<longleftrightarrow> integrable M (\<lambda>x. g (X x))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
896 |
supply distributed_real_measurable[measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
897 |
by (auto simp: distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
898 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
899 |
lemma distributed_transform_integrable: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
900 |
assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
901 |
assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
902 |
assumes Y: "Y = (\<lambda>x. T (X x))" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
903 |
shows "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
904 |
proof - |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
905 |
have "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable M (\<lambda>x. f (Y x))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
906 |
by (rule distributed_integrable) fact+ |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
907 |
also have "\<dots> \<longleftrightarrow> integrable M (\<lambda>x. f (T (X x)))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
908 |
using Y by simp |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
909 |
also have "\<dots> \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
910 |
using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
911 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
912 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
913 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
914 |
lemma distributed_integrable_var: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
915 |
fixes X :: "'a \<Rightarrow> real" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
916 |
shows "distributed M lborel X (\<lambda>x. ennreal (f x)) \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
917 |
integrable lborel (\<lambda>x. f x * x) \<Longrightarrow> integrable M X" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
918 |
using distributed_integrable[of M lborel X f "\<lambda>x. x"] by simp |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
919 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
920 |
lemma (in prob_space) distributed_variance: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
921 |
fixes f::"real \<Rightarrow> real" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
922 |
assumes D: "distributed M lborel X f" and [simp]: "\<And>x. 0 \<le> f x" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
923 |
shows "variance X = (\<integral>x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
924 |
proof (subst distributed_integral[OF D, symmetric]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
925 |
show "(\<integral> x. f x * (x - expectation X)\<^sup>2 \<partial>lborel) = (\<integral> x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
926 |
by (subst lborel_integral_real_affine[where c=1 and t="expectation X"]) (auto simp: ac_simps) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
927 |
qed simp_all |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
928 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
929 |
lemma (in prob_space) variance_affine: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
930 |
fixes f::"real \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
931 |
assumes [arith]: "b \<noteq> 0" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
932 |
assumes D[intro]: "distributed M lborel X f" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
933 |
assumes [simp]: "prob_space (density lborel f)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
934 |
assumes I[simp]: "integrable M X" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
935 |
assumes I2[simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
936 |
shows "variance (\<lambda>x. a + b * X x) = b\<^sup>2 * variance X" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
937 |
by (subst variance_eq) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
938 |
(auto simp: power2_sum power_mult_distrib prob_space variance_eq right_diff_distrib) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57025
diff
changeset
|
939 |
|
47694 | 940 |
definition |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
941 |
"simple_distributed M X f \<longleftrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
942 |
(\<forall>x. 0 \<le> f x) \<and> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
943 |
distributed M (count_space (X`space M)) X (\<lambda>x. ennreal (f x)) \<and> |
47694 | 944 |
finite (X`space M)" |
42902 | 945 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
946 |
lemma simple_distributed_nonneg[dest]: "simple_distributed M X f \<Longrightarrow> 0 \<le> f x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
947 |
by (auto simp: simple_distributed_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
948 |
|
47694 | 949 |
lemma simple_distributed: |
950 |
"simple_distributed M X Px \<Longrightarrow> distributed M (count_space (X`space M)) X Px" |
|
951 |
unfolding simple_distributed_def by auto |
|
42902 | 952 |
|
47694 | 953 |
lemma simple_distributed_finite[dest]: "simple_distributed M X P \<Longrightarrow> finite (X`space M)" |
954 |
by (simp add: simple_distributed_def) |
|
42902 | 955 |
|
47694 | 956 |
lemma (in prob_space) distributed_simple_function_superset: |
957 |
assumes X: "simple_function M X" "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)" |
|
958 |
assumes A: "X`space M \<subseteq> A" "finite A" |
|
959 |
defines "S \<equiv> count_space A" and "P' \<equiv> (\<lambda>x. if x \<in> X`space M then P x else 0)" |
|
960 |
shows "distributed M S X P'" |
|
961 |
unfolding distributed_def |
|
962 |
proof safe |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
963 |
show "(\<lambda>x. ennreal (P' x)) \<in> borel_measurable S" unfolding S_def by simp |
47694 | 964 |
show "distr M S X = density S P'" |
965 |
proof (rule measure_eqI_finite) |
|
966 |
show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A" |
|
967 |
using A unfolding S_def by auto |
|
968 |
show "finite A" by fact |
|
969 |
fix a assume a: "a \<in> A" |
|
970 |
then have "a \<notin> X`space M \<Longrightarrow> X -` {a} \<inter> space M = {}" by auto |
|
971 |
with A a X have "emeasure (distr M S X) {a} = P' a" |
|
972 |
by (subst emeasure_distr) |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
973 |
(auto simp add: S_def P'_def simple_functionD emeasure_eq_measure measurable_count_space_eq2 |
47694 | 974 |
intro!: arg_cong[where f=prob]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
975 |
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' a) * indicator {a} x \<partial>S)" |
47694 | 976 |
using A X a |
56996 | 977 |
by (subst nn_integral_cmult_indicator) |
47694 | 978 |
(auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
979 |
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' x) * indicator {a} x \<partial>S)" |
56996 | 980 |
by (auto simp: indicator_def intro!: nn_integral_cong) |
47694 | 981 |
also have "\<dots> = emeasure (density S P') {a}" |
982 |
using a A by (intro emeasure_density[symmetric]) (auto simp: S_def) |
|
983 |
finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" . |
|
984 |
qed |
|
985 |
show "random_variable S X" |
|
986 |
using X(1) A by (auto simp: measurable_def simple_functionD S_def) |
|
987 |
qed |
|
42902 | 988 |
|
47694 | 989 |
lemma (in prob_space) simple_distributedI: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
990 |
assumes X: "simple_function M X" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
991 |
"\<And>x. 0 \<le> P x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
992 |
"\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)" |
47694 | 993 |
shows "simple_distributed M X P" |
994 |
unfolding simple_distributed_def |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
995 |
proof (safe intro!: X) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
996 |
have "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (if x \<in> X`space M then P x else 0))" |
47694 | 997 |
(is "?A") |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
998 |
using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X(1,3)]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
999 |
also have "?A \<longleftrightarrow> distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (P x))" |
47694 | 1000 |
by (rule distributed_cong_density) auto |
1001 |
finally show "\<dots>" . |
|
1002 |
qed (rule simple_functionD[OF X(1)]) |
|
1003 |
||
1004 |
lemma simple_distributed_joint_finite: |
|
1005 |
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px" |
|
1006 |
shows "finite (X ` space M)" "finite (Y ` space M)" |
|
42902 | 1007 |
proof - |
47694 | 1008 |
have "finite ((\<lambda>x. (X x, Y x)) ` space M)" |
1009 |
using X by (auto simp: simple_distributed_def simple_functionD) |
|
1010 |
then have "finite (fst ` (\<lambda>x. (X x, Y x)) ` space M)" "finite (snd ` (\<lambda>x. (X x, Y x)) ` space M)" |
|
1011 |
by auto |
|
1012 |
then show fin: "finite (X ` space M)" "finite (Y ` space M)" |
|
1013 |
by (auto simp: image_image) |
|
1014 |
qed |
|
1015 |
||
1016 |
lemma simple_distributed_joint2_finite: |
|
1017 |
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px" |
|
1018 |
shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" |
|
1019 |
proof - |
|
1020 |
have "finite ((\<lambda>x. (X x, Y x, Z x)) ` space M)" |
|
1021 |
using X by (auto simp: simple_distributed_def simple_functionD) |
|
1022 |
then have "finite (fst ` (\<lambda>x. (X x, Y x, Z x)) ` space M)" |
|
1023 |
"finite ((fst \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)" |
|
1024 |
"finite ((snd \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)" |
|
1025 |
by auto |
|
1026 |
then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" |
|
1027 |
by (auto simp: image_image) |
|
42902 | 1028 |
qed |
1029 |
||
47694 | 1030 |
lemma simple_distributed_simple_function: |
1031 |
"simple_distributed M X Px \<Longrightarrow> simple_function M X" |
|
1032 |
unfolding simple_distributed_def distributed_def |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1033 |
by (auto simp: simple_function_def measurable_count_space_eq2) |
47694 | 1034 |
|
1035 |
lemma simple_distributed_measure: |
|
1036 |
"simple_distributed M X P \<Longrightarrow> a \<in> X`space M \<Longrightarrow> P a = measure M (X -` {a} \<inter> space M)" |
|
1037 |
using distributed_count_space[of M "X`space M" X P a, symmetric] |
|
1038 |
by (auto simp: simple_distributed_def measure_def) |
|
1039 |
||
1040 |
lemma (in prob_space) simple_distributed_joint: |
|
1041 |
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
1042 |
defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)" |
47694 | 1043 |
defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then Px x else 0)" |
1044 |
shows "distributed M S (\<lambda>x. (X x, Y x)) P" |
|
1045 |
proof - |
|
1046 |
from simple_distributed_joint_finite[OF X, simp] |
|
1047 |
have S_eq: "S = count_space (X`space M \<times> Y`space M)" |
|
1048 |
by (simp add: S_def pair_measure_count_space) |
|
1049 |
show ?thesis |
|
1050 |
unfolding S_eq P_def |
|
1051 |
proof (rule distributed_simple_function_superset) |
|
1052 |
show "simple_function M (\<lambda>x. (X x, Y x))" |
|
1053 |
using X by (rule simple_distributed_simple_function) |
|
1054 |
fix x assume "x \<in> (\<lambda>x. (X x, Y x)) ` space M" |
|
1055 |
from simple_distributed_measure[OF X this] |
|
1056 |
show "Px x = prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)" . |
|
1057 |
qed auto |
|
1058 |
qed |
|
42860 | 1059 |
|
47694 | 1060 |
lemma (in prob_space) simple_distributed_joint2: |
1061 |
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
1062 |
defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M) \<Otimes>\<^sub>M count_space (Z`space M)" |
47694 | 1063 |
defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then Px x else 0)" |
1064 |
shows "distributed M S (\<lambda>x. (X x, Y x, Z x)) P" |
|
1065 |
proof - |
|
1066 |
from simple_distributed_joint2_finite[OF X, simp] |
|
1067 |
have S_eq: "S = count_space (X`space M \<times> Y`space M \<times> Z`space M)" |
|
1068 |
by (simp add: S_def pair_measure_count_space) |
|
1069 |
show ?thesis |
|
1070 |
unfolding S_eq P_def |
|
1071 |
proof (rule distributed_simple_function_superset) |
|
1072 |
show "simple_function M (\<lambda>x. (X x, Y x, Z x))" |
|
1073 |
using X by (rule simple_distributed_simple_function) |
|
1074 |
fix x assume "x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M" |
|
1075 |
from simple_distributed_measure[OF X this] |
|
1076 |
show "Px x = prob ((\<lambda>x. (X x, Y x, Z x)) -` {x} \<inter> space M)" . |
|
1077 |
qed auto |
|
1078 |
qed |
|
1079 |
||
64267 | 1080 |
lemma (in prob_space) simple_distributed_sum_space: |
47694 | 1081 |
assumes X: "simple_distributed M X f" |
64267 | 1082 |
shows "sum f (X`space M) = 1" |
47694 | 1083 |
proof - |
64267 | 1084 |
from X have "sum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)" |
47694 | 1085 |
by (subst finite_measure_finite_Union) |
1086 |
(auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD |
|
64267 | 1087 |
intro!: sum.cong arg_cong[where f="prob"]) |
47694 | 1088 |
also have "\<dots> = prob (space M)" |
1089 |
by (auto intro!: arg_cong[where f=prob]) |
|
1090 |
finally show ?thesis |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1091 |
using emeasure_space_1 by (simp add: emeasure_eq_measure) |
47694 | 1092 |
qed |
42860 | 1093 |
|
47694 | 1094 |
lemma (in prob_space) distributed_marginal_eq_joint_simple: |
1095 |
assumes Px: "simple_function M X" |
|
1096 |
assumes Py: "simple_distributed M Y Py" |
|
1097 |
assumes Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy" |
|
1098 |
assumes y: "y \<in> Y`space M" |
|
1099 |
shows "Py y = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" |
|
1100 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1101 |
note Px = simple_distributedI[OF Px measure_nonneg refl] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1102 |
have "AE y in count_space (Y ` space M). ennreal (Py y) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1103 |
\<integral>\<^sup>+ x. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0) \<partial>count_space (X ` space M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1104 |
using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1105 |
simple_distributed[OF Py] simple_distributed_joint[OF Pxy] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1106 |
by (rule distributed_marginal_eq_joint2) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1107 |
(auto intro: Py Px simple_distributed_finite) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1108 |
then have "ennreal (Py y) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1109 |
(\<Sum>x\<in>X`space M. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1110 |
using y Px[THEN simple_distributed_finite] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1111 |
by (auto simp: AE_count_space nn_integral_count_space_finite) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1112 |
also have "\<dots> = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" |
64267 | 1113 |
using Pxy by (intro sum_ennreal) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1114 |
finally show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1115 |
using simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy] |
64267 | 1116 |
by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg) |
47694 | 1117 |
qed |
42860 | 1118 |
|
50419 | 1119 |
lemma distributedI_real: |
1120 |
fixes f :: "'a \<Rightarrow> real" |
|
1121 |
assumes gen: "sets M1 = sigma_sets (space M1) E" and "Int_stable E" |
|
1122 |
and A: "range A \<subseteq> E" "(\<Union>i::nat. A i) = space M1" "\<And>i. emeasure (distr M M1 X) (A i) \<noteq> \<infinity>" |
|
1123 |
and X: "X \<in> measurable M M1" |
|
1124 |
and f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
1125 |
and eq: "\<And>A. A \<in> E \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M1)" |
50419 | 1126 |
shows "distributed M M1 X f" |
1127 |
unfolding distributed_def |
|
1128 |
proof (intro conjI) |
|
1129 |
show "distr M M1 X = density M1 f" |
|
1130 |
proof (rule measure_eqI_generator_eq[where A=A]) |
|
1131 |
{ fix A assume A: "A \<in> E" |
|
1132 |
then have "A \<in> sigma_sets (space M1) E" by auto |
|
1133 |
then have "A \<in> sets M1" |
|
1134 |
using gen by simp |
|
1135 |
with f A eq[of A] X show "emeasure (distr M M1 X) A = emeasure (density M1 f) A" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1136 |
by (auto simp add: emeasure_distr emeasure_density ennreal_indicator |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1137 |
intro!: nn_integral_cong split: split_indicator) } |
50419 | 1138 |
note eq_E = this |
1139 |
show "Int_stable E" by fact |
|
1140 |
{ fix e assume "e \<in> E" |
|
1141 |
then have "e \<in> sigma_sets (space M1) E" by auto |
|
1142 |
then have "e \<in> sets M1" unfolding gen . |
|
1143 |
then have "e \<subseteq> space M1" by (rule sets.sets_into_space) } |
|
1144 |
then show "E \<subseteq> Pow (space M1)" by auto |
|
1145 |
show "sets (distr M M1 X) = sigma_sets (space M1) E" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1146 |
"sets (density M1 (\<lambda>x. ennreal (f x))) = sigma_sets (space M1) E" |
50419 | 1147 |
unfolding gen[symmetric] by auto |
1148 |
qed fact+ |
|
1149 |
qed (insert X f, auto) |
|
1150 |
||
1151 |
lemma distributedI_borel_atMost: |
|
1152 |
fixes f :: "real \<Rightarrow> real" |
|
1153 |
assumes [measurable]: "X \<in> borel_measurable M" |
|
1154 |
and [measurable]: "f \<in> borel_measurable borel" and f[simp]: "AE x in lborel. 0 \<le> f x" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1155 |
and g_eq: "\<And>a. (\<integral>\<^sup>+x. f x * indicator {..a} x \<partial>lborel) = ennreal (g a)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1156 |
and M_eq: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = ennreal (g a)" |
50419 | 1157 |
shows "distributed M lborel X f" |
1158 |
proof (rule distributedI_real) |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1159 |
show "sets (lborel::real measure) = sigma_sets (space lborel) (range atMost)" |
50419 | 1160 |
by (simp add: borel_eq_atMost) |
1161 |
show "Int_stable (range atMost :: real set set)" |
|
1162 |
by (auto simp: Int_stable_def) |
|
1163 |
have vimage_eq: "\<And>a. (X -` {..a} \<inter> space M) = {x\<in>space M. X x \<le> a}" by auto |
|
63040 | 1164 |
define A where "A i = {.. real i}" for i :: nat |
50419 | 1165 |
then show "range A \<subseteq> range atMost" "(\<Union>i. A i) = space lborel" |
1166 |
"\<And>i. emeasure (distr M lborel X) (A i) \<noteq> \<infinity>" |
|
1167 |
by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq) |
|
1168 |
||
1169 |
fix A :: "real set" assume "A \<in> range atMost" |
|
1170 |
then obtain a where A: "A = {..a}" by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51683
diff
changeset
|
1171 |
show "emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>lborel)" |
50419 | 1172 |
unfolding vimage_eq A M_eq g_eq .. |
1173 |
qed auto |
|
1174 |
||
1175 |
lemma (in prob_space) uniform_distributed_params: |
|
1176 |
assumes X: "distributed M MX X (\<lambda>x. indicator A x / measure MX A)" |
|
1177 |
shows "A \<in> sets MX" "measure MX A \<noteq> 0" |
|
1178 |
proof - |
|
1179 |
interpret X: prob_space "distr M MX X" |
|
1180 |
using distributed_measurable[OF X] by (rule prob_space_distr) |
|
1181 |
||
1182 |
show "measure MX A \<noteq> 0" |
|
1183 |
proof |
|
1184 |
assume "measure MX A = 0" |
|
1185 |
with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X] |
|
1186 |
show False |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1187 |
by (simp add: emeasure_density zero_ennreal_def[symmetric]) |
50419 | 1188 |
qed |
1189 |
with measure_notin_sets[of A MX] show "A \<in> sets MX" |
|
1190 |
by blast |
|
1191 |
qed |
|
1192 |
||
47694 | 1193 |
lemma prob_space_uniform_measure: |
1194 |
assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" |
|
1195 |
shows "prob_space (uniform_measure M A)" |
|
1196 |
proof |
|
1197 |
show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1" |
|
1198 |
using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"] |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50104
diff
changeset
|
1199 |
using sets.sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1200 |
by (simp add: Int_absorb2 less_top) |
47694 | 1201 |
qed |
1202 |
||
1203 |
lemma prob_space_uniform_count_measure: "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> prob_space (uniform_count_measure A)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1204 |
by standard (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ennreal_def) |
42860 | 1205 |
|
59000 | 1206 |
lemma (in prob_space) measure_uniform_measure_eq_cond_prob: |
1207 |
assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q" |
|
1208 |
shows "\<P>(x in uniform_measure M {x\<in>space M. Q x}. P x) = \<P>(x in M. P x \<bar> Q x)" |
|
1209 |
proof cases |
|
1210 |
assume Q: "measure M {x\<in>space M. Q x} = 0" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1211 |
then have *: "AE x in M. \<not> Q x" |
59000 | 1212 |
by (simp add: prob_eq_0) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1213 |
then have "density M (\<lambda>x. indicator {x \<in> space M. Q x} x / emeasure M {x \<in> space M. Q x}) = density M (\<lambda>x. 0)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1214 |
by (intro density_cong) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1215 |
with * show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1216 |
unfolding uniform_measure_def |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1217 |
by (simp add: emeasure_density measure_def cond_prob_def emeasure_eq_0_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1218 |
next |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1219 |
assume Q: "measure M {x\<in>space M. Q x} \<noteq> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1220 |
then show "\<P>(x in uniform_measure M {x \<in> space M. Q x}. P x) = cond_prob M P Q" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1221 |
by (subst measure_uniform_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1222 |
(auto simp: emeasure_eq_measure cond_prob_def measure_nonneg intro!: arg_cong[where f=prob]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1223 |
qed |
59000 | 1224 |
|
1225 |
lemma prob_space_point_measure: |
|
1226 |
"finite S \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> 0 \<le> p s) \<Longrightarrow> (\<Sum>s\<in>S. p s) = 1 \<Longrightarrow> prob_space (point_measure S p)" |
|
1227 |
by (rule prob_spaceI) (simp add: space_point_measure emeasure_point_measure_finite) |
|
1228 |
||
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1229 |
lemma (in prob_space) distr_pair_fst: "distr (N \<Otimes>\<^sub>M M) N fst = N" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1230 |
proof (intro measure_eqI) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1231 |
fix A assume A: "A \<in> sets (distr (N \<Otimes>\<^sub>M M) N fst)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1232 |
from A have "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure (N \<Otimes>\<^sub>M M) (A \<times> space M)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1233 |
by (auto simp add: emeasure_distr space_pair_measure dest: sets.sets_into_space intro!: arg_cong2[where f=emeasure]) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1234 |
with A show "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure N A" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1235 |
by (simp add: emeasure_pair_measure_Times emeasure_space_1) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1236 |
qed simp |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1237 |
|
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1238 |
lemma (in product_prob_space) distr_reorder: |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1239 |
assumes "inj_on t J" "t \<in> J \<rightarrow> K" "finite K" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1240 |
shows "distr (PiM K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) = PiM J (\<lambda>x. M (t x))" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1241 |
proof (rule product_sigma_finite.PiM_eqI) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1242 |
show "product_sigma_finite (\<lambda>x. M (t x))" .. |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1243 |
have "t`J \<subseteq> K" using assms by auto |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1244 |
then show [simp]: "finite J" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1245 |
by (rule finite_imageD[OF finite_subset]) fact+ |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1246 |
fix A assume A: "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M (t i))" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1247 |
moreover have "((\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) -` Pi\<^sub>E J A \<inter> space (Pi\<^sub>M K M)) = |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1248 |
(\<Pi>\<^sub>E i\<in>K. if i \<in> t`J then A (the_inv_into J t i) else space (M i))" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1249 |
using A A[THEN sets.sets_into_space] \<open>t \<in> J \<rightarrow> K\<close> \<open>inj_on t J\<close> |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1250 |
by (subst prod_emb_Pi[symmetric]) (auto simp: space_PiM PiE_iff the_inv_into_f_f prod_emb_def) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1251 |
ultimately show "distr (Pi\<^sub>M K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) (Pi\<^sub>E J A) = (\<Prod>i\<in>J. M (t i) (A i))" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1252 |
using assms |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1253 |
apply (subst emeasure_distr) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1254 |
apply (auto intro!: sets_PiM_I_finite simp: Pi_iff) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1255 |
apply (subst emeasure_PiM) |
64272 | 1256 |
apply (auto simp: the_inv_into_f_f \<open>inj_on t J\<close> prod.reindex[OF \<open>inj_on t J\<close>] |
1257 |
if_distrib[where f="emeasure (M _)"] prod.If_cases emeasure_space_1 Int_absorb1 \<open>t`J \<subseteq> K\<close>) |
|
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1258 |
done |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1259 |
qed simp |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1260 |
|
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1261 |
lemma (in product_prob_space) distr_restrict: |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1262 |
"J \<subseteq> K \<Longrightarrow> finite K \<Longrightarrow> (\<Pi>\<^sub>M i\<in>J. M i) = distr (\<Pi>\<^sub>M i\<in>K. M i) (\<Pi>\<^sub>M i\<in>J. M i) (\<lambda>f. restrict f J)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1263 |
using distr_reorder[of "\<lambda>x. x" J K] by (simp add: Pi_iff subset_eq) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1264 |
|
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1265 |
lemma (in product_prob_space) emeasure_prod_emb[simp]: |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1266 |
assumes L: "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^sub>M J M)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1267 |
shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1268 |
by (subst distr_restrict[OF L]) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1269 |
(simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1270 |
|
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1271 |
lemma emeasure_distr_restrict: |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1272 |
assumes "I \<subseteq> K" and Q[measurable_cong]: "sets Q = sets (PiM K M)" and A[measurable]: "A \<in> sets (PiM I M)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1273 |
shows "emeasure (distr Q (PiM I M) (\<lambda>\<omega>. restrict \<omega> I)) A = emeasure Q (prod_emb K M I A)" |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1274 |
using \<open>I\<subseteq>K\<close> sets_eq_imp_space_eq[OF Q] |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1275 |
by (subst emeasure_distr) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1276 |
(auto simp: measurable_cong_sets[OF Q] prod_emb_def space_PiM[symmetric] intro!: measurable_restrict) |
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
61169
diff
changeset
|
1277 |
|
63626
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents:
63099
diff
changeset
|
1278 |
lemma (in prob_space) prob_space_completion: "prob_space (completion M)" |
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents:
63099
diff
changeset
|
1279 |
by (rule prob_spaceI) (simp add: emeasure_space_1) |
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents:
63099
diff
changeset
|
1280 |
|
73253
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1281 |
lemma distr_PiM_finite_prob_space: |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1282 |
assumes fin: "finite I" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1283 |
assumes "product_prob_space M" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1284 |
assumes "product_prob_space M'" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1285 |
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f \<in> measurable (M i) (M' i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1286 |
shows "distr (PiM I M) (PiM I M') (compose I f) = PiM I (\<lambda>i. distr (M i) (M' i) f)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1287 |
proof - |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1288 |
interpret M: product_prob_space M by fact |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1289 |
interpret M': product_prob_space M' by fact |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1290 |
define N where "N = (\<lambda>i. if i \<in> I then distr (M i) (M' i) f else M' i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1291 |
have [intro]: "prob_space (N i)" for i |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1292 |
by (auto simp: N_def intro!: M.M.prob_space_distr M'.prob_space) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1293 |
|
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1294 |
interpret N: product_prob_space N |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1295 |
by (intro product_prob_spaceI) (auto simp: N_def M'.prob_space intro: M.M.prob_space_distr) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1296 |
|
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1297 |
have "distr (PiM I M) (PiM I M') (compose I f) = PiM I N" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1298 |
proof (rule N.PiM_eqI) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1299 |
have N_events_eq: "sets (Pi\<^sub>M I N) = sets (Pi\<^sub>M I M')" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1300 |
unfolding N_def by (intro sets_PiM_cong) auto |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1301 |
also have "\<dots> = sets (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1302 |
by simp |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1303 |
finally show "sets (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) = sets (Pi\<^sub>M I N)" .. |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1304 |
|
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1305 |
fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> N.M.events i" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1306 |
have "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) (Pi\<^sub>E I A) = |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1307 |
emeasure (Pi\<^sub>M I M) (compose I f -` Pi\<^sub>E I A \<inter> space (Pi\<^sub>M I M))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1308 |
proof (intro emeasure_distr) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1309 |
show "compose I f \<in> Pi\<^sub>M I M \<rightarrow>\<^sub>M Pi\<^sub>M I M'" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1310 |
unfolding compose_def by measurable |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1311 |
show "Pi\<^sub>E I A \<in> sets (Pi\<^sub>M I M')" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1312 |
unfolding N_events_eq [symmetric] by (intro sets_PiM_I_finite fin A) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1313 |
qed |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1314 |
also have "compose I f -` Pi\<^sub>E I A \<inter> space (Pi\<^sub>M I M) = Pi\<^sub>E I (\<lambda>i. f -` A i \<inter> space (M i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1315 |
using A by (auto simp: space_PiM PiE_def Pi_def extensional_def N_def compose_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1316 |
also have "emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I (\<lambda>i. f -` A i \<inter> space (M i))) = |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1317 |
(\<Prod>i\<in>I. emeasure (M i) (f -` A i \<inter> space (M i)))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1318 |
using A by (intro M.emeasure_PiM fin) (auto simp: N_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1319 |
also have "\<dots> = (\<Prod>i\<in>I. emeasure (distr (M i) (M' i) f) (A i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1320 |
using A by (intro prod.cong emeasure_distr [symmetric]) (auto simp: N_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1321 |
also have "\<dots> = (\<Prod>i\<in>I. emeasure (N i) (A i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1322 |
unfolding N_def by (intro prod.cong) (auto simp: N_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1323 |
finally show "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) (Pi\<^sub>E I A) = \<dots>" . |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1324 |
qed fact+ |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1325 |
also have "PiM I N = PiM I (\<lambda>i. distr (M i) (M' i) f)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1326 |
by (intro PiM_cong) (auto simp: N_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1327 |
finally show ?thesis . |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1328 |
qed |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69597
diff
changeset
|
1329 |
|
35582 | 1330 |
end |