| author | paulson | 
| Mon, 30 Apr 2018 22:13:21 +0100 | |
| changeset 68059 | 6f7829c14f5a | 
| parent 67399 | eab6ce8368fa | 
| child 69260 | 0a9688695a1b | 
| permissions | -rw-r--r-- | 
| 50091 | 1  | 
(* Title: HOL/Probability/Projective_Limit.thy  | 
| 50088 | 2  | 
Author: Fabian Immler, TU München  | 
3  | 
*)  | 
|
4  | 
||
| 61808 | 5  | 
section \<open>Projective Limit\<close>  | 
| 50088 | 6  | 
|
7  | 
theory Projective_Limit  | 
|
| 
63626
 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 
hoelzl 
parents: 
63040 
diff
changeset
 | 
8  | 
imports  | 
| 
 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 
hoelzl 
parents: 
63040 
diff
changeset
 | 
9  | 
Fin_Map  | 
| 
 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 
hoelzl 
parents: 
63040 
diff
changeset
 | 
10  | 
Infinite_Product_Measure  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
66447 
diff
changeset
 | 
11  | 
"HOL-Library.Diagonal_Subsequence"  | 
| 50088 | 12  | 
begin  | 
13  | 
||
| 61808 | 14  | 
subsection \<open>Sequences of Finite Maps in Compact Sets\<close>  | 
| 50088 | 15  | 
|
16  | 
locale finmap_seqs_into_compact =  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
17  | 
fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a)" and M  | 
| 50088 | 18  | 
assumes compact: "\<And>n. compact (K n)"  | 
19  | 
  assumes f_in_K: "\<And>n. K n \<noteq> {}"
 | 
|
20  | 
assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)"  | 
|
21  | 
assumes proj_in_K:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
22  | 
"\<And>t n m. m \<ge> n \<Longrightarrow> t \<in> domain (f n) \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"  | 
| 50088 | 23  | 
begin  | 
24  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
25  | 
lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n)"  | 
| 50088 | 26  | 
using proj_in_K f_in_K  | 
27  | 
proof cases  | 
|
28  | 
obtain k where "k \<in> K (Suc 0)" using f_in_K by auto  | 
|
29  | 
assume "\<forall>n. t \<notin> domain (f n)"  | 
|
30  | 
thus ?thesis  | 
|
| 61808 | 31  | 
by (auto intro!: exI[where x=1] image_eqI[OF _ \<open>k \<in> K (Suc 0)\<close>]  | 
32  | 
simp: domain_K[OF \<open>k \<in> K (Suc 0)\<close>])  | 
|
| 50088 | 33  | 
qed blast  | 
34  | 
||
35  | 
lemma proj_in_KE:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
36  | 
obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"  | 
| 50088 | 37  | 
using proj_in_K' by blast  | 
38  | 
||
39  | 
lemma compact_projset:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
40  | 
shows "compact ((\<lambda>k. (k)\<^sub>F i) ` K n)"  | 
| 50088 | 41  | 
using continuous_proj compact by (rule compact_continuous_image)  | 
42  | 
||
43  | 
end  | 
|
44  | 
||
45  | 
lemma compactE':  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50252 
diff
changeset
 | 
46  | 
fixes S :: "'a :: metric_space set"  | 
| 50088 | 47  | 
assumes "compact S" "\<forall>n\<ge>m. f n \<in> S"  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
48  | 
obtains l r where "l \<in> S" "strict_mono (r::nat\<Rightarrow>nat)" "((f \<circ> r) \<longlongrightarrow> l) sequentially"  | 
| 50088 | 49  | 
proof atomize_elim  | 
| 67399 | 50  | 
have "strict_mono ((+) m)" by (simp add: strict_mono_def)  | 
| 50088 | 51  | 
have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> S" using assms by auto  | 
| 61808 | 52  | 
from seq_compactE[OF \<open>compact S\<close>[unfolded compact_eq_seq_compact_metric] this] guess l r .  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
53  | 
hence "l \<in> S" "strict_mono ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) \<longlonglongrightarrow> l"  | 
| 67399 | 54  | 
using strict_mono_o[OF \<open>strict_mono ((+) m)\<close> \<open>strict_mono r\<close>] by (auto simp: o_def)  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
55  | 
thus "\<exists>l r. l \<in> S \<and> strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" by blast  | 
| 50088 | 56  | 
qed  | 
57  | 
||
| 61969 | 58  | 
sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^sub>F n) \<longlonglongrightarrow> l)"  | 
| 50088 | 59  | 
proof  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
60  | 
fix n and s :: "nat \<Rightarrow> nat"  | 
| 
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
61  | 
assume "strict_mono s"  | 
| 50088 | 62  | 
from proj_in_KE[of n] guess n0 . note n0 = this  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
63  | 
have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0"  | 
| 50088 | 64  | 
proof safe  | 
65  | 
fix i assume "n0 \<le> i"  | 
|
66  | 
also have "\<dots> \<le> s i" by (rule seq_suble) fact  | 
|
67  | 
finally have "n0 \<le> s i" .  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
68  | 
with n0 show "((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0 "  | 
| 50088 | 69  | 
by auto  | 
70  | 
qed  | 
|
71  | 
from compactE'[OF compact_projset this] guess ls rs .  | 
|
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
72  | 
thus "\<exists>r'. strict_mono r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^sub>F n) \<longlonglongrightarrow> l)" by (auto simp: o_def)  | 
| 50088 | 73  | 
qed  | 
74  | 
||
| 61969 | 75  | 
lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^sub>F n) \<longlonglongrightarrow> l"  | 
| 50088 | 76  | 
proof -  | 
| 67399 | 77  | 
obtain l where "(\<lambda>i. ((f o (diagseq o (+) (Suc n))) i)\<^sub>F n) \<longlonglongrightarrow> l"  | 
| 
52681
 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 
immler 
parents: 
51351 
diff
changeset
 | 
78  | 
proof (atomize_elim, rule diagseq_holds)  | 
| 
 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 
immler 
parents: 
51351 
diff
changeset
 | 
79  | 
fix r s n  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
80  | 
assume "strict_mono (r :: nat \<Rightarrow> nat)"  | 
| 61969 | 81  | 
assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) \<longlonglongrightarrow> l"  | 
82  | 
then obtain l where "((\<lambda>i. (f i)\<^sub>F n) o s) \<longlonglongrightarrow> l"  | 
|
| 
52681
 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 
immler 
parents: 
51351 
diff
changeset
 | 
83  | 
by (auto simp: o_def)  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
84  | 
hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) \<longlonglongrightarrow> l" using \<open>strict_mono r\<close>  | 
| 
52681
 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 
immler 
parents: 
51351 
diff
changeset
 | 
85  | 
by (rule LIMSEQ_subseq_LIMSEQ)  | 
| 61969 | 86  | 
thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^sub>F n) \<longlonglongrightarrow> l" by (auto simp add: o_def)  | 
| 50088 | 87  | 
qed  | 
| 61969 | 88  | 
hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^sub>F n) \<longlonglongrightarrow> l" by (simp add: ac_simps)  | 
89  | 
hence "(\<lambda>i. (f (diagseq i))\<^sub>F n) \<longlonglongrightarrow> l" by (rule LIMSEQ_offset)  | 
|
| 
52681
 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 
immler 
parents: 
51351 
diff
changeset
 | 
90  | 
thus ?thesis ..  | 
| 50088 | 91  | 
qed  | 
92  | 
||
| 61808 | 93  | 
subsection \<open>Daniell-Kolmogorov Theorem\<close>  | 
| 50088 | 94  | 
|
| 61808 | 95  | 
text \<open>Existence of Projective Limit\<close>  | 
| 50088 | 96  | 
|
97  | 
locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure"  | 
|
98  | 
for I::"'i set" and P  | 
|
99  | 
begin  | 
|
100  | 
||
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
101  | 
lemma emeasure_lim_emb:  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
102  | 
assumes X: "J \<subseteq> I" "finite J" "X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
103  | 
shows "lim (emb I J X) = P J X"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
104  | 
proof (rule emeasure_lim)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
105  | 
  write mu_G ("\<mu>G")
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
106  | 
interpret generator: algebra "space (PiM I (\<lambda>i. borel))" generator  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
107  | 
by (rule algebra_generator)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
108  | 
|
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
109  | 
  fix J and B :: "nat \<Rightarrow> ('i \<Rightarrow> 'a) set"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
110  | 
assume J: "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)" "incseq J"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
111  | 
and B: "decseq (\<lambda>n. emb I (J n) (B n))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
112  | 
and "0 < (INF i. P (J i) (B i))" (is "0 < ?a")  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
113  | 
moreover have "?a \<le> 1"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
114  | 
using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
115  | 
ultimately obtain r where r: "?a = ennreal r" "0 < r" "r \<le> 1"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
116  | 
by (cases "?a") (auto simp: top_unique)  | 
| 63040 | 117  | 
define Z where "Z n = emb I (J n) (B n)" for n  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
118  | 
have Z_mono: "n \<le> m \<Longrightarrow> Z m \<subseteq> Z n" for n m  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
119  | 
unfolding Z_def using B[THEN antimonoD, of n m] .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
120  | 
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
121  | 
using \<open>incseq J\<close> by (force simp: incseq_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
122  | 
note [simp] = \<open>\<And>n. finite (J n)\<close>  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
123  | 
interpret prob_space "P (J i)" for i using J prob_space_P by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
124  | 
|
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
125  | 
have P_eq[simp]:  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
126  | 
"sets (P (J i)) = sets (\<Pi>\<^sub>M i\<in>J i. borel)" "space (P (J i)) = space (\<Pi>\<^sub>M i\<in>J i. borel)" for i  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
127  | 
using J by (auto simp: sets_P space_P)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
128  | 
|
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
129  | 
have "Z i \<in> generator" for i  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
130  | 
unfolding Z_def by (auto intro!: generator.intros J)  | 
| 50088 | 131  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
132  | 
have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite)  | 
| 63040 | 133  | 
define Utn where "Utn = to_nat_on (\<Union>n. J n)"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
134  | 
interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>n. J n)" for n  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
135  | 
by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J])  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
136  | 
have inj_on_Utn: "inj_on Utn (\<Union>n. J n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
137  | 
unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
138  | 
hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto  | 
| 63040 | 139  | 
define P' where "P' n = mapmeasure n (P (J n)) (\<lambda>_. borel)" for n  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
140  | 
interpret P': prob_space "P' n" for n  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
141  | 
unfolding P'_def mapmeasure_def using J  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
142  | 
by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
143  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
144  | 
  let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
145  | 
  { fix n
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
146  | 
have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
147  | 
using J by (auto simp: P'_def mapmeasure_PiM space_P sets_P)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
148  | 
also  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
149  | 
have "\<dots> = ?SUP n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
150  | 
proof (rule inner_regular)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
151  | 
show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
152  | 
next  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
153  | 
show "fm n ` B n \<in> sets borel"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
154  | 
unfolding borel_eq_PiF_borel by (auto simp: P'_def fm_image_measurable_finite sets_P J(3))  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
155  | 
qed simp  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
156  | 
finally have *: "emeasure (P (J n)) (B n) = ?SUP n" .  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
157  | 
have "?SUP n \<noteq> \<infinity>"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
158  | 
unfolding *[symmetric] by simp  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
159  | 
note * this  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
160  | 
} note R = this  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
161  | 
have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a \<and> compact K \<and> K \<subseteq> fm n ` B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
162  | 
proof  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
163  | 
fix n show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ennreal (2 powr - real n) * ?a \<and>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
164  | 
compact K \<and> K \<subseteq> fm n ` B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
165  | 
unfolding R[of n]  | 
| 50088 | 166  | 
proof (rule ccontr)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
167  | 
assume H: "\<nexists>K'. ?SUP n - emeasure (P' n) K' \<le> ennreal (2 powr - real n) * ?a \<and>  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
168  | 
compact K' \<and> K' \<subseteq> fm n ` B n"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
169  | 
have "?SUP n + 0 < ?SUP n + 2 powr (-n) * ?a"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
170  | 
using R[of n] unfolding ennreal_add_left_cancel_less ennreal_zero_less_mult_iff  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
171  | 
by (auto intro: \<open>0 < ?a\<close>)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
172  | 
      also have "\<dots> = (SUP K:{K. K \<subseteq> fm n ` B n \<and> compact K}. emeasure (P' n) K + 2 powr (-n) * ?a)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
173  | 
by (rule ennreal_SUP_add_left[symmetric]) auto  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
174  | 
also have "\<dots> \<le> ?SUP n"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
175  | 
proof (intro SUP_least)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
176  | 
        fix K assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
177  | 
with H have "2 powr (-n) * ?a < ?SUP n - emeasure (P' n) K"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
178  | 
by auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
179  | 
then show "emeasure (P' n) K + (2 powr (-n)) * ?a \<le> ?SUP n"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
180  | 
by (subst (asm) less_diff_eq_ennreal) (auto simp: less_top[symmetric] R(1)[symmetric] ac_simps)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
181  | 
qed  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
182  | 
finally show False by simp  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
183  | 
qed  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
184  | 
qed  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
185  | 
then obtain K' where K':  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
186  | 
"\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ennreal (2 powr - real n) * ?a"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
187  | 
"\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
188  | 
unfolding choice_iff by blast  | 
| 63040 | 189  | 
define K where "K n = fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" for n  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
190  | 
have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
191  | 
unfolding K_def  | 
| 61808 | 192  | 
using compact_imp_closed[OF \<open>compact (K' _)\<close>]  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
193  | 
by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
194  | 
(auto simp: borel_eq_PiF_borel[symmetric])  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
195  | 
have K_B: "\<And>n. K n \<subseteq> B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
196  | 
proof  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
197  | 
fix x n assume "x \<in> K n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
198  | 
then have fm_in: "fm n x \<in> fm n ` B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
199  | 
using K' by (force simp: K_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
200  | 
show "x \<in> B n"  | 
| 61808 | 201  | 
using \<open>x \<in> K n\<close> K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm]  | 
| 
61520
 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
202  | 
by (metis (no_types) Int_iff K_def fm_in space_borel)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
203  | 
qed  | 
| 63040 | 204  | 
define Z' where "Z' n = emb I (J n) (K n)" for n  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
205  | 
have Z': "\<And>n. Z' n \<subseteq> Z n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
206  | 
unfolding Z'_def Z_def  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
207  | 
proof (rule prod_emb_mono, safe)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
208  | 
fix n x assume "x \<in> K n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
209  | 
hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
210  | 
by (simp_all add: K_def space_P)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
211  | 
note this(1)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
212  | 
also have "K' n \<subseteq> fm n ` B n" by (simp add: K')  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
213  | 
finally have "fm n x \<in> fm n ` B n" .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
214  | 
thus "x \<in> B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
215  | 
proof safe  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
216  | 
fix y assume y: "y \<in> B n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
217  | 
hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
218  | 
by (auto simp add: space_P sets_P)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
219  | 
assume "fm n x = fm n y"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
220  | 
note inj_onD[OF inj_on_fm[OF space_borel],  | 
| 61808 | 221  | 
OF \<open>fm n x = fm n y\<close> \<open>x \<in> space _\<close> \<open>y \<in> space _\<close>]  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
222  | 
with y show "x \<in> B n" by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
223  | 
qed  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
224  | 
qed  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
225  | 
have "\<And>n. Z' n \<in> generator" using J K'(2) unfolding Z'_def  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
226  | 
by (auto intro!: generator.intros measurable_sets[OF fm_measurable[of _ "Collect finite"]]  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
227  | 
simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed)  | 
| 63040 | 228  | 
  define Y where "Y n = (\<Inter>i\<in>{1..n}. Z' i)" for n
 | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
229  | 
hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
230  | 
hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
231  | 
have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
232  | 
hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto  | 
| 50088 | 233  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
234  | 
  have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
235  | 
proof -  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
236  | 
fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
237  | 
    have "Y n = (\<Inter>i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
238  | 
by (auto simp: Y_def Z'_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
239  | 
    also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
 | 
| 61808 | 240  | 
using \<open>n \<ge> 1\<close>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
241  | 
by (subst prod_emb_INT) auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
242  | 
finally  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
243  | 
have Y_emb:  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
244  | 
      "Y n = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
 | 
| 61808 | 245  | 
hence "Y n \<in> generator" using J J_mono K_sets \<open>n \<ge> 1\<close>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
246  | 
by (auto simp del: prod_emb_INT intro!: generator.intros)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
247  | 
have *: "\<mu>G (Z n) = P (J n) (B n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
248  | 
unfolding Z_def using J by (intro mu_G_spec) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
249  | 
then have "\<mu>G (Z n) \<noteq> \<infinity>" by auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
250  | 
note *  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
251  | 
    moreover have *: "\<mu>G (Y n) = P (J n) (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
 | 
| 61808 | 252  | 
unfolding Y_emb using J J_mono K_sets \<open>n \<ge> 1\<close> by (subst mu_G_spec) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
253  | 
then have "\<mu>G (Y n) \<noteq> \<infinity>" by auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
254  | 
note *  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
255  | 
moreover have "\<mu>G (Z n - Y n) =  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
256  | 
        P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
 | 
| 61808 | 257  | 
unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets \<open>n \<ge> 1\<close>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
258  | 
by (subst mu_G_spec) (auto intro!: sets.Diff)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
259  | 
ultimately  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
260  | 
have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"  | 
| 61808 | 261  | 
using J J_mono K_sets \<open>n \<ge> 1\<close>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
262  | 
by (simp only: emeasure_eq_measure Z_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
263  | 
(auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
264  | 
intro!: arg_cong[where f=ennreal]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
265  | 
simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
266  | 
ennreal_minus measure_nonneg)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
267  | 
    also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
 | 
| 61808 | 268  | 
using \<open>n \<ge> 1\<close> unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
269  | 
    have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
 | 
| 61808 | 270  | 
using \<open>Z' _ \<in> generator\<close> \<open>Z _ \<in> generator\<close> \<open>Y _ \<in> generator\<close> by auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
271  | 
    hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union>i\<in>{1..n}. (Z i - Z' i))"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
272  | 
using subs generator.additive_increasing[OF positive_mu_G additive_mu_G]  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
273  | 
unfolding increasing_def by auto  | 
| 61808 | 274  | 
    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using \<open>Z _ \<in> generator\<close> \<open>Z' _ \<in> generator\<close>
 | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
275  | 
by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
276  | 
    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
 | 
| 64267 | 277  | 
proof (rule sum_mono)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
278  | 
      fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
279  | 
have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
280  | 
unfolding Z'_def Z_def by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
281  | 
also have "\<dots> = P (J i) (B i - K i)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
282  | 
using J K_sets by (subst mu_G_spec) auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
283  | 
also have "\<dots> = P (J i) (B i) - P (J i) (K i)"  | 
| 61808 | 284  | 
using K_sets J \<open>K _ \<subseteq> B _\<close> by (simp add: emeasure_Diff)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
285  | 
also have "\<dots> = P (J i) (B i) - P' i (K' i)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
286  | 
unfolding K_def P'_def  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
287  | 
by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric]  | 
| 61808 | 288  | 
compact_imp_closed[OF \<open>compact (K' _)\<close>] space_PiM PiE_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
289  | 
also have "\<dots> \<le> ennreal (2 powr - real i) * ?a" using K'(1)[of i] .  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
290  | 
finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
291  | 
qed  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
292  | 
    also have "\<dots> = ennreal ((\<Sum> i\<in>{1..n}. (2 powr -enn2real i)) * enn2real ?a)"
 | 
| 64267 | 293  | 
using r by (simp add: sum_distrib_right ennreal_mult[symmetric])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
294  | 
also have "\<dots> < ennreal (1 * enn2real ?a)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
295  | 
proof (intro ennreal_lessI mult_strict_right_mono)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
296  | 
have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"  | 
| 64267 | 297  | 
by (rule sum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
298  | 
      also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
 | 
| 67399 | 299  | 
      also have "sum ((^) (1 / 2::real)) ({..<Suc n} - {0}) =
 | 
300  | 
        sum ((^) (1 / 2)) ({..<Suc n}) - 1" by (auto simp: sum_diff1)
 | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
301  | 
also have "\<dots> < 1" by (subst geometric_sum) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
302  | 
finally show "(\<Sum>i = 1..n. 2 powr - enn2real i) < 1" by simp  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
303  | 
qed (auto simp: r enn2real_positive_iff)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
304  | 
also have "\<dots> = ?a" by (auto simp: r)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
305  | 
also have "\<dots> \<le> \<mu>G (Z n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
306  | 
using J by (auto intro: INF_lower simp: Z_def mu_G_spec)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
307  | 
finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
308  | 
hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
309  | 
using \<open>\<mu>G (Y n) \<noteq> \<infinity>\<close> by (auto simp: zero_less_iff_neq_zero)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
310  | 
then have "\<mu>G (Y n) > 0"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
311  | 
by simp  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
312  | 
    thus "Y n \<noteq> {}" using positive_mu_G by (auto simp add: positive_def)
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
313  | 
qed  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
314  | 
  hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
315  | 
then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
316  | 
  {
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
317  | 
fix t and n m::nat  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
318  | 
assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp  | 
| 61808 | 319  | 
from Y_mono[OF \<open>m \<ge> n\<close>] y[OF \<open>1 \<le> m\<close>] have "y m \<in> Y n" by auto  | 
320  | 
also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF \<open>1 \<le> n\<close>] .  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
321  | 
finally  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
322  | 
have "fm n (restrict (y m) (J n)) \<in> K' n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
323  | 
unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
324  | 
moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
325  | 
using J by (simp add: fm_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
326  | 
ultimately have "fm n (y m) \<in> K' n" by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
327  | 
} note fm_in_K' = this  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
328  | 
interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
329  | 
proof  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
330  | 
fix n show "compact (K' n)" by fact  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
331  | 
next  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
332  | 
fix n  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
333  | 
from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
334  | 
also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
335  | 
finally  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
336  | 
have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
337  | 
unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
338  | 
    thus "K' (Suc n) \<noteq> {}" by auto
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
339  | 
fix k  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
340  | 
assume "k \<in> K' (Suc n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
341  | 
with K'[of "Suc n"] sets.sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
342  | 
then obtain b where "k = fm (Suc n) b" by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
343  | 
thus "domain k = domain (fm (Suc n) (y (Suc n)))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
344  | 
by (simp_all add: fm_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
345  | 
next  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
346  | 
fix t and n m::nat  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
347  | 
assume "n \<le> m" hence "Suc n \<le> Suc m" by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
348  | 
assume "t \<in> domain (fm (Suc n) (y (Suc n)))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
349  | 
then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto  | 
| 61808 | 350  | 
hence "j \<in> J (Suc m)" using J_mono[OF \<open>Suc n \<le> Suc m\<close>] by auto  | 
351  | 
have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using \<open>n \<le> m\<close>  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
352  | 
by (intro fm_in_K') simp_all  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
353  | 
show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
354  | 
apply (rule image_eqI[OF _ img])  | 
| 61808 | 355  | 
using \<open>j \<in> J (Suc n)\<close> \<open>j \<in> J (Suc m)\<close>  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
356  | 
unfolding j by (subst proj_fm, auto)+  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
357  | 
qed  | 
| 61969 | 358  | 
have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
359  | 
using diagonal_tendsto ..  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
360  | 
then obtain z where z:  | 
| 61969 | 361  | 
"\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z t"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
362  | 
unfolding choice_iff by blast  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
363  | 
  {
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
364  | 
fix n :: nat assume "n \<ge> 1"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
365  | 
have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
366  | 
by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
367  | 
moreover  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
368  | 
    {
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
369  | 
fix t  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
370  | 
assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
371  | 
hence "t \<in> Utn ` J n" by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
372  | 
then obtain j where j: "t = Utn j" "j \<in> J n" by auto  | 
| 61969 | 373  | 
have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z t"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
374  | 
apply (subst (2) tendsto_iff, subst eventually_sequentially)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
375  | 
proof safe  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
376  | 
fix e :: real assume "0 < e"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
377  | 
        { fix i and x :: "'i \<Rightarrow> 'a" assume i: "i \<ge> n"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
378  | 
assume "t \<in> domain (fm n x)"  | 
| 61808 | 379  | 
hence "t \<in> domain (fm i x)" using J_mono[OF \<open>i \<ge> n\<close>] by auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
380  | 
with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
381  | 
using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn])  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
382  | 
} note index_shift = this  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
383  | 
have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
384  | 
apply (rule le_SucI)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
385  | 
apply (rule order_trans) apply simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
386  | 
apply (rule seq_suble[OF subseq_diagseq])  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
387  | 
done  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
388  | 
from z  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
389  | 
have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e"  | 
| 61808 | 390  | 
unfolding tendsto_iff eventually_sequentially using \<open>0 < e\<close> by auto  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
391  | 
then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
392  | 
dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
393  | 
show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e "  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
394  | 
proof (rule exI[where x="max N n"], safe)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
395  | 
fix na assume "max N n \<le> na"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
396  | 
hence "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) =  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
397  | 
dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
398  | 
by (subst index_shift[OF I]) auto  | 
| 61808 | 399  | 
also have "\<dots> < e" using \<open>max N n \<le> na\<close> by (intro N) simp  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
400  | 
finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" .  | 
| 50088 | 401  | 
qed  | 
402  | 
qed  | 
|
| 61969 | 403  | 
hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> (finmap_of (Utn ` J n) z)\<^sub>F t"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
404  | 
by (simp add: tendsto_intros)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
405  | 
} ultimately  | 
| 61969 | 406  | 
have "(\<lambda>i. fm n (y (Suc (diagseq i)))) \<longlonglongrightarrow> finmap_of (Utn ` J n) z"  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
407  | 
by (rule tendsto_finmap)  | 
| 61969 | 408  | 
hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) \<longlonglongrightarrow> finmap_of (Utn ` J n) z"  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
64267 
diff
changeset
 | 
409  | 
by (rule LIMSEQ_subseq_LIMSEQ) (simp add: strict_mono_def)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
410  | 
moreover  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
411  | 
have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"  | 
| 61808 | 412  | 
apply (auto simp add: o_def intro!: fm_in_K' \<open>1 \<le> n\<close> le_SucI)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
413  | 
apply (rule le_trans)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
414  | 
apply (rule le_add2)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
415  | 
using seq_suble[OF subseq_diagseq]  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
416  | 
apply auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
417  | 
done  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
418  | 
moreover  | 
| 61808 | 419  | 
from \<open>compact (K' n)\<close> have "closed (K' n)" by (rule compact_imp_closed)  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
420  | 
ultimately  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
421  | 
have "finmap_of (Utn ` J n) z \<in> K' n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
422  | 
unfolding closed_sequential_limits by blast  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
423  | 
also have "finmap_of (Utn ` J n) z = fm n (\<lambda>i. z (Utn i))"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
424  | 
unfolding finmap_eq_iff  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
425  | 
proof clarsimp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
426  | 
fix i assume i: "i \<in> J n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
427  | 
hence "from_nat_into (\<Union>n. J n) (Utn i) = i"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
428  | 
unfolding Utn_def  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
429  | 
by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
430  | 
with i show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
431  | 
by (simp add: finmap_eq_iff fm_def compose_def)  | 
| 50088 | 432  | 
qed  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
433  | 
finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
434  | 
moreover  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
435  | 
let ?J = "\<Union>n. J n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
436  | 
have "(?J \<inter> J n) = J n" by auto  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
437  | 
ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
438  | 
unfolding K_def by (auto simp: space_P space_PiM)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
439  | 
hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
440  | 
using J by (auto simp: prod_emb_def PiE_def extensional_def)  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
441  | 
also have "\<dots> \<subseteq> Z n" using Z' by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
442  | 
finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
443  | 
} note in_Z = this  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
444  | 
  hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
445  | 
  thus "(\<Inter>i. Z i) \<noteq> {}"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
446  | 
using INT_decseq_offset[OF antimonoI[OF Z_mono]] by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
447  | 
qed fact+  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
448  | 
|
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
449  | 
lemma measure_lim_emb:  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
450  | 
"J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel) \<Longrightarrow> measure lim (emb I J X) = measure (P J) X"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
451  | 
unfolding measure_def by (subst emeasure_lim_emb) auto  | 
| 50088 | 452  | 
|
453  | 
end  | 
|
454  | 
||
| 50090 | 455  | 
hide_const (open) PiF  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52681 
diff
changeset
 | 
456  | 
hide_const (open) Pi\<^sub>F  | 
| 50090 | 457  | 
hide_const (open) Pi'  | 
458  | 
hide_const (open) finmap_of  | 
|
459  | 
hide_const (open) proj  | 
|
460  | 
hide_const (open) domain  | 
|
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
461  | 
hide_const (open) basis_finmap  | 
| 50090 | 462  | 
|
| 61605 | 463  | 
sublocale polish_projective \<subseteq> P: prob_space lim  | 
| 50088 | 464  | 
proof  | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
465  | 
  have *: "emb I {} {\<lambda>x. undefined} = space (\<Pi>\<^sub>M i\<in>I. borel)"
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
466  | 
by (auto simp: prod_emb_def space_PiM)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62397 
diff
changeset
 | 
467  | 
  interpret prob_space "P {}"
 | 
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
468  | 
using prob_space_P by simp  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
469  | 
show "emeasure lim (space lim) = 1"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
470  | 
    using emeasure_lim_emb[of "{}" "{\<lambda>x. undefined}"] emeasure_space_1
 | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
471  | 
by (simp add: * PiM_empty space_P)  | 
| 50088 | 472  | 
qed  | 
473  | 
||
474  | 
locale polish_product_prob_space =  | 
|
475  | 
  product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
 | 
|
476  | 
||
477  | 
sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
 | 
|
| 
63885
 
a6cd18af8bf9
new type for finite maps; use it in HOL-Probability
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
63626 
diff
changeset
 | 
478  | 
..  | 
| 50088 | 479  | 
|
| 
61359
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
480  | 
lemma (in polish_product_prob_space) limP_eq_PiM: "lim = PiM I (\<lambda>_. borel)"  | 
| 
 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 
hoelzl 
parents: 
60809 
diff
changeset
 | 
481  | 
by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_lim_emb)  | 
| 50088 | 482  | 
|
483  | 
end  |