author | blanchet |
Mon, 21 May 2012 10:39:31 +0200 | |
changeset 47944 | e6b51fab96f7 |
parent 47762 | d31085f07f60 |
child 49776 | 199d1d5bb17e |
permissions | -rw-r--r-- |
42147 | 1 |
(* Title: HOL/Probability/Infinite_Product_Measure.thy |
2 |
Author: Johannes Hölzl, TU München |
|
3 |
*) |
|
4 |
||
5 |
header {*Infinite Product Measure*} |
|
6 |
||
7 |
theory Infinite_Product_Measure |
|
47694 | 8 |
imports Probability_Measure Caratheodory |
42147 | 9 |
begin |
10 |
||
47694 | 11 |
lemma (in product_sigma_finite) |
12 |
assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)" |
|
13 |
shows emeasure_fold_integral: |
|
14 |
"emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I) |
|
15 |
and emeasure_fold_measurable: |
|
16 |
"(\<lambda>x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B) |
|
17 |
proof - |
|
18 |
interpret I: finite_product_sigma_finite M I by default fact |
|
19 |
interpret J: finite_product_sigma_finite M J by default fact |
|
20 |
interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" .. |
|
21 |
have merge: "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)" |
|
22 |
by (intro measurable_sets[OF _ A] measurable_merge assms) |
|
23 |
||
24 |
show ?I |
|
25 |
apply (subst distr_merge[symmetric, OF IJ]) |
|
26 |
apply (subst emeasure_distr[OF measurable_merge[OF IJ(1)] A]) |
|
27 |
apply (subst IJ.emeasure_pair_measure_alt[OF merge]) |
|
28 |
apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure) |
|
29 |
done |
|
30 |
||
31 |
show ?B |
|
32 |
using IJ.measurable_emeasure_Pair1[OF merge] |
|
33 |
by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong) |
|
34 |
qed |
|
35 |
||
42147 | 36 |
lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B" |
37 |
unfolding restrict_def extensional_def by auto |
|
38 |
||
39 |
lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \<inter> B)" |
|
40 |
unfolding restrict_def by (simp add: fun_eq_iff) |
|
41 |
||
42 |
lemma split_merge: "P (merge I x J y i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)" |
|
43 |
unfolding merge_def by auto |
|
44 |
||
45 |
lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I x J y \<in> extensional K" |
|
46 |
unfolding merge_def extensional_def by auto |
|
47 |
||
48 |
lemma injective_vimage_restrict: |
|
49 |
assumes J: "J \<subseteq> I" |
|
50 |
and sets: "A \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" and ne: "(\<Pi>\<^isub>E i\<in>I. S i) \<noteq> {}" |
|
51 |
and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^isub>E i\<in>I. S i)" |
|
52 |
shows "A = B" |
|
53 |
proof (intro set_eqI) |
|
54 |
fix x |
|
55 |
from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto |
|
56 |
have "J \<inter> (I - J) = {}" by auto |
|
57 |
show "x \<in> A \<longleftrightarrow> x \<in> B" |
|
58 |
proof cases |
|
59 |
assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)" |
|
60 |
have "x \<in> A \<longleftrightarrow> merge J x (I - J) y \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)" |
|
61 |
using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge) |
|
62 |
then show "x \<in> A \<longleftrightarrow> x \<in> B" |
|
63 |
using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge) |
|
64 |
next |
|
65 |
assume "x \<notin> (\<Pi>\<^isub>E i\<in>J. S i)" with sets show "x \<in> A \<longleftrightarrow> x \<in> B" by auto |
|
66 |
qed |
|
67 |
qed |
|
68 |
||
47694 | 69 |
lemma prod_algebraI_finite: |
70 |
"finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^isub>E I E) \<in> prod_algebra I M" |
|
71 |
using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets_into_space] by simp |
|
72 |
||
73 |
lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}" |
|
74 |
proof (safe intro!: Int_stableI) |
|
75 |
fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)" |
|
76 |
then show "\<exists>G. Pi\<^isub>E J E \<inter> Pi\<^isub>E J F = Pi\<^isub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))" |
|
77 |
by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"]) |
|
78 |
qed |
|
79 |
||
80 |
lemma prod_emb_trans[simp]: |
|
81 |
"J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X" |
|
82 |
by (auto simp add: Int_absorb1 prod_emb_def) |
|
83 |
||
84 |
lemma prod_emb_Pi: |
|
85 |
assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K" |
|
86 |
shows "prod_emb K M J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))" |
|
87 |
using assms space_closed |
|
88 |
by (auto simp: prod_emb_def Pi_iff split: split_if_asm) blast+ |
|
89 |
||
90 |
lemma prod_emb_id: |
|
91 |
"B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B" |
|
92 |
by (auto simp: prod_emb_def Pi_iff subset_eq extensional_restrict) |
|
93 |
||
94 |
lemma measurable_prod_emb[intro, simp]: |
|
95 |
"J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^isub>M L M)" |
|
96 |
unfolding prod_emb_def space_PiM[symmetric] |
|
97 |
by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton) |
|
98 |
||
99 |
lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)" |
|
100 |
by (intro measurable_restrict measurable_component_singleton) auto |
|
101 |
||
102 |
lemma (in product_prob_space) distr_restrict: |
|
42147 | 103 |
assumes "J \<noteq> {}" "J \<subseteq> K" "finite K" |
47694 | 104 |
shows "(\<Pi>\<^isub>M i\<in>J. M i) = distr (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D") |
105 |
proof (rule measure_eqI_generator_eq) |
|
106 |
have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset) |
|
107 |
interpret J: finite_product_prob_space M J proof qed fact |
|
108 |
interpret K: finite_product_prob_space M K proof qed fact |
|
109 |
||
110 |
let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}" |
|
111 |
let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)" |
|
112 |
let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))" |
|
113 |
show "Int_stable ?J" |
|
114 |
by (rule Int_stable_PiE) |
|
115 |
show "range ?F \<subseteq> ?J" "incseq ?F" "(\<Union>i. ?F i) = ?\<Omega>" |
|
116 |
using `finite J` by (auto intro!: prod_algebraI_finite) |
|
117 |
{ fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp } |
|
118 |
show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space) |
|
119 |
show "sets (\<Pi>\<^isub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J" |
|
120 |
using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff) |
|
121 |
||
122 |
fix X assume "X \<in> ?J" |
|
123 |
then obtain E where [simp]: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto |
|
124 |
with `finite J` have X: "X \<in> sets (Pi\<^isub>M J M)" by auto |
|
125 |
||
126 |
have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))" |
|
127 |
using E by (simp add: J.measure_times) |
|
128 |
also have "\<dots> = (\<Prod> i\<in>J. emeasure (M i) (if i \<in> J then E i else space (M i)))" |
|
129 |
by simp |
|
130 |
also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))" |
|
131 |
using `finite K` `J \<subseteq> K` |
|
132 |
by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1) |
|
133 |
also have "\<dots> = emeasure (Pi\<^isub>M K M) (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))" |
|
134 |
using E by (simp add: K.measure_times) |
|
135 |
also have "(\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i))" |
|
136 |
using `J \<subseteq> K` sets_into_space E by (force simp: Pi_iff split: split_if_asm) |
|
137 |
finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X" |
|
138 |
using X `J \<subseteq> K` apply (subst emeasure_distr) |
|
139 |
by (auto intro!: measurable_restrict_subset simp: space_PiM) |
|
42147 | 140 |
qed |
141 |
||
47694 | 142 |
abbreviation (in product_prob_space) |
143 |
"emb L K X \<equiv> prod_emb L M K X" |
|
144 |
||
145 |
lemma (in product_prob_space) emeasure_prod_emb[simp]: |
|
146 |
assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)" |
|
147 |
shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X" |
|
148 |
by (subst distr_restrict[OF L]) |
|
149 |
(simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X) |
|
42147 | 150 |
|
47694 | 151 |
lemma (in product_prob_space) prod_emb_injective: |
152 |
assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)" |
|
153 |
assumes "prod_emb L M J X = prod_emb L M J Y" |
|
154 |
shows "X = Y" |
|
155 |
proof (rule injective_vimage_restrict) |
|
156 |
show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" |
|
157 |
using sets[THEN sets_into_space] by (auto simp: space_PiM) |
|
158 |
have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)" |
|
159 |
using M.not_empty by auto |
|
160 |
from bchoice[OF this] |
|
161 |
show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto |
|
162 |
show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))" |
|
163 |
using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def) |
|
164 |
qed fact |
|
42147 | 165 |
|
47694 | 166 |
definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) set set" where |
167 |
"generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))" |
|
42147 | 168 |
|
47694 | 169 |
lemma (in product_prob_space) generatorI': |
170 |
"J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> generator" |
|
171 |
unfolding generator_def by auto |
|
42147 | 172 |
|
47694 | 173 |
lemma (in product_prob_space) algebra_generator: |
174 |
assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G") |
|
47762 | 175 |
unfolding algebra_def algebra_axioms_def ring_of_sets_iff |
176 |
proof (intro conjI ballI) |
|
47694 | 177 |
let ?G = generator |
178 |
show "?G \<subseteq> Pow ?\<Omega>" |
|
179 |
by (auto simp: generator_def prod_emb_def) |
|
180 |
from `I \<noteq> {}` obtain i where "i \<in> I" by auto |
|
181 |
then show "{} \<in> ?G" |
|
182 |
by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"] |
|
183 |
simp: sigma_sets.Empty generator_def prod_emb_def) |
|
184 |
from `i \<in> I` show "?\<Omega> \<in> ?G" |
|
185 |
by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"] |
|
186 |
simp: generator_def prod_emb_def) |
|
187 |
fix A assume "A \<in> ?G" |
|
188 |
then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA" |
|
189 |
by (auto simp: generator_def) |
|
190 |
fix B assume "B \<in> ?G" |
|
191 |
then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB" |
|
192 |
by (auto simp: generator_def) |
|
193 |
let ?RA = "emb (JA \<union> JB) JA XA" |
|
194 |
let ?RB = "emb (JA \<union> JB) JB XB" |
|
195 |
have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)" |
|
196 |
using XA A XB B by auto |
|
197 |
show "A - B \<in> ?G" "A \<union> B \<in> ?G" |
|
198 |
unfolding * using XA XB by (safe intro!: generatorI') auto |
|
42147 | 199 |
qed |
200 |
||
47694 | 201 |
lemma (in product_prob_space) sets_PiM_generator: |
202 |
assumes "I \<noteq> {}" shows "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" |
|
203 |
proof |
|
204 |
show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" |
|
205 |
unfolding sets_PiM |
|
206 |
proof (safe intro!: sigma_sets_subseteq) |
|
207 |
fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator" |
|
208 |
by (auto intro!: generatorI' elim!: prod_algebraE) |
|
209 |
qed |
|
210 |
qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset) |
|
42147 | 211 |
|
212 |
lemma (in product_prob_space) generatorI: |
|
47694 | 213 |
"J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator" |
42147 | 214 |
unfolding generator_def by auto |
215 |
||
216 |
definition (in product_prob_space) |
|
217 |
"\<mu>G A = |
|
47694 | 218 |
(THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (Pi\<^isub>M J M) X))" |
42147 | 219 |
|
220 |
lemma (in product_prob_space) \<mu>G_spec: |
|
221 |
assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)" |
|
47694 | 222 |
shows "\<mu>G A = emeasure (Pi\<^isub>M J M) X" |
42147 | 223 |
unfolding \<mu>G_def |
224 |
proof (intro the_equality allI impI ballI) |
|
225 |
fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)" |
|
47694 | 226 |
have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)" |
42147 | 227 |
using K J by simp |
228 |
also have "emb (K \<union> J) K Y = emb (K \<union> J) J X" |
|
47694 | 229 |
using K J by (simp add: prod_emb_injective[of "K \<union> J" I]) |
230 |
also have "emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = emeasure (Pi\<^isub>M J M) X" |
|
42147 | 231 |
using K J by simp |
47694 | 232 |
finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" .. |
42147 | 233 |
qed (insert J, force) |
234 |
||
235 |
lemma (in product_prob_space) \<mu>G_eq: |
|
47694 | 236 |
"J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (Pi\<^isub>M J M) X" |
42147 | 237 |
by (intro \<mu>G_spec) auto |
238 |
||
239 |
lemma (in product_prob_space) generator_Ex: |
|
47694 | 240 |
assumes *: "A \<in> generator" |
241 |
shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (Pi\<^isub>M J M) X" |
|
42147 | 242 |
proof - |
243 |
from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)" |
|
244 |
unfolding generator_def by auto |
|
245 |
with \<mu>G_spec[OF this] show ?thesis by auto |
|
246 |
qed |
|
247 |
||
248 |
lemma (in product_prob_space) generatorE: |
|
47694 | 249 |
assumes A: "A \<in> generator" |
250 |
obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (Pi\<^isub>M J M) X" |
|
42147 | 251 |
proof - |
252 |
from generator_Ex[OF A] obtain X J where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" |
|
47694 | 253 |
"\<mu>G A = emeasure (Pi\<^isub>M J M) X" by auto |
42147 | 254 |
then show thesis by (intro that) auto |
255 |
qed |
|
256 |
||
257 |
lemma (in product_prob_space) merge_sets: |
|
258 |
assumes "finite J" "finite K" "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)" |
|
259 |
shows "merge J x K -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)" |
|
260 |
proof - |
|
47694 | 261 |
from sets_Pair1[OF |
42147 | 262 |
measurable_merge[THEN measurable_sets, OF `J \<inter> K = {}`], OF A, of x] x |
263 |
show ?thesis |
|
264 |
by (simp add: space_pair_measure comp_def vimage_compose[symmetric]) |
|
265 |
qed |
|
266 |
||
267 |
lemma (in product_prob_space) merge_emb: |
|
268 |
assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)" |
|
269 |
shows "(merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = |
|
270 |
emb I (K - J) (merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))" |
|
271 |
proof - |
|
272 |
have [simp]: "\<And>x J K L. merge J y K (restrict x L) = merge J y (K \<inter> L) x" |
|
273 |
by (auto simp: restrict_def merge_def) |
|
274 |
have [simp]: "\<And>x J K L. restrict (merge J y K x) L = merge (J \<inter> L) y (K \<inter> L) x" |
|
275 |
by (auto simp: restrict_def merge_def) |
|
276 |
have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto |
|
277 |
have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto |
|
278 |
have [simp]: "(K - J) \<inter> K = K - J" by auto |
|
279 |
from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis |
|
47694 | 280 |
by (simp split: split_merge add: prod_emb_def Pi_iff extensional_merge_sub set_eq_iff space_PiM) |
281 |
auto |
|
42147 | 282 |
qed |
283 |
||
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:
44928
diff
changeset
|
284 |
lemma (in product_prob_space) positive_\<mu>G: |
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:
44928
diff
changeset
|
285 |
assumes "I \<noteq> {}" |
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:
44928
diff
changeset
|
286 |
shows "positive generator \<mu>G" |
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:
44928
diff
changeset
|
287 |
proof - |
47694 | 288 |
interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact |
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:
44928
diff
changeset
|
289 |
show ?thesis |
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:
44928
diff
changeset
|
290 |
proof (intro positive_def[THEN iffD2] conjI ballI) |
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:
44928
diff
changeset
|
291 |
from generatorE[OF G.empty_sets] guess J X . note this[simp] |
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:
44928
diff
changeset
|
292 |
interpret J: finite_product_sigma_finite M J by default fact |
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:
44928
diff
changeset
|
293 |
have "X = {}" |
47694 | 294 |
by (rule prod_emb_injective[of J I]) simp_all |
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:
44928
diff
changeset
|
295 |
then show "\<mu>G {} = 0" by simp |
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:
44928
diff
changeset
|
296 |
next |
47694 | 297 |
fix A assume "A \<in> generator" |
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:
44928
diff
changeset
|
298 |
from generatorE[OF this] guess J X . note this[simp] |
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:
44928
diff
changeset
|
299 |
interpret J: finite_product_sigma_finite M J by default fact |
47694 | 300 |
show "0 \<le> \<mu>G A" by (simp add: emeasure_nonneg) |
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:
44928
diff
changeset
|
301 |
qed |
42147 | 302 |
qed |
303 |
||
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:
44928
diff
changeset
|
304 |
lemma (in product_prob_space) additive_\<mu>G: |
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:
44928
diff
changeset
|
305 |
assumes "I \<noteq> {}" |
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:
44928
diff
changeset
|
306 |
shows "additive generator \<mu>G" |
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:
44928
diff
changeset
|
307 |
proof - |
47694 | 308 |
interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact |
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:
44928
diff
changeset
|
309 |
show ?thesis |
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:
44928
diff
changeset
|
310 |
proof (intro additive_def[THEN iffD2] ballI impI) |
47694 | 311 |
fix A assume "A \<in> generator" with generatorE guess J X . note J = this |
312 |
fix B assume "B \<in> generator" with generatorE guess K Y . note K = this |
|
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:
44928
diff
changeset
|
313 |
assume "A \<inter> B = {}" |
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:
44928
diff
changeset
|
314 |
have JK: "J \<union> K \<noteq> {}" "J \<union> K \<subseteq> I" "finite (J \<union> K)" |
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:
44928
diff
changeset
|
315 |
using J K by auto |
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:
44928
diff
changeset
|
316 |
interpret JK: finite_product_sigma_finite M "J \<union> K" by default fact |
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:
44928
diff
changeset
|
317 |
have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}" |
47694 | 318 |
apply (rule prod_emb_injective[of "J \<union> K" 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:
44928
diff
changeset
|
319 |
apply (insert `A \<inter> B = {}` JK J K) |
47694 | 320 |
apply (simp_all add: Int prod_emb_Int) |
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:
44928
diff
changeset
|
321 |
done |
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:
44928
diff
changeset
|
322 |
have AB: "A = emb I (J \<union> K) (emb (J \<union> K) J X)" "B = emb I (J \<union> K) (emb (J \<union> K) K Y)" |
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:
44928
diff
changeset
|
323 |
using J K by simp_all |
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:
44928
diff
changeset
|
324 |
then have "\<mu>G (A \<union> B) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))" |
47694 | 325 |
by simp |
326 |
also have "\<dots> = emeasure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)" |
|
327 |
using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq Un del: prod_emb_Un) |
|
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:
44928
diff
changeset
|
328 |
also have "\<dots> = \<mu>G A + \<mu>G B" |
47694 | 329 |
using J K JK_disj by (simp add: plus_emeasure[symmetric]) |
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:
44928
diff
changeset
|
330 |
finally show "\<mu>G (A \<union> B) = \<mu>G A + \<mu>G B" . |
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:
44928
diff
changeset
|
331 |
qed |
42147 | 332 |
qed |
333 |
||
47694 | 334 |
lemma (in product_prob_space) emeasure_PiM_emb_not_empty: |
335 |
assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)" |
|
336 |
shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)" |
|
42147 | 337 |
proof cases |
47694 | 338 |
assume "finite I" with X show ?thesis by simp |
42147 | 339 |
next |
47694 | 340 |
let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space (M i)" |
42147 | 341 |
let ?G = generator |
342 |
assume "\<not> finite 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:
44928
diff
changeset
|
343 |
then have I_not_empty: "I \<noteq> {}" by auto |
47694 | 344 |
interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact |
42147 | 345 |
note \<mu>G_mono = |
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:
44928
diff
changeset
|
346 |
G.additive_increasing[OF positive_\<mu>G[OF I_not_empty] additive_\<mu>G[OF I_not_empty], THEN increasingD] |
42147 | 347 |
|
47694 | 348 |
{ fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G" |
42147 | 349 |
|
350 |
from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J" |
|
351 |
by (metis rev_finite_subset subsetI) |
|
352 |
moreover from Z guess K' X' by (rule generatorE) |
|
353 |
moreover def K \<equiv> "insert k K'" |
|
354 |
moreover def X \<equiv> "emb K K' X'" |
|
355 |
ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X" |
|
47694 | 356 |
"K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X" |
42147 | 357 |
by (auto simp: subset_insertI) |
358 |
||
46731 | 359 |
let ?M = "\<lambda>y. merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)" |
42147 | 360 |
{ fix y assume y: "y \<in> space (Pi\<^isub>M J M)" |
361 |
note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X] |
|
362 |
moreover |
|
363 |
have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)" |
|
364 |
using J K y by (intro merge_sets) auto |
|
365 |
ultimately |
|
47694 | 366 |
have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G" |
42147 | 367 |
using J K by (intro generatorI) auto |
47694 | 368 |
have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)" |
42147 | 369 |
unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto |
370 |
note * ** *** this } |
|
371 |
note merge_in_G = this |
|
372 |
||
373 |
have "finite (K - J)" using K by auto |
|
374 |
||
375 |
interpret J: finite_product_prob_space M J by default fact+ |
|
376 |
interpret KmJ: finite_product_prob_space M "K - J" by default fact+ |
|
377 |
||
47694 | 378 |
have "\<mu>G Z = emeasure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)" |
42147 | 379 |
using K J by simp |
47694 | 380 |
also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)" |
381 |
using K J by (subst emeasure_fold_integral) auto |
|
42147 | 382 |
also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)" |
383 |
(is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)") |
|
47694 | 384 |
proof (intro positive_integral_cong) |
42147 | 385 |
fix x assume x: "x \<in> space (Pi\<^isub>M J M)" |
386 |
with K merge_in_G(2)[OF this] |
|
47694 | 387 |
show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)" |
42147 | 388 |
unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \<mu>G_eq) auto |
389 |
qed |
|
390 |
finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" . |
|
391 |
||
392 |
{ fix x assume x: "x \<in> space (Pi\<^isub>M J M)" |
|
393 |
then have "\<mu>G (?MZ x) \<le> 1" |
|
394 |
unfolding merge_in_G(4)[OF x] `Z = emb I K X` |
|
395 |
by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) } |
|
396 |
note le_1 = this |
|
397 |
||
46731 | 398 |
let ?q = "\<lambda>y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M))" |
42147 | 399 |
have "?q \<in> borel_measurable (Pi\<^isub>M J M)" |
400 |
unfolding `Z = emb I K X` using J K merge_in_G(3) |
|
47694 | 401 |
by (simp add: merge_in_G \<mu>G_eq emeasure_fold_measurable cong: measurable_cong) |
42147 | 402 |
note this fold le_1 merge_in_G(3) } |
403 |
note fold = this |
|
404 |
||
47694 | 405 |
have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>" |
42147 | 406 |
proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G]) |
47694 | 407 |
fix A assume "A \<in> ?G" |
42147 | 408 |
with generatorE guess J X . note JX = this |
409 |
interpret JK: finite_product_prob_space M J by default fact+ |
|
46898
1570b30ee040
tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents:
46731
diff
changeset
|
410 |
from JX show "\<mu>G A \<noteq> \<infinity>" by simp |
42147 | 411 |
next |
47694 | 412 |
fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}" |
42147 | 413 |
then have "decseq (\<lambda>i. \<mu>G (A i))" |
414 |
by (auto intro!: \<mu>G_mono simp: decseq_def) |
|
415 |
moreover |
|
416 |
have "(INF i. \<mu>G (A i)) = 0" |
|
417 |
proof (rule ccontr) |
|
418 |
assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0") |
|
419 |
moreover have "0 \<le> ?a" |
|
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:
44928
diff
changeset
|
420 |
using A positive_\<mu>G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def) |
42147 | 421 |
ultimately have "0 < ?a" by auto |
422 |
||
47694 | 423 |
have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (Pi\<^isub>M J M) X" |
42147 | 424 |
using A by (intro allI generator_Ex) auto |
425 |
then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)" |
|
426 |
and A': "\<And>n. A n = emb I (J' n) (X' n)" |
|
427 |
unfolding choice_iff by blast |
|
428 |
moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)" |
|
429 |
moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)" |
|
430 |
ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^isub>M (J n) M)" |
|
431 |
by auto |
|
47694 | 432 |
with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G" |
433 |
unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto) |
|
42147 | 434 |
|
435 |
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m" |
|
436 |
unfolding J_def by force |
|
437 |
||
438 |
interpret J: finite_product_prob_space M "J i" for i by default fact+ |
|
439 |
||
440 |
have a_le_1: "?a \<le> 1" |
|
441 |
using \<mu>G_spec[of "J 0" "A 0" "X 0"] J A_eq |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
43920
diff
changeset
|
442 |
by (auto intro!: INF_lower2[of 0] J.measure_le_1) |
42147 | 443 |
|
46731 | 444 |
let ?M = "\<lambda>K Z y. merge K y (I - K) -` Z \<inter> space (Pi\<^isub>M I M)" |
42147 | 445 |
|
47694 | 446 |
{ fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)" |
447 |
then have Z_sets: "\<And>n. Z n \<in> ?G" by auto |
|
42147 | 448 |
fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I" |
449 |
interpret J': finite_product_prob_space M J' by default fact+ |
|
450 |
||
46731 | 451 |
let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)" |
452 |
let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^isub>M J' M)" |
|
42147 | 453 |
{ fix n |
454 |
have "?q n \<in> borel_measurable (Pi\<^isub>M J' M)" |
|
455 |
using Z J' by (intro fold(1)) auto |
|
456 |
then have "?Q n \<in> sets (Pi\<^isub>M J' M)" |
|
457 |
by (rule measurable_sets) auto } |
|
458 |
note Q_sets = this |
|
459 |
||
47694 | 460 |
have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^isub>M J' M) (?Q n))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
43920
diff
changeset
|
461 |
proof (intro INF_greatest) |
42147 | 462 |
fix n |
463 |
have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto |
|
464 |
also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)" |
|
47694 | 465 |
unfolding fold(2)[OF J' `Z n \<in> ?G`] |
466 |
proof (intro positive_integral_mono) |
|
42147 | 467 |
fix x assume x: "x \<in> space (Pi\<^isub>M J' M)" |
468 |
then have "?q n x \<le> 1 + 0" |
|
469 |
using J' Z fold(3) Z_sets by auto |
|
470 |
also have "\<dots> \<le> 1 + ?a / 2^(k+1)" |
|
471 |
using `0 < ?a` by (intro add_mono) auto |
|
472 |
finally have "?q n x \<le> 1 + ?a / 2^(k+1)" . |
|
473 |
with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)" |
|
474 |
by (auto split: split_indicator simp del: power_Suc) |
|
475 |
qed |
|
47694 | 476 |
also have "\<dots> = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)" |
477 |
using `0 \<le> ?a` Q_sets J'.emeasure_space_1 |
|
478 |
by (subst positive_integral_add) auto |
|
479 |
finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1` |
|
480 |
by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"]) |
|
42147 | 481 |
(auto simp: field_simps) |
482 |
qed |
|
47694 | 483 |
also have "\<dots> = emeasure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)" |
484 |
proof (intro INF_emeasure_decseq) |
|
42147 | 485 |
show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto |
486 |
show "decseq ?Q" |
|
487 |
unfolding decseq_def |
|
488 |
proof (safe intro!: vimageI[OF refl]) |
|
489 |
fix m n :: nat assume "m \<le> n" |
|
490 |
fix x assume x: "x \<in> space (Pi\<^isub>M J' M)" |
|
491 |
assume "?a / 2^(k+1) \<le> ?q n x" |
|
492 |
also have "?q n x \<le> ?q m x" |
|
493 |
proof (rule \<mu>G_mono) |
|
494 |
from fold(4)[OF J', OF Z_sets x] |
|
47694 | 495 |
show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto |
42147 | 496 |
show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x" |
497 |
using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto |
|
498 |
qed |
|
499 |
finally show "?a / 2^(k+1) \<le> ?q m x" . |
|
500 |
qed |
|
47694 | 501 |
qed simp |
42147 | 502 |
finally have "(\<Inter>n. ?Q n) \<noteq> {}" |
503 |
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) |
|
504 |
then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto } |
|
505 |
note Ex_w = this |
|
506 |
||
46731 | 507 |
let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)" |
42147 | 508 |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
43920
diff
changeset
|
509 |
have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower) |
42147 | 510 |
from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this |
511 |
||
46731 | 512 |
let ?P = |
513 |
"\<lambda>k wk w. w \<in> space (Pi\<^isub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and> |
|
514 |
(\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)" |
|
42147 | 515 |
def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))" |
516 |
||
517 |
{ fix k have w: "w k \<in> space (Pi\<^isub>M (J k) M) \<and> |
|
518 |
(\<forall>n. ?a / 2 ^ (k + 1) \<le> ?q k n (w k)) \<and> (k \<noteq> 0 \<longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1))" |
|
519 |
proof (induct k) |
|
520 |
case 0 with w0 show ?case |
|
521 |
unfolding w_def nat_rec_0 by auto |
|
522 |
next |
|
523 |
case (Suc k) |
|
524 |
then have wk: "w k \<in> space (Pi\<^isub>M (J k) M)" by auto |
|
525 |
have "\<exists>w'. ?P k (w k) w'" |
|
526 |
proof cases |
|
527 |
assume [simp]: "J k = J (Suc k)" |
|
528 |
show ?thesis |
|
529 |
proof (intro exI[of _ "w k"] conjI allI) |
|
530 |
fix n |
|
531 |
have "?a / 2 ^ (Suc k + 1) \<le> ?a / 2 ^ (k + 1)" |
|
532 |
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps) |
|
533 |
also have "\<dots> \<le> ?q k n (w k)" using Suc by auto |
|
534 |
finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp |
|
535 |
next |
|
536 |
show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)" |
|
537 |
using Suc by simp |
|
538 |
then show "restrict (w k) (J k) = w k" |
|
47694 | 539 |
by (simp add: extensional_restrict space_PiM) |
42147 | 540 |
qed |
541 |
next |
|
542 |
assume "J k \<noteq> J (Suc k)" |
|
543 |
with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto |
|
47694 | 544 |
have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> ?G" |
42147 | 545 |
"decseq (\<lambda>n. ?M (J k) (A n) (w k))" |
546 |
"\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))" |
|
547 |
using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc |
|
548 |
by (auto simp: decseq_def) |
|
549 |
from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"] |
|
550 |
obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)" |
|
551 |
"\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto |
|
552 |
let ?w = "merge (J k) (w k) ?D w'" |
|
553 |
have [simp]: "\<And>x. merge (J k) (w k) (I - J k) (merge ?D w' (I - ?D) x) = |
|
554 |
merge (J (Suc k)) ?w (I - (J (Suc k))) x" |
|
555 |
using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"] |
|
556 |
by (auto intro!: ext split: split_merge) |
|
557 |
have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w" |
|
558 |
using w'(1) J(3)[of "Suc k"] |
|
47694 | 559 |
by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+ |
42147 | 560 |
show ?thesis |
561 |
apply (rule exI[of _ ?w]) |
|
562 |
using w' J_mono[of k "Suc k"] wk unfolding * |
|
47694 | 563 |
apply (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM) |
42147 | 564 |
apply (force simp: extensional_def) |
565 |
done |
|
566 |
qed |
|
567 |
then have "?P k (w k) (w (Suc k))" |
|
568 |
unfolding w_def nat_rec_Suc unfolding w_def[symmetric] |
|
569 |
by (rule someI_ex) |
|
570 |
then show ?case by auto |
|
571 |
qed |
|
572 |
moreover |
|
573 |
then have "w k \<in> space (Pi\<^isub>M (J k) M)" by auto |
|
574 |
moreover |
|
575 |
from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto |
|
576 |
then have "?M (J k) (A k) (w k) \<noteq> {}" |
|
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:
44928
diff
changeset
|
577 |
using positive_\<mu>G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1` |
42147 | 578 |
by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) |
579 |
then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto |
|
580 |
then have "merge (J k) (w k) (I - J k) x \<in> A k" by auto |
|
581 |
then have "\<exists>x\<in>A k. restrict x (J k) = w k" |
|
582 |
using `w k \<in> space (Pi\<^isub>M (J k) M)` |
|
47694 | 583 |
by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) |
42147 | 584 |
ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)" |
585 |
"\<exists>x\<in>A k. restrict x (J k) = w k" |
|
586 |
"k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)" |
|
587 |
by auto } |
|
588 |
note w = this |
|
589 |
||
590 |
{ fix k l i assume "k \<le> l" "i \<in> J k" |
|
591 |
{ fix l have "w k i = w (k + l) i" |
|
592 |
proof (induct l) |
|
593 |
case (Suc l) |
|
594 |
from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto |
|
595 |
with w(3)[of "k + Suc l"] |
|
596 |
have "w (k + l) i = w (k + Suc l) i" |
|
597 |
by (auto simp: restrict_def fun_eq_iff split: split_if_asm) |
|
598 |
with Suc show ?case by simp |
|
599 |
qed simp } |
|
600 |
from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp } |
|
601 |
note w_mono = this |
|
602 |
||
603 |
def w' \<equiv> "\<lambda>i. if i \<in> (\<Union>k. J k) then w (LEAST k. i \<in> J k) i else if i \<in> I then (SOME x. x \<in> space (M i)) else undefined" |
|
604 |
{ fix i k assume k: "i \<in> J k" |
|
605 |
have "w k i = w (LEAST k. i \<in> J k) i" |
|
606 |
by (intro w_mono Least_le k LeastI[of _ k]) |
|
607 |
then have "w' i = w k i" |
|
608 |
unfolding w'_def using k by auto } |
|
609 |
note w'_eq = this |
|
610 |
have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined" |
|
611 |
using J by (auto simp: w'_def) |
|
612 |
have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)" |
|
613 |
using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]]) |
|
614 |
{ fix i assume "i \<in> I" then have "w' i \<in> space (M i)" |
|
47694 | 615 |
using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ } |
42147 | 616 |
note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this |
617 |
||
618 |
have w': "w' \<in> space (Pi\<^isub>M I M)" |
|
47694 | 619 |
using w(1) by (auto simp add: Pi_iff extensional_def space_PiM) |
42147 | 620 |
|
621 |
{ fix n |
|
622 |
have "restrict w' (J n) = w n" using w(1) |
|
47694 | 623 |
by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def space_PiM) |
42147 | 624 |
with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto |
47694 | 625 |
then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) } |
42147 | 626 |
then have "w' \<in> (\<Inter>i. A i)" by auto |
627 |
with `(\<Inter>i. A i) = {}` show False by auto |
|
628 |
qed |
|
629 |
ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0" |
|
43920 | 630 |
using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp |
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:
44928
diff
changeset
|
631 |
qed fact+ |
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:
44928
diff
changeset
|
632 |
then guess \<mu> .. note \<mu> = this |
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:
44928
diff
changeset
|
633 |
show ?thesis |
47694 | 634 |
proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X]) |
635 |
from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))" |
|
636 |
by (simp add: Pi_iff) |
|
637 |
next |
|
638 |
fix J X assume J: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))" |
|
639 |
then show "emb I J (Pi\<^isub>E J X) \<in> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))" |
|
640 |
by (auto simp: Pi_iff prod_emb_def dest: sets_into_space) |
|
641 |
have "emb I J (Pi\<^isub>E J X) \<in> generator" |
|
642 |
using J `I \<noteq> {}` by (intro generatorI') auto |
|
643 |
then have "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))" |
|
644 |
using \<mu> by simp |
|
645 |
also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" |
|
646 |
using J `I \<noteq> {}` by (subst \<mu>G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff) |
|
647 |
also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. |
|
648 |
if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" |
|
649 |
using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) |
|
650 |
finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = \<dots>" . |
|
651 |
next |
|
652 |
let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))" |
|
653 |
have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)" |
|
654 |
using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) |
|
655 |
then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = |
|
656 |
emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)" |
|
657 |
using X by (auto simp add: emeasure_PiM) |
|
658 |
next |
|
659 |
show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>" |
|
660 |
using \<mu> unfolding sets_PiM_generator[OF `I \<noteq> {}`] by (auto simp: measure_space_def) |
|
42147 | 661 |
qed |
662 |
qed |
|
663 |
||
47694 | 664 |
sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^isub>M I M" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
665 |
proof |
47694 | 666 |
show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1" |
667 |
proof cases |
|
668 |
assume "I = {}" then show ?thesis by (simp add: space_PiM_empty) |
|
669 |
next |
|
670 |
assume "I \<noteq> {}" |
|
671 |
then obtain i where "i \<in> I" by auto |
|
672 |
moreover then have "emb I {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i)) = (space (Pi\<^isub>M I M))" |
|
673 |
by (auto simp: prod_emb_def space_PiM) |
|
674 |
ultimately show ?thesis |
|
675 |
using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"] |
|
676 |
by (simp add: emeasure_PiM emeasure_space_1) |
|
677 |
qed |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
678 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
679 |
|
47694 | 680 |
lemma (in product_prob_space) emeasure_PiM_emb: |
681 |
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" |
|
682 |
shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))" |
|
683 |
proof cases |
|
684 |
assume "J = {}" |
|
685 |
moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^isub>M I M)" |
|
686 |
by (auto simp: space_PiM prod_emb_def) |
|
687 |
ultimately show ?thesis |
|
688 |
by (simp add: space_PiM_empty P.emeasure_space_1) |
|
689 |
next |
|
690 |
assume "J \<noteq> {}" with X show ?thesis |
|
691 |
by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM) |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
692 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
693 |
|
47694 | 694 |
lemma (in product_prob_space) measure_PiM_emb: |
695 |
assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" |
|
696 |
shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))" |
|
697 |
using emeasure_PiM_emb[OF assms] |
|
698 |
unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal) |
|
42865 | 699 |
|
47694 | 700 |
lemma (in finite_product_prob_space) finite_measure_PiM_emb: |
701 |
"(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))" |
|
702 |
using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets_into_space, of I A M] |
|
703 |
by auto |
|
42865 | 704 |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
705 |
subsection {* Sequence space *} |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
706 |
|
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
707 |
locale sequence_space = product_prob_space M "UNIV :: nat set" for M |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
708 |
|
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
709 |
lemma (in sequence_space) infprod_in_sets[intro]: |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
710 |
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)" |
47694 | 711 |
shows "Pi UNIV E \<in> sets (Pi\<^isub>M UNIV M)" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
712 |
proof - |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
713 |
have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^isub>E j\<in>{..i}. E j))" |
47694 | 714 |
using E E[THEN sets_into_space] |
715 |
by (auto simp: prod_emb_def Pi_iff extensional_def) blast |
|
716 |
with E show ?thesis by auto |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
717 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
718 |
|
47694 | 719 |
lemma (in sequence_space) measure_PiM_countable: |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
720 |
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)" |
47694 | 721 |
shows "(\<lambda>n. \<Prod>i\<le>n. measure (M i) (E i)) ----> measure (Pi\<^isub>M UNIV M) (Pi UNIV E)" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
722 |
proof - |
46731 | 723 |
let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)" |
47694 | 724 |
have "\<And>n. (\<Prod>i\<le>n. measure (M i) (E i)) = measure (Pi\<^isub>M UNIV M) (?E n)" |
725 |
using E by (simp add: measure_PiM_emb) |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
726 |
moreover have "Pi UNIV E = (\<Inter>n. ?E n)" |
47694 | 727 |
using E E[THEN sets_into_space] |
728 |
by (auto simp: prod_emb_def extensional_def Pi_iff) blast |
|
729 |
moreover have "range ?E \<subseteq> sets (Pi\<^isub>M UNIV M)" |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
730 |
using E by auto |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
731 |
moreover have "decseq ?E" |
47694 | 732 |
by (auto simp: prod_emb_def Pi_iff decseq_def) |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
733 |
ultimately show ?thesis |
47694 | 734 |
by (simp add: finite_Lim_measure_decseq) |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
735 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
736 |
|
42147 | 737 |
end |