| author | wenzelm | 
| Wed, 23 Aug 2023 11:00:30 +0200 | |
| changeset 78568 | a97d2b6b5c3e | 
| parent 74362 | 0135a0c77b64 | 
| child 80914 | d97fdabd9e2b | 
| 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: 
63040diff
changeset | 8 | imports | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63040diff
changeset | 9 | Fin_Map | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63040diff
changeset | 10 | Infinite_Product_Measure | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66447diff
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: 
52681diff
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: 
52681diff
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: 
52681diff
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: 
52681diff
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: 
52681diff
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: 
50252diff
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: 
64267diff
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 | 
| 74362 | 52 | from seq_compactE[OF \<open>compact S\<close>[unfolded compact_eq_seq_compact_metric] this] | 
| 53 | obtain l r where "l \<in> S" "strict_mono r" "(f \<circ> (+) m \<circ> r) \<longlonglongrightarrow> l" by blast | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 54 | hence "l \<in> S" "strict_mono ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) \<longlonglongrightarrow> l" | 
| 67399 | 55 | 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: 
64267diff
changeset | 56 | thus "\<exists>l r. l \<in> S \<and> strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" by blast | 
| 50088 | 57 | qed | 
| 58 | ||
| 61969 | 59 | sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^sub>F n) \<longlonglongrightarrow> l)" | 
| 50088 | 60 | proof | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 61 | fix n and s :: "nat \<Rightarrow> nat" | 
| 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 62 | assume "strict_mono s" | 
| 74362 | 63 | from proj_in_KE[of n] obtain n0 where n0: "\<And>m. n0 \<le> m \<Longrightarrow> (f m)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0" | 
| 64 | by blast | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 65 | have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0" | 
| 50088 | 66 | proof safe | 
| 67 | fix i assume "n0 \<le> i" | |
| 68 | also have "\<dots> \<le> s i" by (rule seq_suble) fact | |
| 69 | 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: 
52681diff
changeset | 70 | with n0 show "((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0 " | 
| 50088 | 71 | by auto | 
| 72 | qed | |
| 74362 | 73 | then obtain ls rs | 
| 74 | where "ls \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0" "strict_mono rs" "((\<lambda>i. ((f \<circ> s) i)\<^sub>F n) \<circ> rs) \<longlonglongrightarrow> ls" | |
| 75 | by (rule compactE'[OF compact_projset]) | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 76 | 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 | 77 | qed | 
| 78 | ||
| 61969 | 79 | lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^sub>F n) \<longlonglongrightarrow> l" | 
| 50088 | 80 | proof - | 
| 67399 | 81 | 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: 
51351diff
changeset | 82 | proof (atomize_elim, rule diagseq_holds) | 
| 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 83 | fix r s n | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 84 | assume "strict_mono (r :: nat \<Rightarrow> nat)" | 
| 61969 | 85 | assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) \<longlonglongrightarrow> l" | 
| 86 | 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: 
51351diff
changeset | 87 | by (auto simp: o_def) | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 88 | 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: 
51351diff
changeset | 89 | by (rule LIMSEQ_subseq_LIMSEQ) | 
| 61969 | 90 | thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^sub>F n) \<longlonglongrightarrow> l" by (auto simp add: o_def) | 
| 50088 | 91 | qed | 
| 61969 | 92 | hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^sub>F n) \<longlonglongrightarrow> l" by (simp add: ac_simps) | 
| 93 | 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: 
51351diff
changeset | 94 | thus ?thesis .. | 
| 50088 | 95 | qed | 
| 96 | ||
| 61808 | 97 | subsection \<open>Daniell-Kolmogorov Theorem\<close> | 
| 50088 | 98 | |
| 61808 | 99 | text \<open>Existence of Projective Limit\<close> | 
| 50088 | 100 | |
| 101 | locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure" | |
| 102 | for I::"'i set" and P | |
| 103 | begin | |
| 104 | ||
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 105 | lemma emeasure_lim_emb: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 106 | 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: 
60809diff
changeset | 107 | 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: 
60809diff
changeset | 108 | proof (rule emeasure_lim) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 109 |   write mu_G ("\<mu>G")
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 110 | 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: 
60809diff
changeset | 111 | by (rule algebra_generator) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 112 | |
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 113 |   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: 
60809diff
changeset | 114 | 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: 
60809diff
changeset | 115 | 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: 
60809diff
changeset | 116 | 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: 
60809diff
changeset | 117 | moreover have "?a \<le> 1" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 118 | 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: 
62397diff
changeset | 119 | 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: 
62397diff
changeset | 120 | by (cases "?a") (auto simp: top_unique) | 
| 63040 | 121 | 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: 
60809diff
changeset | 122 | 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: 
60809diff
changeset | 123 | 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: 
60809diff
changeset | 124 | 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: 
60809diff
changeset | 125 | 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: 
60809diff
changeset | 126 | 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: 
60809diff
changeset | 127 | 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: 
60809diff
changeset | 128 | |
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 129 | have P_eq[simp]: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 130 | "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: 
60809diff
changeset | 131 | 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: 
60809diff
changeset | 132 | |
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 133 | 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: 
60809diff
changeset | 134 | unfolding Z_def by (auto intro!: generator.intros J) | 
| 50088 | 135 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 136 | have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite) | 
| 63040 | 137 | 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: 
60809diff
changeset | 138 | 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: 
60809diff
changeset | 139 | 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: 
60809diff
changeset | 140 | 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: 
60809diff
changeset | 141 | 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: 
60809diff
changeset | 142 | hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto | 
| 63040 | 143 | 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: 
60809diff
changeset | 144 | 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: 
60809diff
changeset | 145 | 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: 
60809diff
changeset | 146 | 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: 
62397diff
changeset | 147 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 148 |   let ?SUP = "\<lambda>n. SUP K \<in> {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
 | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 149 |   { fix n
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 150 | 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: 
60809diff
changeset | 151 | 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: 
60809diff
changeset | 152 | also | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 153 | have "\<dots> = ?SUP n" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 154 | proof (rule inner_regular) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 155 | 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: 
60809diff
changeset | 156 | next | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 157 | 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: 
60809diff
changeset | 158 | 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: 
60809diff
changeset | 159 | qed simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 160 | finally have *: "emeasure (P (J n)) (B n) = ?SUP n" . | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 161 | have "?SUP n \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 162 | unfolding *[symmetric] by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 163 | note * this | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 164 | } note R = this | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 165 | 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: 
60809diff
changeset | 166 | proof | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 167 | 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: 
60809diff
changeset | 168 | 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: 
60809diff
changeset | 169 | unfolding R[of n] | 
| 50088 | 170 | proof (rule ccontr) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 171 | 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: 
62397diff
changeset | 172 | compact K' \<and> K' \<subseteq> fm n ` B n" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 173 | have "?SUP n + 0 < ?SUP n + 2 powr (-n) * ?a" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 174 | 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: 
62397diff
changeset | 175 | by (auto intro: \<open>0 < ?a\<close>) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67399diff
changeset | 176 |       also have "\<dots> = (SUP K\<in>{K. K \<subseteq> fm n ` B n \<and> compact K}. emeasure (P' n) K + 2 powr (-n) * ?a)"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 177 | by (rule ennreal_SUP_add_left[symmetric]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 178 | 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: 
60809diff
changeset | 179 | proof (intro SUP_least) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 180 |         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: 
62397diff
changeset | 181 | 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: 
60809diff
changeset | 182 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 183 | 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: 
62397diff
changeset | 184 | 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: 
60809diff
changeset | 185 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 186 | 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: 
60809diff
changeset | 187 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 188 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 189 | then obtain K' where K': | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 190 | "\<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: 
60809diff
changeset | 191 | "\<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: 
60809diff
changeset | 192 | unfolding choice_iff by blast | 
| 63040 | 193 | 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: 
60809diff
changeset | 194 | 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: 
60809diff
changeset | 195 | unfolding K_def | 
| 61808 | 196 | 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: 
60809diff
changeset | 197 | 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: 
60809diff
changeset | 198 | (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: 
60809diff
changeset | 199 | 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: 
60809diff
changeset | 200 | proof | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 201 | 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: 
60809diff
changeset | 202 | 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: 
60809diff
changeset | 203 | 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: 
60809diff
changeset | 204 | show "x \<in> B n" | 
| 61808 | 205 | 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: 
61359diff
changeset | 206 | 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: 
60809diff
changeset | 207 | qed | 
| 63040 | 208 | 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: 
60809diff
changeset | 209 | 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: 
60809diff
changeset | 210 | unfolding Z'_def Z_def | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 211 | 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: 
60809diff
changeset | 212 | 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: 
60809diff
changeset | 213 | 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: 
60809diff
changeset | 214 | 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: 
60809diff
changeset | 215 | note this(1) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 216 | 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: 
60809diff
changeset | 217 | 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: 
60809diff
changeset | 218 | thus "x \<in> B n" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 219 | proof safe | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 220 | 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: 
60809diff
changeset | 221 | 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: 
60809diff
changeset | 222 | 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: 
60809diff
changeset | 223 | 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: 
60809diff
changeset | 224 | note inj_onD[OF inj_on_fm[OF space_borel], | 
| 61808 | 225 | 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: 
60809diff
changeset | 226 | 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: 
60809diff
changeset | 227 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 228 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 229 | 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: 
60809diff
changeset | 230 | 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: 
60809diff
changeset | 231 | simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed) | 
| 63040 | 232 |   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: 
60809diff
changeset | 233 | 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: 
60809diff
changeset | 234 | 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: 
60809diff
changeset | 235 | 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: 
60809diff
changeset | 236 | hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto | 
| 50088 | 237 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 238 |   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: 
60809diff
changeset | 239 | proof - | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 240 | 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: 
60809diff
changeset | 241 |     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: 
60809diff
changeset | 242 | 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: 
60809diff
changeset | 243 |     also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
 | 
| 61808 | 244 | 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: 
60809diff
changeset | 245 | 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: 
60809diff
changeset | 246 | finally | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 247 | have Y_emb: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 248 |       "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 | 249 | 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: 
60809diff
changeset | 250 | 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: 
60809diff
changeset | 251 | 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: 
60809diff
changeset | 252 | 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: 
62397diff
changeset | 253 | 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: 
60809diff
changeset | 254 | note * | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 255 |     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 | 256 | 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: 
62397diff
changeset | 257 | 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: 
60809diff
changeset | 258 | note * | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 259 | 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: 
60809diff
changeset | 260 |         P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
 | 
| 61808 | 261 | 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: 
60809diff
changeset | 262 | 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: 
60809diff
changeset | 263 | ultimately | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 264 | have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)" | 
| 61808 | 265 | 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: 
60809diff
changeset | 266 | by (simp only: emeasure_eq_measure Z_def) | 
| 69712 | 267 | (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] subsetD[OF K_B] | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 268 | intro!: arg_cong[where f=ennreal] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 269 | 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: 
62397diff
changeset | 270 | ennreal_minus measure_nonneg) | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 271 |     also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
 | 
| 61808 | 272 | 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: 
60809diff
changeset | 273 |     have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
 | 
| 61808 | 274 | 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: 
60809diff
changeset | 275 |     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: 
60809diff
changeset | 276 | 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: 
60809diff
changeset | 277 | unfolding increasing_def by auto | 
| 61808 | 278 |     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: 
60809diff
changeset | 279 | 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: 
60809diff
changeset | 280 |     also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
 | 
| 64267 | 281 | proof (rule sum_mono) | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 282 |       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: 
60809diff
changeset | 283 | 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: 
60809diff
changeset | 284 | 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: 
60809diff
changeset | 285 | 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: 
60809diff
changeset | 286 | 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: 
60809diff
changeset | 287 | also have "\<dots> = P (J i) (B i) - P (J i) (K i)" | 
| 61808 | 288 | 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: 
60809diff
changeset | 289 | 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: 
60809diff
changeset | 290 | unfolding K_def P'_def | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 291 | by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric] | 
| 61808 | 292 | 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: 
62397diff
changeset | 293 | 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: 
60809diff
changeset | 294 | 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: 
60809diff
changeset | 295 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 296 |     also have "\<dots> = ennreal ((\<Sum> i\<in>{1..n}. (2 powr -enn2real i)) * enn2real ?a)"
 | 
| 64267 | 297 | 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: 
62397diff
changeset | 298 | also have "\<dots> < ennreal (1 * enn2real ?a)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 299 | 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: 
60809diff
changeset | 300 | have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)" | 
| 64267 | 301 | 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: 
60809diff
changeset | 302 |       also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
 | 
| 67399 | 303 |       also have "sum ((^) (1 / 2::real)) ({..<Suc n} - {0}) =
 | 
| 304 |         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: 
60809diff
changeset | 305 | also have "\<dots> < 1" by (subst geometric_sum) auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 306 | 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: 
62397diff
changeset | 307 | 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: 
60809diff
changeset | 308 | 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: 
60809diff
changeset | 309 | 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: 
60809diff
changeset | 310 | 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: 
60809diff
changeset | 311 | 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: 
60809diff
changeset | 312 | 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: 
62397diff
changeset | 313 | 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: 
62397diff
changeset | 314 | then have "\<mu>G (Y n) > 0" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 315 | by simp | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 316 |     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: 
60809diff
changeset | 317 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 318 |   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: 
60809diff
changeset | 319 | 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: 
60809diff
changeset | 320 |   {
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 321 | 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: 
60809diff
changeset | 322 | assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp | 
| 61808 | 323 | from Y_mono[OF \<open>m \<ge> n\<close>] y[OF \<open>1 \<le> m\<close>] have "y m \<in> Y n" by auto | 
| 324 | 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: 
60809diff
changeset | 325 | finally | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 326 | 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: 
60809diff
changeset | 327 | 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: 
60809diff
changeset | 328 | 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: 
60809diff
changeset | 329 | 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: 
60809diff
changeset | 330 | 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: 
60809diff
changeset | 331 | } note fm_in_K' = this | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 332 | 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: 
60809diff
changeset | 333 | proof | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 334 | 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: 
60809diff
changeset | 335 | next | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 336 | fix n | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 337 | 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: 
60809diff
changeset | 338 | 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: 
60809diff
changeset | 339 | finally | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 340 | 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: 
60809diff
changeset | 341 | 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: 
60809diff
changeset | 342 |     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: 
60809diff
changeset | 343 | fix k | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 344 | 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: 
60809diff
changeset | 345 | 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: 
60809diff
changeset | 346 | 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: 
60809diff
changeset | 347 | 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: 
60809diff
changeset | 348 | 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: 
60809diff
changeset | 349 | next | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 350 | 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: 
60809diff
changeset | 351 | 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: 
60809diff
changeset | 352 | 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: 
60809diff
changeset | 353 | then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto | 
| 61808 | 354 | hence "j \<in> J (Suc m)" using J_mono[OF \<open>Suc n \<le> Suc m\<close>] by auto | 
| 355 | 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: 
60809diff
changeset | 356 | 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: 
60809diff
changeset | 357 | 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: 
60809diff
changeset | 358 | apply (rule image_eqI[OF _ img]) | 
| 61808 | 359 | 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: 
60809diff
changeset | 360 | 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: 
60809diff
changeset | 361 | qed | 
| 61969 | 362 | 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: 
60809diff
changeset | 363 | using diagonal_tendsto .. | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 364 | then obtain z where z: | 
| 61969 | 365 | "\<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: 
60809diff
changeset | 366 | unfolding choice_iff by blast | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 367 |   {
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 368 | 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: 
60809diff
changeset | 369 | 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: 
60809diff
changeset | 370 | by simp | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 371 | moreover | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 372 |     {
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 373 | fix t | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 374 | 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: 
60809diff
changeset | 375 | 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: 
60809diff
changeset | 376 | then obtain j where j: "t = Utn j" "j \<in> J n" by auto | 
| 61969 | 377 | 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: 
60809diff
changeset | 378 | 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: 
60809diff
changeset | 379 | proof safe | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 380 | 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: 
60809diff
changeset | 381 |         { 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: 
60809diff
changeset | 382 | assume "t \<in> domain (fm n x)" | 
| 61808 | 383 | 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: 
60809diff
changeset | 384 | 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: 
60809diff
changeset | 385 | 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: 
60809diff
changeset | 386 | } note index_shift = this | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 387 | 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: 
60809diff
changeset | 388 | apply (rule le_SucI) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 389 | 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: 
60809diff
changeset | 390 | 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: 
60809diff
changeset | 391 | done | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 392 | from z | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 393 | have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" | 
| 61808 | 394 | 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: 
60809diff
changeset | 395 | 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: 
60809diff
changeset | 396 | 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: 
60809diff
changeset | 397 | 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: 
60809diff
changeset | 398 | 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: 
60809diff
changeset | 399 | 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: 
60809diff
changeset | 400 | 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: 
60809diff
changeset | 401 | 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: 
60809diff
changeset | 402 | by (subst index_shift[OF I]) auto | 
| 61808 | 403 | 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: 
60809diff
changeset | 404 | finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" . | 
| 50088 | 405 | qed | 
| 406 | qed | |
| 61969 | 407 | 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: 
60809diff
changeset | 408 | by (simp add: tendsto_intros) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 409 | } ultimately | 
| 61969 | 410 | 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: 
60809diff
changeset | 411 | by (rule tendsto_finmap) | 
| 61969 | 412 | 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: 
64267diff
changeset | 413 | 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: 
60809diff
changeset | 414 | moreover | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 415 | have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)" | 
| 61808 | 416 | 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: 
60809diff
changeset | 417 | apply (rule le_trans) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 418 | apply (rule le_add2) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 419 | 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: 
60809diff
changeset | 420 | apply auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 421 | done | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 422 | moreover | 
| 61808 | 423 | 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: 
60809diff
changeset | 424 | ultimately | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 425 | 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: 
60809diff
changeset | 426 | 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: 
60809diff
changeset | 427 | 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: 
60809diff
changeset | 428 | unfolding finmap_eq_iff | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 429 | proof clarsimp | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 430 | 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: 
60809diff
changeset | 431 | 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: 
60809diff
changeset | 432 | unfolding Utn_def | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 433 | 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: 
60809diff
changeset | 434 | 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: 
60809diff
changeset | 435 | by (simp add: finmap_eq_iff fm_def compose_def) | 
| 50088 | 436 | qed | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 437 | 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: 
60809diff
changeset | 438 | moreover | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 439 | 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: 
60809diff
changeset | 440 | 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: 
60809diff
changeset | 441 | 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: 
60809diff
changeset | 442 | 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: 
60809diff
changeset | 443 | 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: 
60809diff
changeset | 444 | 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: 
60809diff
changeset | 445 | 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: 
60809diff
changeset | 446 | 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: 
60809diff
changeset | 447 | } note in_Z = this | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 448 |   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: 
60809diff
changeset | 449 |   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: 
60809diff
changeset | 450 | 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: 
60809diff
changeset | 451 | qed fact+ | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 452 | |
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 453 | lemma measure_lim_emb: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 454 | "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: 
60809diff
changeset | 455 | unfolding measure_def by (subst emeasure_lim_emb) auto | 
| 50088 | 456 | |
| 457 | end | |
| 458 | ||
| 50090 | 459 | hide_const (open) PiF | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 460 | hide_const (open) Pi\<^sub>F | 
| 50090 | 461 | hide_const (open) Pi' | 
| 462 | hide_const (open) finmap_of | |
| 463 | hide_const (open) proj | |
| 464 | hide_const (open) domain | |
| 50245 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50244diff
changeset | 465 | hide_const (open) basis_finmap | 
| 50090 | 466 | |
| 61605 | 467 | sublocale polish_projective \<subseteq> P: prob_space lim | 
| 50088 | 468 | proof | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 469 |   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: 
60809diff
changeset | 470 | by (auto simp: prod_emb_def space_PiM) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62397diff
changeset | 471 |   interpret prob_space "P {}"
 | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 472 | 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: 
60809diff
changeset | 473 | 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: 
60809diff
changeset | 474 |     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: 
60809diff
changeset | 475 | by (simp add: * PiM_empty space_P) | 
| 50088 | 476 | qed | 
| 477 | ||
| 478 | locale polish_product_prob_space = | |
| 479 |   product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
 | |
| 480 | ||
| 481 | 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: 
63626diff
changeset | 482 | .. | 
| 50088 | 483 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
60809diff
changeset | 484 | 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: 
60809diff
changeset | 485 | by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_lim_emb) | 
| 50088 | 486 | |
| 487 | end |