author | haftmann |
Sat, 28 Dec 2013 17:51:54 +0100 | |
changeset 54870 | 1b5f2485757b |
parent 53374 | a14d2a854c02 |
child 55414 | eab03e9cee8a |
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 |
|
50039
bfd5198cbe40
added projective_family; generalized generator in product_prob_space to projective_family
immler@in.tum.de
parents:
50038
diff
changeset
|
8 |
imports Probability_Measure Caratheodory Projective_Family |
42147 | 9 |
begin |
10 |
||
47694 | 11 |
lemma (in product_prob_space) emeasure_PiM_emb_not_empty: |
12 |
assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
13 |
shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)" |
42147 | 14 |
proof cases |
47694 | 15 |
assume "finite I" with X show ?thesis by simp |
42147 | 16 |
next |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
17 |
let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space (M i)" |
42147 | 18 |
let ?G = generator |
19 |
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
|
20 |
then have I_not_empty: "I \<noteq> {}" by auto |
47694 | 21 |
interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact |
50252 | 22 |
note mu_G_mono = |
23 |
G.additive_increasing[OF positive_mu_G[OF I_not_empty] additive_mu_G[OF I_not_empty], |
|
24 |
THEN increasingD] |
|
25 |
write mu_G ("\<mu>G") |
|
42147 | 26 |
|
47694 | 27 |
{ fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G" |
42147 | 28 |
|
29 |
from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J" |
|
30 |
by (metis rev_finite_subset subsetI) |
|
31 |
moreover from Z guess K' X' by (rule generatorE) |
|
32 |
moreover def K \<equiv> "insert k K'" |
|
33 |
moreover def X \<equiv> "emb K K' X'" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
34 |
ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^sub>M K M)" "Z = emb I K X" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
35 |
"K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^sub>M K M) X" |
42147 | 36 |
by (auto simp: subset_insertI) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
37 |
let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
38 |
{ fix y assume y: "y \<in> space (Pi\<^sub>M J M)" |
42147 | 39 |
note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X] |
40 |
moreover |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
41 |
have **: "?M y \<in> sets (Pi\<^sub>M (K - J) M)" |
42147 | 42 |
using J K y by (intro merge_sets) auto |
43 |
ultimately |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
44 |
have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<in> ?G" |
42147 | 45 |
using J K by (intro generatorI) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
46 |
have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) = emeasure (Pi\<^sub>M (K - J) M) (?M y)" |
50252 | 47 |
unfolding * using K J by (subst mu_G_eq[OF _ _ _ **]) auto |
42147 | 48 |
note * ** *** this } |
49 |
note merge_in_G = this |
|
50 |
||
51 |
have "finite (K - J)" using K by auto |
|
52 |
||
53 |
interpret J: finite_product_prob_space M J by default fact+ |
|
54 |
interpret KmJ: finite_product_prob_space M "K - J" by default fact+ |
|
55 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
56 |
have "\<mu>G Z = emeasure (Pi\<^sub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)" |
42147 | 57 |
using K J by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
58 |
also have "\<dots> = (\<integral>\<^sup>+ x. emeasure (Pi\<^sub>M (K - J) M) (?M x) \<partial>Pi\<^sub>M J M)" |
47694 | 59 |
using K J by (subst emeasure_fold_integral) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
60 |
also have "\<dots> = (\<integral>\<^sup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<partial>Pi\<^sub>M J M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
61 |
(is "_ = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)") |
47694 | 62 |
proof (intro positive_integral_cong) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
63 |
fix x assume x: "x \<in> space (Pi\<^sub>M J M)" |
42147 | 64 |
with K merge_in_G(2)[OF this] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
65 |
show "emeasure (Pi\<^sub>M (K - J) M) (?M x) = \<mu>G (?MZ x)" |
50252 | 66 |
unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst mu_G_eq) auto |
42147 | 67 |
qed |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
68 |
finally have fold: "\<mu>G Z = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)" . |
42147 | 69 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
70 |
{ fix x assume x: "x \<in> space (Pi\<^sub>M J M)" |
42147 | 71 |
then have "\<mu>G (?MZ x) \<le> 1" |
72 |
unfolding merge_in_G(4)[OF x] `Z = emb I K X` |
|
73 |
by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) } |
|
74 |
note le_1 = this |
|
75 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
76 |
let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^sub>M I M))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
77 |
have "?q \<in> borel_measurable (Pi\<^sub>M J M)" |
42147 | 78 |
unfolding `Z = emb I K X` using J K merge_in_G(3) |
50252 | 79 |
by (simp add: merge_in_G mu_G_eq emeasure_fold_measurable cong: measurable_cong) |
42147 | 80 |
note this fold le_1 merge_in_G(3) } |
81 |
note fold = this |
|
82 |
||
47694 | 83 |
have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>" |
50252 | 84 |
proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G]) |
47694 | 85 |
fix A assume "A \<in> ?G" |
42147 | 86 |
with generatorE guess J X . note JX = this |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
87 |
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
|
88 |
from JX show "\<mu>G A \<noteq> \<infinity>" by simp |
42147 | 89 |
next |
47694 | 90 |
fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}" |
42147 | 91 |
then have "decseq (\<lambda>i. \<mu>G (A i))" |
50252 | 92 |
by (auto intro!: mu_G_mono simp: decseq_def) |
42147 | 93 |
moreover |
94 |
have "(INF i. \<mu>G (A i)) = 0" |
|
95 |
proof (rule ccontr) |
|
96 |
assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0") |
|
97 |
moreover have "0 \<le> ?a" |
|
50252 | 98 |
using A positive_mu_G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def) |
42147 | 99 |
ultimately have "0 < ?a" by auto |
100 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
101 |
have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (limP J M (\<lambda>J. (Pi\<^sub>M J M))) X" |
42147 | 102 |
using A by (intro allI generator_Ex) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
103 |
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\<^sub>M (J' n) M)" |
42147 | 104 |
and A': "\<And>n. A n = emb I (J' n) (X' n)" |
105 |
unfolding choice_iff by blast |
|
106 |
moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)" |
|
107 |
moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
108 |
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\<^sub>M (J n) M)" |
42147 | 109 |
by auto |
47694 | 110 |
with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G" |
111 |
unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto) |
|
42147 | 112 |
|
113 |
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m" |
|
114 |
unfolding J_def by force |
|
115 |
||
116 |
interpret J: finite_product_prob_space M "J i" for i by default fact+ |
|
117 |
||
118 |
have a_le_1: "?a \<le> 1" |
|
50252 | 119 |
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
|
120 |
by (auto intro!: INF_lower2[of 0] J.measure_le_1) |
42147 | 121 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
122 |
let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)" |
42147 | 123 |
|
47694 | 124 |
{ fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)" |
125 |
then have Z_sets: "\<And>n. Z n \<in> ?G" by auto |
|
42147 | 126 |
fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I" |
127 |
interpret J': finite_product_prob_space M J' by default fact+ |
|
128 |
||
46731 | 129 |
let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
130 |
let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^sub>M J' M)" |
42147 | 131 |
{ fix n |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
132 |
have "?q n \<in> borel_measurable (Pi\<^sub>M J' M)" |
42147 | 133 |
using Z J' by (intro fold(1)) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
134 |
then have "?Q n \<in> sets (Pi\<^sub>M J' M)" |
42147 | 135 |
by (rule measurable_sets) auto } |
136 |
note Q_sets = this |
|
137 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
138 |
have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^sub>M J' M) (?Q n))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
43920
diff
changeset
|
139 |
proof (intro INF_greatest) |
42147 | 140 |
fix n |
141 |
have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
142 |
also have "\<dots> \<le> (\<integral>\<^sup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^sub>M J' M)" |
47694 | 143 |
unfolding fold(2)[OF J' `Z n \<in> ?G`] |
144 |
proof (intro positive_integral_mono) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
145 |
fix x assume x: "x \<in> space (Pi\<^sub>M J' M)" |
42147 | 146 |
then have "?q n x \<le> 1 + 0" |
147 |
using J' Z fold(3) Z_sets by auto |
|
148 |
also have "\<dots> \<le> 1 + ?a / 2^(k+1)" |
|
149 |
using `0 < ?a` by (intro add_mono) auto |
|
150 |
finally have "?q n x \<le> 1 + ?a / 2^(k+1)" . |
|
151 |
with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)" |
|
152 |
by (auto split: split_indicator simp del: power_Suc) |
|
153 |
qed |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
154 |
also have "\<dots> = emeasure (Pi\<^sub>M J' M) (?Q n) + ?a / 2^(k+1)" |
47694 | 155 |
using `0 \<le> ?a` Q_sets J'.emeasure_space_1 |
156 |
by (subst positive_integral_add) auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
157 |
finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^sub>M J' M) (?Q n)" using `?a \<le> 1` |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
158 |
by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^sub>M J' M) (?Q n)"]) |
42147 | 159 |
(auto simp: field_simps) |
160 |
qed |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
161 |
also have "\<dots> = emeasure (Pi\<^sub>M J' M) (\<Inter>n. ?Q n)" |
47694 | 162 |
proof (intro INF_emeasure_decseq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
163 |
show "range ?Q \<subseteq> sets (Pi\<^sub>M J' M)" using Q_sets by auto |
42147 | 164 |
show "decseq ?Q" |
165 |
unfolding decseq_def |
|
166 |
proof (safe intro!: vimageI[OF refl]) |
|
167 |
fix m n :: nat assume "m \<le> n" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
168 |
fix x assume x: "x \<in> space (Pi\<^sub>M J' M)" |
42147 | 169 |
assume "?a / 2^(k+1) \<le> ?q n x" |
170 |
also have "?q n x \<le> ?q m x" |
|
50252 | 171 |
proof (rule mu_G_mono) |
42147 | 172 |
from fold(4)[OF J', OF Z_sets x] |
47694 | 173 |
show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto |
42147 | 174 |
show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x" |
175 |
using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto |
|
176 |
qed |
|
177 |
finally show "?a / 2^(k+1) \<le> ?q m x" . |
|
178 |
qed |
|
47694 | 179 |
qed simp |
42147 | 180 |
finally have "(\<Inter>n. ?Q n) \<noteq> {}" |
181 |
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
182 |
then have "\<exists>w\<in>space (Pi\<^sub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto } |
42147 | 183 |
note Ex_w = this |
184 |
||
46731 | 185 |
let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)" |
42147 | 186 |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
43920
diff
changeset
|
187 |
have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower) |
42147 | 188 |
from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this |
189 |
||
46731 | 190 |
let ?P = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
191 |
"\<lambda>k wk w. w \<in> space (Pi\<^sub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and> |
46731 | 192 |
(\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)" |
42147 | 193 |
def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))" |
194 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
195 |
{ fix k have w: "w k \<in> space (Pi\<^sub>M (J k) M) \<and> |
42147 | 196 |
(\<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))" |
197 |
proof (induct k) |
|
198 |
case 0 with w0 show ?case |
|
199 |
unfolding w_def nat_rec_0 by auto |
|
200 |
next |
|
201 |
case (Suc k) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
202 |
then have wk: "w k \<in> space (Pi\<^sub>M (J k) M)" by auto |
42147 | 203 |
have "\<exists>w'. ?P k (w k) w'" |
204 |
proof cases |
|
205 |
assume [simp]: "J k = J (Suc k)" |
|
206 |
show ?thesis |
|
207 |
proof (intro exI[of _ "w k"] conjI allI) |
|
208 |
fix n |
|
209 |
have "?a / 2 ^ (Suc k + 1) \<le> ?a / 2 ^ (k + 1)" |
|
210 |
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps) |
|
211 |
also have "\<dots> \<le> ?q k n (w k)" using Suc by auto |
|
212 |
finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp |
|
213 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
214 |
show "w k \<in> space (Pi\<^sub>M (J (Suc k)) M)" |
42147 | 215 |
using Suc by simp |
216 |
then show "restrict (w k) (J k) = w k" |
|
47694 | 217 |
by (simp add: extensional_restrict space_PiM) |
42147 | 218 |
qed |
219 |
next |
|
220 |
assume "J k \<noteq> J (Suc k)" |
|
221 |
with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto |
|
47694 | 222 |
have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> ?G" |
42147 | 223 |
"decseq (\<lambda>n. ?M (J k) (A n) (w k))" |
224 |
"\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))" |
|
225 |
using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc |
|
226 |
by (auto simp: decseq_def) |
|
227 |
from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"] |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
228 |
obtain w' where w': "w' \<in> space (Pi\<^sub>M ?D M)" |
42147 | 229 |
"\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto |
49780 | 230 |
let ?w = "merge (J k) ?D (w k, w')" |
231 |
have [simp]: "\<And>x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) = |
|
232 |
merge (J (Suc k)) (I - (J (Suc k))) (?w, x)" |
|
42147 | 233 |
using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"] |
234 |
by (auto intro!: ext split: split_merge) |
|
235 |
have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w" |
|
236 |
using w'(1) J(3)[of "Suc k"] |
|
47694 | 237 |
by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+ |
42147 | 238 |
show ?thesis |
239 |
using w' J_mono[of k "Suc k"] wk unfolding * |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
240 |
by (intro exI[of _ ?w]) |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
241 |
(auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM PiE_iff) |
42147 | 242 |
qed |
243 |
then have "?P k (w k) (w (Suc k))" |
|
244 |
unfolding w_def nat_rec_Suc unfolding w_def[symmetric] |
|
245 |
by (rule someI_ex) |
|
246 |
then show ?case by auto |
|
247 |
qed |
|
248 |
moreover |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
249 |
from w have "w k \<in> space (Pi\<^sub>M (J k) M)" by auto |
42147 | 250 |
moreover |
251 |
from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto |
|
252 |
then have "?M (J k) (A k) (w k) \<noteq> {}" |
|
50252 | 253 |
using positive_mu_G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1` |
42147 | 254 |
by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) |
255 |
then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto |
|
49780 | 256 |
then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto |
42147 | 257 |
then have "\<exists>x\<in>A k. restrict x (J k) = w k" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
258 |
using `w k \<in> space (Pi\<^sub>M (J k) M)` |
47694 | 259 |
by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
260 |
ultimately have "w k \<in> space (Pi\<^sub>M (J k) M)" |
42147 | 261 |
"\<exists>x\<in>A k. restrict x (J k) = w k" |
262 |
"k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)" |
|
263 |
by auto } |
|
264 |
note w = this |
|
265 |
||
266 |
{ fix k l i assume "k \<le> l" "i \<in> J k" |
|
267 |
{ fix l have "w k i = w (k + l) i" |
|
268 |
proof (induct l) |
|
269 |
case (Suc l) |
|
270 |
from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto |
|
271 |
with w(3)[of "k + Suc l"] |
|
272 |
have "w (k + l) i = w (k + Suc l) i" |
|
273 |
by (auto simp: restrict_def fun_eq_iff split: split_if_asm) |
|
274 |
with Suc show ?case by simp |
|
275 |
qed simp } |
|
276 |
from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp } |
|
277 |
note w_mono = this |
|
278 |
||
279 |
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" |
|
280 |
{ fix i k assume k: "i \<in> J k" |
|
281 |
have "w k i = w (LEAST k. i \<in> J k) i" |
|
282 |
by (intro w_mono Least_le k LeastI[of _ k]) |
|
283 |
then have "w' i = w k i" |
|
284 |
unfolding w'_def using k by auto } |
|
285 |
note w'_eq = this |
|
286 |
have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined" |
|
287 |
using J by (auto simp: w'_def) |
|
288 |
have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)" |
|
289 |
using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]]) |
|
290 |
{ fix i assume "i \<in> I" then have "w' i \<in> space (M i)" |
|
47694 | 291 |
using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ } |
42147 | 292 |
note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this |
293 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
294 |
have w': "w' \<in> space (Pi\<^sub>M I M)" |
47694 | 295 |
using w(1) by (auto simp add: Pi_iff extensional_def space_PiM) |
42147 | 296 |
|
297 |
{ fix n |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
298 |
have "restrict w' (J n) = w n" using w(1)[of n] |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
299 |
by (auto simp add: fun_eq_iff space_PiM) |
42147 | 300 |
with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto |
47694 | 301 |
then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) } |
42147 | 302 |
then have "w' \<in> (\<Inter>i. A i)" by auto |
303 |
with `(\<Inter>i. A i) = {}` show False by auto |
|
304 |
qed |
|
305 |
ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0" |
|
51351 | 306 |
using LIMSEQ_INF[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
|
307 |
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
|
308 |
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
|
309 |
show ?thesis |
47694 | 310 |
proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X]) |
311 |
from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))" |
|
312 |
by (simp add: Pi_iff) |
|
313 |
next |
|
314 |
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))" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
315 |
then show "emb I J (Pi\<^sub>E J X) \<in> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
316 |
by (auto simp: Pi_iff prod_emb_def dest: sets.sets_into_space) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
317 |
have "emb I J (Pi\<^sub>E J X) \<in> generator" |
50003 | 318 |
using J `I \<noteq> {}` by (intro generatorI') (auto simp: Pi_iff) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
319 |
then have "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))" |
47694 | 320 |
using \<mu> by simp |
321 |
also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" |
|
50252 | 322 |
using J `I \<noteq> {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff) |
47694 | 323 |
also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. |
324 |
if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" |
|
325 |
using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
326 |
finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = \<dots>" . |
47694 | 327 |
next |
328 |
let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))" |
|
329 |
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)" |
|
330 |
using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) |
|
331 |
then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
332 |
emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)" |
47694 | 333 |
using X by (auto simp add: emeasure_PiM) |
334 |
next |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
335 |
show "positive (sets (Pi\<^sub>M I M)) \<mu>" "countably_additive (sets (Pi\<^sub>M I M)) \<mu>" |
49804 | 336 |
using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def) |
42147 | 337 |
qed |
338 |
qed |
|
339 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
340 |
sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^sub>M I M" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
341 |
proof |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
342 |
show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1" |
47694 | 343 |
proof cases |
344 |
assume "I = {}" then show ?thesis by (simp add: space_PiM_empty) |
|
345 |
next |
|
346 |
assume "I \<noteq> {}" |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
347 |
then obtain i where i: "i \<in> I" by auto |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
348 |
then have "emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)) = (space (Pi\<^sub>M I M))" |
47694 | 349 |
by (auto simp: prod_emb_def space_PiM) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
350 |
with i show ?thesis |
47694 | 351 |
using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"] |
352 |
by (simp add: emeasure_PiM emeasure_space_1) |
|
353 |
qed |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
354 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
355 |
|
47694 | 356 |
lemma (in product_prob_space) emeasure_PiM_emb: |
357 |
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
358 |
shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))" |
47694 | 359 |
proof cases |
360 |
assume "J = {}" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
361 |
moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^sub>M I M)" |
47694 | 362 |
by (auto simp: space_PiM prod_emb_def) |
363 |
ultimately show ?thesis |
|
364 |
by (simp add: space_PiM_empty P.emeasure_space_1) |
|
365 |
next |
|
366 |
assume "J \<noteq> {}" with X show ?thesis |
|
367 |
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
|
368 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
369 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
370 |
lemma (in product_prob_space) emeasure_PiM_Collect: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
371 |
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
372 |
shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
373 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
374 |
have "{x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^sub>E J X)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
375 |
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
|
376 |
with emeasure_PiM_emb[OF assms] show ?thesis by simp |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
377 |
qed |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
378 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
379 |
lemma (in product_prob_space) emeasure_PiM_Collect_single: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
380 |
assumes X: "i \<in> I" "A \<in> sets (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
381 |
shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). x i \<in> A} = emeasure (M i) A" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
382 |
using emeasure_PiM_Collect[of "{i}" "\<lambda>i. A"] assms |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
383 |
by simp |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
384 |
|
47694 | 385 |
lemma (in product_prob_space) measure_PiM_emb: |
386 |
assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
387 |
shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))" |
47694 | 388 |
using emeasure_PiM_emb[OF assms] |
389 |
unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal) |
|
42865 | 390 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
391 |
lemma sets_Collect_single': |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
392 |
"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
|
393 |
using sets_Collect_single[of i I "{x\<in>space (M i). P x}" M] |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
394 |
by (simp add: space_PiM PiE_iff cong: conj_cong) |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
395 |
|
47694 | 396 |
lemma (in finite_product_prob_space) finite_measure_PiM_emb: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
397 |
"(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
398 |
using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M] |
47694 | 399 |
by auto |
42865 | 400 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
401 |
lemma (in product_prob_space) PiM_component: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
402 |
assumes "i \<in> I" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
403 |
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
|
404 |
proof (rule measure_eqI[symmetric]) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
405 |
fix A assume "A \<in> sets (M i)" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
406 |
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
|
407 |
by auto |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
408 |
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
|
409 |
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
|
410 |
qed simp |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
411 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
412 |
lemma (in product_prob_space) PiM_eq: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
413 |
assumes "I \<noteq> {}" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
414 |
assumes "sets M' = sets (PiM I M)" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
415 |
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> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
416 |
emeasure M' (prod_emb I M J (\<Pi>\<^sub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
417 |
shows "M' = (PiM I M)" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
418 |
proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
419 |
show "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
420 |
by (rule sets_PiM) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
421 |
then show "sets M' = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
422 |
unfolding `sets M' = sets (PiM I M)` by simp |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
423 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
424 |
def i \<equiv> "SOME i. i \<in> I" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
425 |
with `I \<noteq> {}` have i: "i \<in> I" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
426 |
by (auto intro: someI_ex) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
427 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
428 |
def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. space (M i))" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
429 |
then show "range A \<subseteq> prod_algebra I M" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
430 |
by (auto intro!: prod_algebraI i) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
431 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
432 |
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
|
433 |
by (auto simp: prod_emb_def space_PiM Pi_iff A_def i) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
434 |
show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
435 |
unfolding A_eq by (auto simp: space_PiM) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
436 |
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
|
437 |
unfolding A_eq P.emeasure_space_1 by simp |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
438 |
next |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
439 |
fix X assume X: "X \<in> prod_algebra I M" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
440 |
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
|
441 |
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
|
442 |
by (force elim!: prod_algebraE) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
443 |
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
|
444 |
by (simp add: X) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
445 |
also have "\<dots> = emeasure (PiM I M) X" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
446 |
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
|
447 |
finally show "emeasure (PiM I M) X = emeasure M' X" .. |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
448 |
qed |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
449 |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
450 |
subsection {* Sequence space *} |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
451 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
452 |
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
|
453 |
"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
|
454 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
455 |
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
|
456 |
by (auto simp: comb_seq_def not_less) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
457 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
458 |
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
|
459 |
by (auto simp: comb_seq_def) |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
460 |
|
50099 | 461 |
lemma measurable_comb_seq: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
462 |
"(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) (\<Pi>\<^sub>M i\<in>UNIV. M)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
463 |
proof (rule measurable_PiM_single) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
464 |
show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^sub>E space M)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
465 |
by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq) |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
466 |
fix j :: nat and A assume A: "A \<in> sets M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
467 |
then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
468 |
(if j < i then {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^sub>M i\<in>UNIV. M) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
469 |
else space (\<Pi>\<^sub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
470 |
by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
471 |
show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
472 |
unfolding * by (auto simp: A intro!: sets_Collect_single) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
473 |
qed |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
474 |
|
50099 | 475 |
lemma measurable_comb_seq'[measurable (raw)]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
476 |
assumes f: "f \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
477 |
shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
478 |
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
|
479 |
|
50099 | 480 |
lemma comb_seq_0: "comb_seq 0 \<omega> \<omega>' = \<omega>'" |
481 |
by (auto simp add: comb_seq_def) |
|
482 |
||
483 |
lemma comb_seq_Suc: "comb_seq (Suc n) \<omega> \<omega>' = comb_seq n \<omega> (nat_case (\<omega> n) \<omega>')" |
|
484 |
by (auto simp add: comb_seq_def not_less less_Suc_eq le_imp_diff_is_add intro!: ext split: nat.split) |
|
485 |
||
486 |
lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \<omega> = nat_case (\<omega> 0)" |
|
487 |
by (intro ext) (simp add: comb_seq_Suc comb_seq_0) |
|
488 |
||
489 |
lemma comb_seq_less: "i < n \<Longrightarrow> comb_seq n \<omega> \<omega>' i = \<omega> i" |
|
490 |
by (auto split: split_comb_seq) |
|
491 |
||
492 |
lemma comb_seq_add: "comb_seq n \<omega> \<omega>' (i + n) = \<omega>' i" |
|
493 |
by (auto split: nat.split split_comb_seq) |
|
494 |
||
495 |
lemma nat_case_comb_seq: "nat_case s' (comb_seq n \<omega> \<omega>') (i + n) = nat_case (nat_case s' \<omega> n) \<omega>' i" |
|
496 |
by (auto split: nat.split split_comb_seq) |
|
497 |
||
498 |
lemma nat_case_comb_seq': |
|
499 |
"nat_case s (comb_seq i \<omega> \<omega>') = comb_seq (Suc i) (nat_case s \<omega>) \<omega>'" |
|
500 |
by (auto split: split_comb_seq nat.split) |
|
501 |
||
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
502 |
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
|
503 |
begin |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
504 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
505 |
abbreviation "S \<equiv> \<Pi>\<^sub>M i\<in>UNIV::nat set. M" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
506 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
507 |
lemma infprod_in_sets[intro]: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
508 |
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
|
509 |
shows "Pi UNIV E \<in> sets S" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
510 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
511 |
have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^sub>E j\<in>{..i}. E j))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
512 |
using E E[THEN sets.sets_into_space] |
47694 | 513 |
by (auto simp: prod_emb_def Pi_iff extensional_def) blast |
514 |
with E show ?thesis by auto |
|
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
515 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
516 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
517 |
lemma measure_PiM_countable: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
518 |
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
|
519 |
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
|
520 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
521 |
let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^sub>E {.. n} E)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
522 |
have "\<And>n. (\<Prod>i\<le>n. measure M (E i)) = measure S (?E n)" |
47694 | 523 |
using E by (simp add: measure_PiM_emb) |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
524 |
moreover have "Pi UNIV E = (\<Inter>n. ?E n)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
525 |
using E E[THEN sets.sets_into_space] |
47694 | 526 |
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
|
527 |
moreover have "range ?E \<subseteq> sets S" |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
528 |
using E by auto |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
529 |
moreover have "decseq ?E" |
47694 | 530 |
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
|
531 |
ultimately show ?thesis |
47694 | 532 |
by (simp add: finite_Lim_measure_decseq) |
42257
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
533 |
qed |
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents:
42166
diff
changeset
|
534 |
|
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
535 |
lemma nat_eq_diff_eq: |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
536 |
fixes a b c :: nat |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
537 |
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
|
538 |
by auto |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
539 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
540 |
lemma PiM_comb_seq: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
541 |
"distr (S \<Otimes>\<^sub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _") |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
542 |
proof (rule PiM_eq) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
543 |
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
544 |
let "distr _ _ ?f" = "?D" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
545 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
546 |
fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
547 |
let ?X = "prod_emb ?I ?M J (\<Pi>\<^sub>E j\<in>J. E j)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
548 |
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
549 |
using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
550 |
with J have "?f -` ?X \<inter> space (S \<Otimes>\<^sub>M S) = |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
551 |
(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
|
552 |
(prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \<times> ?F") |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
553 |
by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib PiE_iff |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
554 |
split: split_comb_seq split_comb_seq_asm) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
555 |
then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^sub>M S) (?E \<times> ?F)" |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
556 |
by (subst emeasure_distr[OF measurable_comb_seq]) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
557 |
(auto intro!: sets_PiM_I simp: split_beta' J) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
558 |
also have "\<dots> = emeasure S ?E * emeasure S ?F" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
559 |
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
|
560 |
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
|
561 |
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
|
562 |
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
|
563 |
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
|
564 |
(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
|
565 |
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
|
566 |
using J by (intro emeasure_PiM_emb) simp_all |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
567 |
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
|
568 |
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
|
569 |
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
|
570 |
qed simp_all |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
571 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
572 |
lemma PiM_iter: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
573 |
"distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). nat_case s \<omega>) = S" (is "?D = _") |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
574 |
proof (rule PiM_eq) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
575 |
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
576 |
let "distr _ _ ?f" = "?D" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
577 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
578 |
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
|
579 |
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
|
580 |
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
581 |
using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
582 |
with J have "?f -` ?X \<inter> space (M \<Otimes>\<^sub>M S) = (if 0 \<in> J then E 0 else space M) \<times> |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
583 |
(prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \<times> ?F") |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50099
diff
changeset
|
584 |
by (auto simp: space_pair_measure space_PiM PiE_iff prod_emb_def all_conj_distrib |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
585 |
split: nat.split nat.split_asm) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51351
diff
changeset
|
586 |
then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^sub>M S) (?E \<times> ?F)" |
50099 | 587 |
by (subst emeasure_distr) |
50000
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
588 |
(auto intro!: sets_PiM_I simp: split_beta' J) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
589 |
also have "\<dots> = emeasure M ?E * emeasure S ?F" |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
590 |
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
|
591 |
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
|
592 |
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
|
593 |
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
|
594 |
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
|
595 |
(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
|
596 |
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
|
597 |
by (auto simp: M.emeasure_space_1 setprod.remove J) |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
598 |
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
|
599 |
qed simp_all |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
600 |
|
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
601 |
end |
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
hoelzl
parents:
49804
diff
changeset
|
602 |
|
42147 | 603 |
end |