| author | immler@in.tum.de | 
| Tue, 06 Nov 2012 11:03:28 +0100 | |
| changeset 50038 | 8e32c9254535 | 
| parent 50003 | 8c213922ed49 | 
| child 50039 | bfd5198cbe40 | 
| 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  | 
||
| 49780 | 11  | 
lemma split_merge: "P (merge I J (x,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)"  | 
| 42147 | 12  | 
unfolding merge_def by auto  | 
13  | 
||
| 49780 | 14  | 
lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K"  | 
| 42147 | 15  | 
unfolding merge_def extensional_def by auto  | 
16  | 
||
17  | 
lemma injective_vimage_restrict:  | 
|
18  | 
assumes J: "J \<subseteq> I"  | 
|
19  | 
  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> {}"
 | 
|
20  | 
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)"  | 
|
21  | 
shows "A = B"  | 
|
22  | 
proof (intro set_eqI)  | 
|
23  | 
fix x  | 
|
24  | 
from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto  | 
|
25  | 
  have "J \<inter> (I - J) = {}" by auto
 | 
|
26  | 
show "x \<in> A \<longleftrightarrow> x \<in> B"  | 
|
27  | 
proof cases  | 
|
28  | 
assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"  | 
|
| 49780 | 29  | 
have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"  | 
| 42147 | 30  | 
using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge)  | 
31  | 
then show "x \<in> A \<longleftrightarrow> x \<in> B"  | 
|
32  | 
using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge)  | 
|
33  | 
next  | 
|
34  | 
assume "x \<notin> (\<Pi>\<^isub>E i\<in>J. S i)" with sets show "x \<in> A \<longleftrightarrow> x \<in> B" by auto  | 
|
35  | 
qed  | 
|
36  | 
qed  | 
|
37  | 
||
| 47694 | 38  | 
lemma (in product_prob_space) distr_restrict:  | 
| 42147 | 39  | 
  assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
 | 
| 47694 | 40  | 
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")  | 
41  | 
proof (rule measure_eqI_generator_eq)  | 
|
42  | 
have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset)  | 
|
43  | 
interpret J: finite_product_prob_space M J proof qed fact  | 
|
44  | 
interpret K: finite_product_prob_space M K proof qed fact  | 
|
45  | 
||
46  | 
  let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
 | 
|
47  | 
let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)"  | 
|
48  | 
let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"  | 
|
49  | 
show "Int_stable ?J"  | 
|
50  | 
by (rule Int_stable_PiE)  | 
|
| 
49784
 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 
hoelzl 
parents: 
49780 
diff
changeset
 | 
51  | 
show "range ?F \<subseteq> ?J" "(\<Union>i. ?F i) = ?\<Omega>"  | 
| 47694 | 52  | 
using `finite J` by (auto intro!: prod_algebraI_finite)  | 
53  | 
  { fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
 | 
|
54  | 
show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space)  | 
|
55  | 
show "sets (\<Pi>\<^isub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"  | 
|
56  | 
using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)  | 
|
57  | 
||
58  | 
fix X assume "X \<in> ?J"  | 
|
59  | 
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  | 
|
| 50003 | 60  | 
with `finite J` have X: "X \<in> sets (Pi\<^isub>M J M)"  | 
61  | 
by simp  | 
|
| 47694 | 62  | 
|
63  | 
have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))"  | 
|
64  | 
using E by (simp add: J.measure_times)  | 
|
65  | 
also have "\<dots> = (\<Prod> i\<in>J. emeasure (M i) (if i \<in> J then E i else space (M i)))"  | 
|
66  | 
by simp  | 
|
67  | 
also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"  | 
|
68  | 
using `finite K` `J \<subseteq> K`  | 
|
69  | 
by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1)  | 
|
70  | 
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))"  | 
|
71  | 
using E by (simp add: K.measure_times)  | 
|
72  | 
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))"  | 
|
73  | 
using `J \<subseteq> K` sets_into_space E by (force simp: Pi_iff split: split_if_asm)  | 
|
74  | 
finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X"  | 
|
75  | 
using X `J \<subseteq> K` apply (subst emeasure_distr)  | 
|
76  | 
by (auto intro!: measurable_restrict_subset simp: space_PiM)  | 
|
| 42147 | 77  | 
qed  | 
78  | 
||
| 47694 | 79  | 
abbreviation (in product_prob_space)  | 
80  | 
"emb L K X \<equiv> prod_emb L M K X"  | 
|
81  | 
||
82  | 
lemma (in product_prob_space) emeasure_prod_emb[simp]:  | 
|
83  | 
  assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)"
 | 
|
84  | 
shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X"  | 
|
85  | 
by (subst distr_restrict[OF L])  | 
|
86  | 
(simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)  | 
|
| 42147 | 87  | 
|
| 47694 | 88  | 
lemma (in product_prob_space) prod_emb_injective:  | 
89  | 
  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)"
 | 
|
90  | 
assumes "prod_emb L M J X = prod_emb L M J Y"  | 
|
91  | 
shows "X = Y"  | 
|
92  | 
proof (rule injective_vimage_restrict)  | 
|
93  | 
show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"  | 
|
94  | 
using sets[THEN sets_into_space] by (auto simp: space_PiM)  | 
|
95  | 
have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"  | 
|
| 49780 | 96  | 
using M.not_empty by auto  | 
| 47694 | 97  | 
from bchoice[OF this]  | 
98  | 
  show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
 | 
|
99  | 
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))"  | 
|
100  | 
using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)  | 
|
101  | 
qed fact  | 
|
| 42147 | 102  | 
|
| 47694 | 103  | 
definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) set set" where
 | 
104  | 
  "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))"
 | 
|
| 42147 | 105  | 
|
| 47694 | 106  | 
lemma (in product_prob_space) generatorI':  | 
107  | 
  "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"
 | 
|
108  | 
unfolding generator_def by auto  | 
|
| 42147 | 109  | 
|
| 47694 | 110  | 
lemma (in product_prob_space) algebra_generator:  | 
111  | 
  assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
 | 
|
| 47762 | 112  | 
unfolding algebra_def algebra_axioms_def ring_of_sets_iff  | 
113  | 
proof (intro conjI ballI)  | 
|
| 47694 | 114  | 
let ?G = generator  | 
115  | 
show "?G \<subseteq> Pow ?\<Omega>"  | 
|
116  | 
by (auto simp: generator_def prod_emb_def)  | 
|
117  | 
  from `I \<noteq> {}` obtain i where "i \<in> I" by auto
 | 
|
118  | 
  then show "{} \<in> ?G"
 | 
|
119  | 
    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
 | 
|
120  | 
simp: sigma_sets.Empty generator_def prod_emb_def)  | 
|
121  | 
from `i \<in> I` show "?\<Omega> \<in> ?G"  | 
|
122  | 
    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
 | 
|
123  | 
simp: generator_def prod_emb_def)  | 
|
124  | 
fix A assume "A \<in> ?G"  | 
|
125  | 
  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"
 | 
|
126  | 
by (auto simp: generator_def)  | 
|
127  | 
fix B assume "B \<in> ?G"  | 
|
128  | 
  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"
 | 
|
129  | 
by (auto simp: generator_def)  | 
|
130  | 
let ?RA = "emb (JA \<union> JB) JA XA"  | 
|
131  | 
let ?RB = "emb (JA \<union> JB) JB XB"  | 
|
132  | 
have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"  | 
|
133  | 
using XA A XB B by auto  | 
|
134  | 
show "A - B \<in> ?G" "A \<union> B \<in> ?G"  | 
|
135  | 
unfolding * using XA XB by (safe intro!: generatorI') auto  | 
|
| 42147 | 136  | 
qed  | 
137  | 
||
| 47694 | 138  | 
lemma (in product_prob_space) sets_PiM_generator:  | 
| 49804 | 139  | 
"sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"  | 
140  | 
proof cases  | 
|
141  | 
  assume "I = {}" then show ?thesis
 | 
|
142  | 
unfolding generator_def  | 
|
143  | 
by (auto simp: sets_PiM_empty sigma_sets_empty_eq cong: conj_cong)  | 
|
144  | 
next  | 
|
145  | 
  assume "I \<noteq> {}"
 | 
|
146  | 
show ?thesis  | 
|
147  | 
proof  | 
|
148  | 
show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"  | 
|
149  | 
unfolding sets_PiM  | 
|
150  | 
proof (safe intro!: sigma_sets_subseteq)  | 
|
151  | 
      fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
 | 
|
| 50003 | 152  | 
by (auto intro!: generatorI' sets_PiM_I_finite elim!: prod_algebraE)  | 
| 49804 | 153  | 
qed  | 
154  | 
qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)  | 
|
155  | 
qed  | 
|
156  | 
||
| 42147 | 157  | 
|
158  | 
lemma (in product_prob_space) generatorI:  | 
|
| 47694 | 159  | 
  "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 | 160  | 
unfolding generator_def by auto  | 
161  | 
||
162  | 
definition (in product_prob_space)  | 
|
163  | 
"\<mu>G A =  | 
|
| 47694 | 164  | 
    (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 | 165  | 
|
166  | 
lemma (in product_prob_space) \<mu>G_spec:  | 
|
167  | 
  assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
 | 
|
| 47694 | 168  | 
shows "\<mu>G A = emeasure (Pi\<^isub>M J M) X"  | 
| 42147 | 169  | 
unfolding \<mu>G_def  | 
170  | 
proof (intro the_equality allI impI ballI)  | 
|
171  | 
  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 | 172  | 
have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"  | 
| 42147 | 173  | 
using K J by simp  | 
174  | 
also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"  | 
|
| 47694 | 175  | 
using K J by (simp add: prod_emb_injective[of "K \<union> J" I])  | 
176  | 
also have "emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = emeasure (Pi\<^isub>M J M) X"  | 
|
| 42147 | 177  | 
using K J by simp  | 
| 47694 | 178  | 
finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" ..  | 
| 42147 | 179  | 
qed (insert J, force)  | 
180  | 
||
181  | 
lemma (in product_prob_space) \<mu>G_eq:  | 
|
| 47694 | 182  | 
  "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 | 183  | 
by (intro \<mu>G_spec) auto  | 
184  | 
||
185  | 
lemma (in product_prob_space) generator_Ex:  | 
|
| 47694 | 186  | 
assumes *: "A \<in> generator"  | 
187  | 
  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 | 188  | 
proof -  | 
189  | 
  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)"
 | 
|
190  | 
unfolding generator_def by auto  | 
|
191  | 
with \<mu>G_spec[OF this] show ?thesis by auto  | 
|
192  | 
qed  | 
|
193  | 
||
194  | 
lemma (in product_prob_space) generatorE:  | 
|
| 47694 | 195  | 
assumes A: "A \<in> generator"  | 
196  | 
  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 | 197  | 
proof -  | 
198  | 
  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 | 199  | 
"\<mu>G A = emeasure (Pi\<^isub>M J M) X" by auto  | 
| 42147 | 200  | 
then show thesis by (intro that) auto  | 
201  | 
qed  | 
|
202  | 
||
203  | 
lemma (in product_prob_space) merge_sets:  | 
|
| 50003 | 204  | 
  "J \<inter> K = {} \<Longrightarrow> A \<in> sets (Pi\<^isub>M (J \<union> K) M) \<Longrightarrow> x \<in> space (Pi\<^isub>M J M) \<Longrightarrow> (\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
 | 
205  | 
by simp  | 
|
| 42147 | 206  | 
|
207  | 
lemma (in product_prob_space) merge_emb:  | 
|
208  | 
assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"  | 
|
| 49780 | 209  | 
shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =  | 
210  | 
emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"  | 
|
| 42147 | 211  | 
proof -  | 
| 49780 | 212  | 
have [simp]: "\<And>x J K L. merge J K (y, restrict x L) = merge J (K \<inter> L) (y, x)"  | 
| 42147 | 213  | 
by (auto simp: restrict_def merge_def)  | 
| 49780 | 214  | 
have [simp]: "\<And>x J K L. restrict (merge J K (y, x)) L = merge (J \<inter> L) (K \<inter> L) (y, x)"  | 
| 42147 | 215  | 
by (auto simp: restrict_def merge_def)  | 
216  | 
have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto  | 
|
217  | 
have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto  | 
|
218  | 
have [simp]: "(K - J) \<inter> K = K - J" by auto  | 
|
219  | 
from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis  | 
|
| 47694 | 220  | 
by (simp split: split_merge add: prod_emb_def Pi_iff extensional_merge_sub set_eq_iff space_PiM)  | 
221  | 
auto  | 
|
| 42147 | 222  | 
qed  | 
223  | 
||
| 
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
 | 
224  | 
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
 | 
225  | 
  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
 | 
226  | 
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
 | 
227  | 
proof -  | 
| 47694 | 228  | 
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
 | 
229  | 
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
 | 
230  | 
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
 | 
231  | 
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
 | 
232  | 
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
 | 
233  | 
    have "X = {}"
 | 
| 47694 | 234  | 
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
 | 
235  | 
    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
 | 
236  | 
next  | 
| 47694 | 237  | 
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
 | 
238  | 
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
 | 
239  | 
interpret J: finite_product_sigma_finite M J by default fact  | 
| 47694 | 240  | 
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
 | 
241  | 
qed  | 
| 42147 | 242  | 
qed  | 
243  | 
||
| 
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
 | 
244  | 
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
 | 
245  | 
  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
 | 
246  | 
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
 | 
247  | 
proof -  | 
| 47694 | 248  | 
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
 | 
249  | 
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
 | 
250  | 
proof (intro additive_def[THEN iffD2] ballI impI)  | 
| 47694 | 251  | 
fix A assume "A \<in> generator" with generatorE guess J X . note J = this  | 
252  | 
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
 | 
253  | 
    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
 | 
254  | 
    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
 | 
255  | 
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
 | 
256  | 
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
 | 
257  | 
    have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
 | 
| 47694 | 258  | 
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
 | 
259  | 
      apply (insert `A \<inter> B = {}` JK J K)
 | 
| 47694 | 260  | 
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
 | 
261  | 
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
 | 
262  | 
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
 | 
263  | 
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
 | 
264  | 
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 | 265  | 
by simp  | 
266  | 
also have "\<dots> = emeasure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"  | 
|
267  | 
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
 | 
268  | 
also have "\<dots> = \<mu>G A + \<mu>G B"  | 
| 47694 | 269  | 
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
 | 
270  | 
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
 | 
271  | 
qed  | 
| 42147 | 272  | 
qed  | 
273  | 
||
| 47694 | 274  | 
lemma (in product_prob_space) emeasure_PiM_emb_not_empty:  | 
275  | 
  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
 | 
|
276  | 
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 | 277  | 
proof cases  | 
| 47694 | 278  | 
assume "finite I" with X show ?thesis by simp  | 
| 42147 | 279  | 
next  | 
| 47694 | 280  | 
let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space (M i)"  | 
| 42147 | 281  | 
let ?G = generator  | 
282  | 
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
 | 
283  | 
  then have I_not_empty: "I \<noteq> {}" by auto
 | 
| 47694 | 284  | 
interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact  | 
| 42147 | 285  | 
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
 | 
286  | 
G.additive_increasing[OF positive_\<mu>G[OF I_not_empty] additive_\<mu>G[OF I_not_empty], THEN increasingD]  | 
| 42147 | 287  | 
|
| 47694 | 288  | 
  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G"
 | 
| 42147 | 289  | 
|
290  | 
from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"  | 
|
291  | 
by (metis rev_finite_subset subsetI)  | 
|
292  | 
moreover from Z guess K' X' by (rule generatorE)  | 
|
293  | 
moreover def K \<equiv> "insert k K'"  | 
|
294  | 
moreover def X \<equiv> "emb K K' X'"  | 
|
295  | 
    ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
 | 
|
| 47694 | 296  | 
      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
 | 
| 42147 | 297  | 
by (auto simp: subset_insertI)  | 
298  | 
||
| 49780 | 299  | 
let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"  | 
| 42147 | 300  | 
    { fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
 | 
301  | 
note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]  | 
|
302  | 
moreover  | 
|
303  | 
have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"  | 
|
304  | 
using J K y by (intro merge_sets) auto  | 
|
305  | 
ultimately  | 
|
| 49780 | 306  | 
have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"  | 
| 42147 | 307  | 
using J K by (intro generatorI) auto  | 
| 49780 | 308  | 
have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"  | 
| 42147 | 309  | 
unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto  | 
310  | 
note * ** *** this }  | 
|
311  | 
note merge_in_G = this  | 
|
312  | 
||
313  | 
have "finite (K - J)" using K by auto  | 
|
314  | 
||
315  | 
interpret J: finite_product_prob_space M J by default fact+  | 
|
316  | 
interpret KmJ: finite_product_prob_space M "K - J" by default fact+  | 
|
317  | 
||
| 47694 | 318  | 
have "\<mu>G Z = emeasure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"  | 
| 42147 | 319  | 
using K J by simp  | 
| 47694 | 320  | 
also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"  | 
321  | 
using K J by (subst emeasure_fold_integral) auto  | 
|
| 49780 | 322  | 
also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"  | 
| 42147 | 323  | 
(is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")  | 
| 47694 | 324  | 
proof (intro positive_integral_cong)  | 
| 42147 | 325  | 
fix x assume x: "x \<in> space (Pi\<^isub>M J M)"  | 
326  | 
with K merge_in_G(2)[OF this]  | 
|
| 47694 | 327  | 
show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"  | 
| 42147 | 328  | 
unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \<mu>G_eq) auto  | 
329  | 
qed  | 
|
330  | 
finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" .  | 
|
331  | 
||
332  | 
    { fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
 | 
|
333  | 
then have "\<mu>G (?MZ x) \<le> 1"  | 
|
334  | 
unfolding merge_in_G(4)[OF x] `Z = emb I K X`  | 
|
335  | 
by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }  | 
|
336  | 
note le_1 = this  | 
|
337  | 
||
| 49780 | 338  | 
let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^isub>M I M))"  | 
| 42147 | 339  | 
have "?q \<in> borel_measurable (Pi\<^isub>M J M)"  | 
340  | 
unfolding `Z = emb I K X` using J K merge_in_G(3)  | 
|
| 47694 | 341  | 
by (simp add: merge_in_G \<mu>G_eq emeasure_fold_measurable cong: measurable_cong)  | 
| 42147 | 342  | 
note this fold le_1 merge_in_G(3) }  | 
343  | 
note fold = this  | 
|
344  | 
||
| 47694 | 345  | 
have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"  | 
| 42147 | 346  | 
proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G])  | 
| 47694 | 347  | 
fix A assume "A \<in> ?G"  | 
| 42147 | 348  | 
with generatorE guess J X . note JX = this  | 
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
349  | 
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
 | 
350  | 
from JX show "\<mu>G A \<noteq> \<infinity>" by simp  | 
| 42147 | 351  | 
next  | 
| 47694 | 352  | 
    fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}"
 | 
| 42147 | 353  | 
then have "decseq (\<lambda>i. \<mu>G (A i))"  | 
354  | 
by (auto intro!: \<mu>G_mono simp: decseq_def)  | 
|
355  | 
moreover  | 
|
356  | 
have "(INF i. \<mu>G (A i)) = 0"  | 
|
357  | 
proof (rule ccontr)  | 
|
358  | 
assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0")  | 
|
359  | 
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
 | 
360  | 
using A positive_\<mu>G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)  | 
| 42147 | 361  | 
ultimately have "0 < ?a" by auto  | 
362  | 
||
| 47694 | 363  | 
      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 | 364  | 
using A by (intro allI generator_Ex) auto  | 
365  | 
      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)"
 | 
|
366  | 
and A': "\<And>n. A n = emb I (J' n) (X' n)"  | 
|
367  | 
unfolding choice_iff by blast  | 
|
368  | 
moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"  | 
|
369  | 
moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"  | 
|
370  | 
      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)"
 | 
|
371  | 
by auto  | 
|
| 47694 | 372  | 
with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"  | 
373  | 
unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)  | 
|
| 42147 | 374  | 
|
375  | 
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"  | 
|
376  | 
unfolding J_def by force  | 
|
377  | 
||
378  | 
interpret J: finite_product_prob_space M "J i" for i by default fact+  | 
|
379  | 
||
380  | 
have a_le_1: "?a \<le> 1"  | 
|
381  | 
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
 | 
382  | 
by (auto intro!: INF_lower2[of 0] J.measure_le_1)  | 
| 42147 | 383  | 
|
| 49780 | 384  | 
let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)"  | 
| 42147 | 385  | 
|
| 47694 | 386  | 
      { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
 | 
387  | 
then have Z_sets: "\<And>n. Z n \<in> ?G" by auto  | 
|
| 42147 | 388  | 
        fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
 | 
389  | 
interpret J': finite_product_prob_space M J' by default fact+  | 
|
390  | 
||
| 46731 | 391  | 
let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)"  | 
392  | 
        let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^isub>M J' M)"
 | 
|
| 42147 | 393  | 
        { fix n
 | 
394  | 
have "?q n \<in> borel_measurable (Pi\<^isub>M J' M)"  | 
|
395  | 
using Z J' by (intro fold(1)) auto  | 
|
396  | 
then have "?Q n \<in> sets (Pi\<^isub>M J' M)"  | 
|
397  | 
by (rule measurable_sets) auto }  | 
|
398  | 
note Q_sets = this  | 
|
399  | 
||
| 47694 | 400  | 
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
 | 
401  | 
proof (intro INF_greatest)  | 
| 42147 | 402  | 
fix n  | 
403  | 
have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto  | 
|
404  | 
also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)"  | 
|
| 47694 | 405  | 
unfolding fold(2)[OF J' `Z n \<in> ?G`]  | 
406  | 
proof (intro positive_integral_mono)  | 
|
| 42147 | 407  | 
fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"  | 
408  | 
then have "?q n x \<le> 1 + 0"  | 
|
409  | 
using J' Z fold(3) Z_sets by auto  | 
|
410  | 
also have "\<dots> \<le> 1 + ?a / 2^(k+1)"  | 
|
411  | 
using `0 < ?a` by (intro add_mono) auto  | 
|
412  | 
finally have "?q n x \<le> 1 + ?a / 2^(k+1)" .  | 
|
413  | 
with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"  | 
|
414  | 
by (auto split: split_indicator simp del: power_Suc)  | 
|
415  | 
qed  | 
|
| 47694 | 416  | 
also have "\<dots> = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"  | 
417  | 
using `0 \<le> ?a` Q_sets J'.emeasure_space_1  | 
|
418  | 
by (subst positive_integral_add) auto  | 
|
419  | 
finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`  | 
|
420  | 
by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"])  | 
|
| 42147 | 421  | 
(auto simp: field_simps)  | 
422  | 
qed  | 
|
| 47694 | 423  | 
also have "\<dots> = emeasure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"  | 
424  | 
proof (intro INF_emeasure_decseq)  | 
|
| 42147 | 425  | 
show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto  | 
426  | 
show "decseq ?Q"  | 
|
427  | 
unfolding decseq_def  | 
|
428  | 
proof (safe intro!: vimageI[OF refl])  | 
|
429  | 
fix m n :: nat assume "m \<le> n"  | 
|
430  | 
fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"  | 
|
431  | 
assume "?a / 2^(k+1) \<le> ?q n x"  | 
|
432  | 
also have "?q n x \<le> ?q m x"  | 
|
433  | 
proof (rule \<mu>G_mono)  | 
|
434  | 
from fold(4)[OF J', OF Z_sets x]  | 
|
| 47694 | 435  | 
show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto  | 
| 42147 | 436  | 
show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"  | 
437  | 
using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto  | 
|
438  | 
qed  | 
|
439  | 
finally show "?a / 2^(k+1) \<le> ?q m x" .  | 
|
440  | 
qed  | 
|
| 47694 | 441  | 
qed simp  | 
| 42147 | 442  | 
        finally have "(\<Inter>n. ?Q n) \<noteq> {}"
 | 
443  | 
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)  | 
|
444  | 
then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }  | 
|
445  | 
note Ex_w = this  | 
|
446  | 
||
| 46731 | 447  | 
let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)"  | 
| 42147 | 448  | 
|
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
43920 
diff
changeset
 | 
449  | 
have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower)  | 
| 42147 | 450  | 
from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this  | 
451  | 
||
| 46731 | 452  | 
let ?P =  | 
453  | 
"\<lambda>k wk w. w \<in> space (Pi\<^isub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and>  | 
|
454  | 
(\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)"  | 
|
| 42147 | 455  | 
def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))"  | 
456  | 
||
457  | 
      { fix k have w: "w k \<in> space (Pi\<^isub>M (J k) M) \<and>
 | 
|
458  | 
(\<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))"  | 
|
459  | 
proof (induct k)  | 
|
460  | 
case 0 with w0 show ?case  | 
|
461  | 
unfolding w_def nat_rec_0 by auto  | 
|
462  | 
next  | 
|
463  | 
case (Suc k)  | 
|
464  | 
then have wk: "w k \<in> space (Pi\<^isub>M (J k) M)" by auto  | 
|
465  | 
have "\<exists>w'. ?P k (w k) w'"  | 
|
466  | 
proof cases  | 
|
467  | 
assume [simp]: "J k = J (Suc k)"  | 
|
468  | 
show ?thesis  | 
|
469  | 
proof (intro exI[of _ "w k"] conjI allI)  | 
|
470  | 
fix n  | 
|
471  | 
have "?a / 2 ^ (Suc k + 1) \<le> ?a / 2 ^ (k + 1)"  | 
|
472  | 
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps)  | 
|
473  | 
also have "\<dots> \<le> ?q k n (w k)" using Suc by auto  | 
|
474  | 
finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp  | 
|
475  | 
next  | 
|
476  | 
show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)"  | 
|
477  | 
using Suc by simp  | 
|
478  | 
then show "restrict (w k) (J k) = w k"  | 
|
| 47694 | 479  | 
by (simp add: extensional_restrict space_PiM)  | 
| 42147 | 480  | 
qed  | 
481  | 
next  | 
|
482  | 
assume "J k \<noteq> J (Suc k)"  | 
|
483  | 
            with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
 | 
|
| 47694 | 484  | 
have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> ?G"  | 
| 42147 | 485  | 
"decseq (\<lambda>n. ?M (J k) (A n) (w k))"  | 
486  | 
"\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))"  | 
|
487  | 
using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc  | 
|
488  | 
by (auto simp: decseq_def)  | 
|
489  | 
            from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
 | 
|
490  | 
obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)"  | 
|
491  | 
"\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto  | 
|
| 49780 | 492  | 
let ?w = "merge (J k) ?D (w k, w')"  | 
493  | 
have [simp]: "\<And>x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) =  | 
|
494  | 
merge (J (Suc k)) (I - (J (Suc k))) (?w, x)"  | 
|
| 42147 | 495  | 
using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]  | 
496  | 
by (auto intro!: ext split: split_merge)  | 
|
497  | 
have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w"  | 
|
498  | 
using w'(1) J(3)[of "Suc k"]  | 
|
| 47694 | 499  | 
by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+  | 
| 42147 | 500  | 
show ?thesis  | 
501  | 
apply (rule exI[of _ ?w])  | 
|
502  | 
using w' J_mono[of k "Suc k"] wk unfolding *  | 
|
| 47694 | 503  | 
apply (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM)  | 
| 42147 | 504  | 
apply (force simp: extensional_def)  | 
505  | 
done  | 
|
506  | 
qed  | 
|
507  | 
then have "?P k (w k) (w (Suc k))"  | 
|
508  | 
unfolding w_def nat_rec_Suc unfolding w_def[symmetric]  | 
|
509  | 
by (rule someI_ex)  | 
|
510  | 
then show ?case by auto  | 
|
511  | 
qed  | 
|
512  | 
moreover  | 
|
513  | 
then have "w k \<in> space (Pi\<^isub>M (J k) M)" by auto  | 
|
514  | 
moreover  | 
|
515  | 
from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto  | 
|
516  | 
        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
 | 
517  | 
using positive_\<mu>G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1`  | 
| 42147 | 518  | 
by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)  | 
519  | 
then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto  | 
|
| 49780 | 520  | 
then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto  | 
| 42147 | 521  | 
then have "\<exists>x\<in>A k. restrict x (J k) = w k"  | 
522  | 
using `w k \<in> space (Pi\<^isub>M (J k) M)`  | 
|
| 47694 | 523  | 
by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM)  | 
| 42147 | 524  | 
ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)"  | 
525  | 
"\<exists>x\<in>A k. restrict x (J k) = w k"  | 
|
526  | 
"k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)"  | 
|
527  | 
by auto }  | 
|
528  | 
note w = this  | 
|
529  | 
||
530  | 
      { fix k l i assume "k \<le> l" "i \<in> J k"
 | 
|
531  | 
        { fix l have "w k i = w (k + l) i"
 | 
|
532  | 
proof (induct l)  | 
|
533  | 
case (Suc l)  | 
|
534  | 
from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto  | 
|
535  | 
with w(3)[of "k + Suc l"]  | 
|
536  | 
have "w (k + l) i = w (k + Suc l) i"  | 
|
537  | 
by (auto simp: restrict_def fun_eq_iff split: split_if_asm)  | 
|
538  | 
with Suc show ?case by simp  | 
|
539  | 
qed simp }  | 
|
540  | 
from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp }  | 
|
541  | 
note w_mono = this  | 
|
542  | 
||
543  | 
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"  | 
|
544  | 
      { fix i k assume k: "i \<in> J k"
 | 
|
545  | 
have "w k i = w (LEAST k. i \<in> J k) i"  | 
|
546  | 
by (intro w_mono Least_le k LeastI[of _ k])  | 
|
547  | 
then have "w' i = w k i"  | 
|
548  | 
unfolding w'_def using k by auto }  | 
|
549  | 
note w'_eq = this  | 
|
550  | 
have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined"  | 
|
551  | 
using J by (auto simp: w'_def)  | 
|
552  | 
have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"  | 
|
553  | 
using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])  | 
|
554  | 
      { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
 | 
|
| 47694 | 555  | 
using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }  | 
| 42147 | 556  | 
note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this  | 
557  | 
||
558  | 
have w': "w' \<in> space (Pi\<^isub>M I M)"  | 
|
| 47694 | 559  | 
using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)  | 
| 42147 | 560  | 
|
561  | 
      { fix n
 | 
|
562  | 
have "restrict w' (J n) = w n" using w(1)  | 
|
| 47694 | 563  | 
by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def space_PiM)  | 
| 42147 | 564  | 
with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto  | 
| 47694 | 565  | 
then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) }  | 
| 42147 | 566  | 
then have "w' \<in> (\<Inter>i. A i)" by auto  | 
567  | 
      with `(\<Inter>i. A i) = {}` show False by auto
 | 
|
568  | 
qed  | 
|
569  | 
ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"  | 
|
| 43920 | 570  | 
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
 | 
571  | 
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
 | 
572  | 
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
 | 
573  | 
show ?thesis  | 
| 47694 | 574  | 
proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X])  | 
575  | 
    from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
 | 
|
576  | 
by (simp add: Pi_iff)  | 
|
577  | 
next  | 
|
578  | 
    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))"
 | 
|
579  | 
then show "emb I J (Pi\<^isub>E J X) \<in> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"  | 
|
580  | 
by (auto simp: Pi_iff prod_emb_def dest: sets_into_space)  | 
|
581  | 
have "emb I J (Pi\<^isub>E J X) \<in> generator"  | 
|
| 50003 | 582  | 
      using J `I \<noteq> {}` by (intro generatorI') (auto simp: Pi_iff)
 | 
| 47694 | 583  | 
then have "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))"  | 
584  | 
using \<mu> by simp  | 
|
585  | 
also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"  | 
|
586  | 
      using J  `I \<noteq> {}` by (subst \<mu>G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
 | 
|
587  | 
    also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
 | 
|
588  | 
if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"  | 
|
589  | 
      using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
 | 
|
590  | 
finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = \<dots>" .  | 
|
591  | 
next  | 
|
592  | 
let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"  | 
|
593  | 
    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)"
 | 
|
594  | 
      using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
 | 
|
595  | 
    then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
 | 
|
596  | 
emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"  | 
|
597  | 
using X by (auto simp add: emeasure_PiM)  | 
|
598  | 
next  | 
|
599  | 
show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>"  | 
|
| 49804 | 600  | 
using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def)  | 
| 42147 | 601  | 
qed  | 
602  | 
qed  | 
|
603  | 
||
| 47694 | 604  | 
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
 | 
605  | 
proof  | 
| 47694 | 606  | 
show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1"  | 
607  | 
proof cases  | 
|
608  | 
    assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
 | 
|
609  | 
next  | 
|
610  | 
    assume "I \<noteq> {}"
 | 
|
611  | 
then obtain i where "i \<in> I" by auto  | 
|
612  | 
    moreover then have "emb I {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i)) = (space (Pi\<^isub>M I M))"
 | 
|
613  | 
by (auto simp: prod_emb_def space_PiM)  | 
|
614  | 
ultimately show ?thesis  | 
|
615  | 
      using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
 | 
|
616  | 
by (simp add: emeasure_PiM emeasure_space_1)  | 
|
617  | 
qed  | 
|
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
618  | 
qed  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
619  | 
|
| 47694 | 620  | 
lemma (in product_prob_space) emeasure_PiM_emb:  | 
621  | 
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"  | 
|
622  | 
shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"  | 
|
623  | 
proof cases  | 
|
624  | 
  assume "J = {}"
 | 
|
625  | 
  moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^isub>M I M)"
 | 
|
626  | 
by (auto simp: space_PiM prod_emb_def)  | 
|
627  | 
ultimately show ?thesis  | 
|
628  | 
by (simp add: space_PiM_empty P.emeasure_space_1)  | 
|
629  | 
next  | 
|
630  | 
  assume "J \<noteq> {}" with X show ?thesis
 | 
|
631  | 
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
 | 
632  | 
qed  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
633  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
634  | 
lemma (in product_prob_space) emeasure_PiM_Collect:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
635  | 
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
636  | 
  shows "emeasure (Pi\<^isub>M I M) {x\<in>space (Pi\<^isub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
637  | 
proof -  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
638  | 
  have "{x\<in>space (Pi\<^isub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^isub>E J X)"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
639  | 
unfolding prod_emb_def using assms by (auto simp: space_PiM Pi_iff)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
640  | 
with emeasure_PiM_emb[OF assms] show ?thesis by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
641  | 
qed  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
642  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
643  | 
lemma (in product_prob_space) emeasure_PiM_Collect_single:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
644  | 
assumes X: "i \<in> I" "A \<in> sets (M i)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
645  | 
  shows "emeasure (Pi\<^isub>M I M) {x\<in>space (Pi\<^isub>M I M). x i \<in> A} = emeasure (M i) A"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
646  | 
  using emeasure_PiM_Collect[of "{i}" "\<lambda>i. A"] assms
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
647  | 
by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
648  | 
|
| 47694 | 649  | 
lemma (in product_prob_space) measure_PiM_emb:  | 
650  | 
assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"  | 
|
651  | 
shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"  | 
|
652  | 
using emeasure_PiM_emb[OF assms]  | 
|
653  | 
unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)  | 
|
| 42865 | 654  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
655  | 
lemma sets_Collect_single':  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
656  | 
  "i \<in> I \<Longrightarrow> {x\<in>space (M i). P x} \<in> sets (M i) \<Longrightarrow> {x\<in>space (PiM I M). P (x i)} \<in> sets (PiM I M)"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
657  | 
  using sets_Collect_single[of i I "{x\<in>space (M i). P x}" M]
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
658  | 
by (simp add: space_PiM Pi_iff cong: conj_cong)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
659  | 
|
| 47694 | 660  | 
lemma (in finite_product_prob_space) finite_measure_PiM_emb:  | 
661  | 
"(\<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))"  | 
|
662  | 
using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets_into_space, of I A M]  | 
|
663  | 
by auto  | 
|
| 42865 | 664  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
665  | 
lemma (in product_prob_space) PiM_component:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
666  | 
assumes "i \<in> I"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
667  | 
shows "distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i) = M i"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
668  | 
proof (rule measure_eqI[symmetric])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
669  | 
fix A assume "A \<in> sets (M i)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
670  | 
  moreover have "((\<lambda>\<omega>. \<omega> i) -` A \<inter> space (PiM I M)) = {x\<in>space (PiM I M). x i \<in> A}"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
671  | 
by auto  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
672  | 
ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i)) A"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
673  | 
by (auto simp: `i\<in>I` emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
674  | 
qed simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
675  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
676  | 
lemma (in product_prob_space) PiM_eq:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
677  | 
  assumes "I \<noteq> {}"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
678  | 
assumes "sets M' = sets (PiM I M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
679  | 
assumes eq: "\<And>J F. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>j. j \<in> J \<Longrightarrow> F j \<in> sets (M j)) \<Longrightarrow>  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
680  | 
emeasure M' (prod_emb I M J (\<Pi>\<^isub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
681  | 
shows "M' = (PiM I M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
682  | 
proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
683  | 
show "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
684  | 
by (rule sets_PiM)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
685  | 
then show "sets M' = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
686  | 
unfolding `sets M' = sets (PiM I M)` by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
687  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
688  | 
def i \<equiv> "SOME i. i \<in> I"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
689  | 
  with `I \<noteq> {}` have i: "i \<in> I"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
690  | 
by (auto intro: someI_ex)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
691  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
692  | 
  def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^isub>E j\<in>{i}. space (M i))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
693  | 
then show "range A \<subseteq> prod_algebra I M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
694  | 
by (auto intro!: prod_algebraI i)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
695  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
696  | 
have A_eq: "\<And>i. A i = space (PiM I M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
697  | 
by (auto simp: prod_emb_def space_PiM Pi_iff A_def i)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
698  | 
show "(\<Union>i. A i) = (\<Pi>\<^isub>E i\<in>I. space (M i))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
699  | 
unfolding A_eq by (auto simp: space_PiM)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
700  | 
show "\<And>i. emeasure (PiM I M) (A i) \<noteq> \<infinity>"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
701  | 
unfolding A_eq P.emeasure_space_1 by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
702  | 
next  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
703  | 
fix X assume X: "X \<in> prod_algebra I M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
704  | 
then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
705  | 
and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
706  | 
by (force elim!: prod_algebraE)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
707  | 
from eq[OF J] have "emeasure M' X = (\<Prod>j\<in>J. emeasure (M j) (E j))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
708  | 
by (simp add: X)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
709  | 
also have "\<dots> = emeasure (PiM I M) X"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
710  | 
unfolding X using J by (intro emeasure_PiM_emb[symmetric]) auto  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
711  | 
finally show "emeasure (PiM I M) X = emeasure M' X" ..  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
712  | 
qed  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
713  | 
|
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
714  | 
subsection {* Sequence space *}
 | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
715  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
716  | 
lemma measurable_nat_case: "(\<lambda>(x, \<omega>). nat_case x \<omega>) \<in> measurable (M \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
717  | 
proof (rule measurable_PiM_single)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
718  | 
show "(\<lambda>(x, \<omega>). nat_case x \<omega>) \<in> space (M \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^isub>E space M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
719  | 
by (auto simp: space_pair_measure space_PiM Pi_iff split: nat.split)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
720  | 
fix i :: nat and A assume A: "A \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
721  | 
  then have *: "{\<omega> \<in> space (M \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case nat_case \<omega> i \<in> A} =
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
722  | 
    (case i of 0 \<Rightarrow> A \<times> space (\<Pi>\<^isub>M i\<in>UNIV. M) | Suc n \<Rightarrow> space M \<times> {\<omega>\<in>space (\<Pi>\<^isub>M i\<in>UNIV. M). \<omega> n \<in> A})"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
723  | 
by (auto simp: space_PiM space_pair_measure split: nat.split dest: sets_into_space)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
724  | 
  show "{\<omega> \<in> space (M \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case nat_case \<omega> i \<in> A} \<in> sets (M \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
725  | 
unfolding * by (auto simp: A split: nat.split intro!: sets_Collect_single)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
726  | 
qed  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
727  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
728  | 
lemma measurable_nat_case':  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
729  | 
assumes f: "f \<in> measurable N M" and g: "g \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
730  | 
shows "(\<lambda>x. nat_case (f x) (g x)) \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
731  | 
using measurable_compose[OF measurable_Pair[OF f g] measurable_nat_case] by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
732  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
733  | 
definition comb_seq :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" where  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
734  | 
"comb_seq i \<omega> \<omega>' j = (if j < i then \<omega> j else \<omega>' (j - i))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
735  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
736  | 
lemma split_comb_seq: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> (j < i \<longrightarrow> P (\<omega> j)) \<and> (\<forall>k. j = i + k \<longrightarrow> P (\<omega>' k))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
737  | 
by (auto simp: comb_seq_def not_less)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
738  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
739  | 
lemma split_comb_seq_asm: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> \<not> ((j < i \<and> \<not> P (\<omega> j)) \<or> (\<exists>k. j = i + k \<and> \<not> P (\<omega>' k)))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
740  | 
by (auto simp: comb_seq_def)  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
741  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
742  | 
lemma measurable_comb_seq: "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
743  | 
proof (rule measurable_PiM_single)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
744  | 
show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^isub>E space M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
745  | 
by (auto simp: space_pair_measure space_PiM Pi_iff split: split_comb_seq)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
746  | 
fix j :: nat and A assume A: "A \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
747  | 
  then have *: "{\<omega> \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} =
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
748  | 
    (if j < i then {\<omega> \<in> space (\<Pi>\<^isub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^isub>M i\<in>UNIV. M)
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
749  | 
              else space (\<Pi>\<^isub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^isub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
750  | 
by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets_into_space)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
751  | 
  show "{\<omega> \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
752  | 
unfolding * by (auto simp: A intro!: sets_Collect_single)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
753  | 
qed  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
754  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
755  | 
lemma measurable_comb_seq':  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
756  | 
assumes f: "f \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
757  | 
shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
758  | 
using measurable_compose[OF measurable_Pair[OF f g] measurable_comb_seq] by simp  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
759  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
760  | 
locale sequence_space = product_prob_space "\<lambda>i. M" "UNIV :: nat set" for M  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
761  | 
begin  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
762  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
763  | 
abbreviation "S \<equiv> \<Pi>\<^isub>M i\<in>UNIV::nat set. M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
764  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
765  | 
lemma infprod_in_sets[intro]:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
766  | 
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
767  | 
shows "Pi UNIV E \<in> sets S"  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
768  | 
proof -  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
769  | 
  have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^isub>E j\<in>{..i}. E j))"
 | 
| 47694 | 770  | 
using E E[THEN sets_into_space]  | 
771  | 
by (auto simp: prod_emb_def Pi_iff extensional_def) blast  | 
|
772  | 
with E show ?thesis by auto  | 
|
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
773  | 
qed  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
774  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
775  | 
lemma measure_PiM_countable:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
776  | 
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
777  | 
shows "(\<lambda>n. \<Prod>i\<le>n. measure M (E i)) ----> measure S (Pi UNIV E)"  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
778  | 
proof -  | 
| 46731 | 779  | 
  let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)"
 | 
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
780  | 
have "\<And>n. (\<Prod>i\<le>n. measure M (E i)) = measure S (?E n)"  | 
| 47694 | 781  | 
using E by (simp add: measure_PiM_emb)  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
782  | 
moreover have "Pi UNIV E = (\<Inter>n. ?E n)"  | 
| 47694 | 783  | 
using E E[THEN sets_into_space]  | 
784  | 
by (auto simp: prod_emb_def extensional_def Pi_iff) blast  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
785  | 
moreover have "range ?E \<subseteq> sets S"  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
786  | 
using E by auto  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
787  | 
moreover have "decseq ?E"  | 
| 47694 | 788  | 
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
 | 
789  | 
ultimately show ?thesis  | 
| 47694 | 790  | 
by (simp add: finite_Lim_measure_decseq)  | 
| 
42257
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
791  | 
qed  | 
| 
 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 
hoelzl 
parents: 
42166 
diff
changeset
 | 
792  | 
|
| 
50000
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
793  | 
lemma nat_eq_diff_eq:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
794  | 
fixes a b c :: nat  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
795  | 
shows "c \<le> b \<Longrightarrow> a = b - c \<longleftrightarrow> a + c = b"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
796  | 
by auto  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
797  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
798  | 
lemma PiM_comb_seq:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
799  | 
"distr (S \<Otimes>\<^isub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _")  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
800  | 
proof (rule PiM_eq)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
801  | 
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
802  | 
let "distr _ _ ?f" = "?D"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
803  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
804  | 
fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
805  | 
let ?X = "prod_emb ?I ?M J (\<Pi>\<^isub>E j\<in>J. E j)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
806  | 
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
807  | 
using J(3)[THEN sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
808  | 
with J have "?f -` ?X \<inter> space (S \<Otimes>\<^isub>M S) =  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
809  | 
    (prod_emb ?I ?M (J \<inter> {..<i}) (PIE j:J \<inter> {..<i}. E j)) \<times>
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
810  | 
(prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \<times> ?F")  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
811  | 
by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib Pi_iff  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
812  | 
split: split_comb_seq split_comb_seq_asm)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
813  | 
then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^isub>M S) (?E \<times> ?F)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
814  | 
by (subst emeasure_distr[OF measurable_comb_seq])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
815  | 
(auto intro!: sets_PiM_I simp: split_beta' J)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
816  | 
also have "\<dots> = emeasure S ?E * emeasure S ?F"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
817  | 
using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI simp: inj_on_def)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
818  | 
also have "emeasure S ?F = (\<Prod>j\<in>(op + i) -` J. emeasure M (E (i + j)))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
819  | 
using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI inj_on_def)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
820  | 
  also have "\<dots> = (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
821  | 
by (rule strong_setprod_reindex_cong[where f="\<lambda>x. x - i"])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
822  | 
(auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
823  | 
  also have "emeasure S ?E = (\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
824  | 
using J by (intro emeasure_PiM_emb) simp_all  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
825  | 
  also have "(\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j)) * (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
826  | 
by (subst mult_commute) (auto simp: J setprod_subset_diff[symmetric])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
827  | 
finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" .  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
828  | 
qed simp_all  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
829  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
830  | 
lemma PiM_iter:  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
831  | 
"distr (M \<Otimes>\<^isub>M S) S (\<lambda>(s, \<omega>). nat_case s \<omega>) = S" (is "?D = _")  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
832  | 
proof (rule PiM_eq)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
833  | 
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
834  | 
let "distr _ _ ?f" = "?D"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
835  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
836  | 
fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
837  | 
let ?X = "prod_emb ?I ?M J (PIE j:J. E j)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
838  | 
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
839  | 
using J(3)[THEN sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
840  | 
with J have "?f -` ?X \<inter> space (M \<Otimes>\<^isub>M S) = (if 0 \<in> J then E 0 else space M) \<times>  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
841  | 
(prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \<times> ?F")  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
842  | 
by (auto simp: space_pair_measure space_PiM Pi_iff prod_emb_def all_conj_distrib  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
843  | 
split: nat.split nat.split_asm)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
844  | 
then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^isub>M S) (?E \<times> ?F)"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
845  | 
by (subst emeasure_distr[OF measurable_nat_case])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
846  | 
(auto intro!: sets_PiM_I simp: split_beta' J)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
847  | 
also have "\<dots> = emeasure M ?E * emeasure S ?F"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
848  | 
using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
849  | 
also have "emeasure S ?F = (\<Prod>j\<in>Suc -` J. emeasure M (E (Suc j)))"  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
850  | 
using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
851  | 
  also have "\<dots> = (\<Prod>j\<in>J - {0}. emeasure M (E j))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
852  | 
by (rule strong_setprod_reindex_cong[where f="\<lambda>x. x - 1"])  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
853  | 
(auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
854  | 
  also have "emeasure M ?E * (\<Prod>j\<in>J - {0}. emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
 | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
855  | 
by (auto simp: M.emeasure_space_1 setprod.remove J)  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
856  | 
finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" .  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
857  | 
qed simp_all  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
858  | 
|
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
859  | 
end  | 
| 
 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 
hoelzl 
parents: 
49804 
diff
changeset
 | 
860  | 
|
| 42147 | 861  | 
end  |