| author | blanchet | 
| Tue, 20 May 2014 09:38:39 +0200 | |
| changeset 57013 | ed95456499e6 | 
| parent 56193 | c726ecfb22b6 | 
| child 57418 | 6ab1c7cb0b8d | 
| permissions | -rw-r--r-- | 
| 50091 | 1 | (* Title: HOL/Probability/Projective_Limit.thy | 
| 50088 | 2 | Author: Fabian Immler, TU München | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Projective Limit *}
 | |
| 6 | ||
| 7 | theory Projective_Limit | |
| 8 | imports | |
| 9 | Caratheodory | |
| 10 | Fin_Map | |
| 11 | Regularity | |
| 12 | Projective_Family | |
| 13 | Infinite_Product_Measure | |
| 50971 | 14 | "~~/src/HOL/Library/Diagonal_Subsequence" | 
| 50088 | 15 | begin | 
| 16 | ||
| 17 | subsection {* Sequences of Finite Maps in Compact Sets *}
 | |
| 18 | ||
| 19 | 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 | 20 | fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a)" and M | 
| 50088 | 21 | assumes compact: "\<And>n. compact (K n)" | 
| 22 |   assumes f_in_K: "\<And>n. K n \<noteq> {}"
 | |
| 23 | assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)" | |
| 24 | 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 | 25 | "\<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 | 26 | begin | 
| 27 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 28 | 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 | 29 | using proj_in_K f_in_K | 
| 30 | proof cases | |
| 31 | obtain k where "k \<in> K (Suc 0)" using f_in_K by auto | |
| 32 | assume "\<forall>n. t \<notin> domain (f n)" | |
| 33 | thus ?thesis | |
| 34 | by (auto intro!: exI[where x=1] image_eqI[OF _ `k \<in> K (Suc 0)`] | |
| 35 | simp: domain_K[OF `k \<in> K (Suc 0)`]) | |
| 36 | qed blast | |
| 37 | ||
| 38 | 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 | 39 | obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n" | 
| 50088 | 40 | using proj_in_K' by blast | 
| 41 | ||
| 42 | lemma compact_projset: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 43 | shows "compact ((\<lambda>k. (k)\<^sub>F i) ` K n)" | 
| 50088 | 44 | using continuous_proj compact by (rule compact_continuous_image) | 
| 45 | ||
| 46 | end | |
| 47 | ||
| 48 | lemma compactE': | |
| 50884 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 hoelzl parents: 
50252diff
changeset | 49 | fixes S :: "'a :: metric_space set" | 
| 50088 | 50 | assumes "compact S" "\<forall>n\<ge>m. f n \<in> S" | 
| 51 | obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially" | |
| 52 | proof atomize_elim | |
| 53 | have "subseq (op + m)" by (simp add: subseq_def) | |
| 54 | have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> S" using assms by auto | |
| 50884 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 hoelzl parents: 
50252diff
changeset | 55 | from seq_compactE[OF `compact S`[unfolded compact_eq_seq_compact_metric] this] guess l r . | 
| 50088 | 56 | hence "l \<in> S" "subseq ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) ----> l" | 
| 57 | using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def) | |
| 58 | thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) ----> l" by blast | |
| 59 | qed | |
| 60 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 61 | sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^sub>F n) ----> l)" | 
| 50088 | 62 | proof | 
| 63 | fix n s | |
| 64 | assume "subseq s" | |
| 65 | 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: 
52681diff
changeset | 66 | have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0" | 
| 50088 | 67 | proof safe | 
| 68 | fix i assume "n0 \<le> i" | |
| 69 | also have "\<dots> \<le> s i" by (rule seq_suble) fact | |
| 70 | 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 | 71 | with n0 show "((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0 " | 
| 50088 | 72 | by auto | 
| 73 | qed | |
| 74 | from compactE'[OF compact_projset this] guess ls rs . | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 75 | thus "\<exists>r'. subseq r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^sub>F n) ----> l)" by (auto simp: o_def) | 
| 50088 | 76 | qed | 
| 77 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 78 | lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^sub>F n) ----> l" | 
| 50088 | 79 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 80 | obtain l where "(\<lambda>i. ((f o (diagseq o op + (Suc n))) i)\<^sub>F n) ----> l" | 
| 52681 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 81 | proof (atomize_elim, rule diagseq_holds) | 
| 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 82 | fix r s n | 
| 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 83 | assume "subseq r" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 84 | assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) ----> l" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 85 | then obtain l where "((\<lambda>i. (f i)\<^sub>F n) o s) ----> l" | 
| 52681 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 86 | by (auto simp: o_def) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 87 | hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) ----> l" using `subseq r` | 
| 52681 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 88 | by (rule LIMSEQ_subseq_LIMSEQ) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 89 | thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^sub>F n) ----> l" by (auto simp add: o_def) | 
| 50088 | 90 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 91 | hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^sub>F n) ----> l" by (simp add: ac_simps) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 92 | hence "(\<lambda>i. (f (diagseq i))\<^sub>F n) ----> l" by (rule LIMSEQ_offset) | 
| 52681 
8cc7f76b827a
tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
 immler parents: 
51351diff
changeset | 93 | thus ?thesis .. | 
| 50088 | 94 | qed | 
| 95 | ||
| 96 | subsection {* Daniell-Kolmogorov Theorem *}
 | |
| 97 | ||
| 98 | text {* Existence of Projective Limit *}
 | |
| 99 | ||
| 100 | locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure" | |
| 101 | for I::"'i set" and P | |
| 102 | begin | |
| 103 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 104 | abbreviation "lim\<^sub>B \<equiv> (\<lambda>J P. limP J (\<lambda>_. borel) P)" | 
| 50088 | 105 | |
| 50125 
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
 hoelzl parents: 
50124diff
changeset | 106 | lemma emeasure_limB_emb_not_empty: | 
| 50088 | 107 |   assumes "I \<noteq> {}"
 | 
| 108 |   assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 109 | shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)" | 
| 50088 | 110 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 111 | let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space borel" | 
| 50088 | 112 | let ?G = generator | 
| 113 | interpret G!: algebra ?\<Omega> generator by (intro algebra_generator) fact | |
| 50252 | 114 | note mu_G_mono = | 
| 115 |     G.additive_increasing[OF positive_mu_G[OF `I \<noteq> {}`] additive_mu_G[OF `I \<noteq> {}`],
 | |
| 116 | THEN increasingD] | |
| 117 |   write mu_G  ("\<mu>G")
 | |
| 118 | ||
| 50088 | 119 | have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>" | 
| 50252 | 120 | proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G, | 
| 50088 | 121 |       OF `I \<noteq> {}`, OF `I \<noteq> {}`])
 | 
| 122 | fix A assume "A \<in> ?G" | |
| 123 | with generatorE guess J X . note JX = this | |
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 124 | interpret prob_space "P J" using proj_prob_space[OF `finite J`] . | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 125 | show "\<mu>G A \<noteq> \<infinity>" using JX by (simp add: limP_finite) | 
| 50088 | 126 | next | 
| 127 |     fix Z assume Z: "range Z \<subseteq> ?G" "decseq Z" "(\<Inter>i. Z i) = {}"
 | |
| 128 | then have "decseq (\<lambda>i. \<mu>G (Z i))" | |
| 50252 | 129 | by (auto intro!: mu_G_mono simp: decseq_def) | 
| 50088 | 130 | moreover | 
| 131 | have "(INF i. \<mu>G (Z i)) = 0" | |
| 132 | proof (rule ccontr) | |
| 133 | assume "(INF i. \<mu>G (Z i)) \<noteq> 0" (is "?a \<noteq> 0") | |
| 134 | moreover have "0 \<le> ?a" | |
| 50252 | 135 |         using Z positive_mu_G[OF `I \<noteq> {}`] by (auto intro!: INF_greatest simp: positive_def)
 | 
| 50088 | 136 | ultimately have "0 < ?a" by auto | 
| 137 | hence "?a \<noteq> -\<infinity>" by auto | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 138 |       have "\<forall>n. \<exists>J B. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> B \<in> sets (Pi\<^sub>M J (\<lambda>_. borel)) \<and>
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 139 | Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (lim\<^sub>B J P) B" | 
| 50088 | 140 | using Z by (intro allI generator_Ex) auto | 
| 141 |       then obtain J' B' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I"
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 142 | "\<And>n. B' n \<in> sets (\<Pi>\<^sub>M i\<in>J' n. borel)" | 
| 50088 | 143 | and Z_emb: "\<And>n. Z n = emb I (J' n) (B' n)" | 
| 144 | unfolding choice_iff by blast | |
| 145 | moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)" | |
| 146 | moreover def B \<equiv> "\<lambda>n. emb (J n) (J' n) (B' n)" | |
| 147 |       ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I"
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 148 | "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)" | 
| 50088 | 149 | by auto | 
| 150 | have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m" | |
| 151 | unfolding J_def by force | |
| 152 | have "\<forall>n. \<exists>j. j \<in> J n" using J by blast | |
| 153 | then obtain j where j: "\<And>n. j n \<in> J n" | |
| 154 | unfolding choice_iff by blast | |
| 155 | note [simp] = `\<And>n. finite (J n)` | |
| 156 | from J Z_emb have Z_eq: "\<And>n. Z n = emb I (J n) (B n)" "\<And>n. Z n \<in> ?G" | |
| 157 | unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto) | |
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 158 | interpret prob_space "P (J i)" for i using proj_prob_space by simp | 
| 50088 | 159 | have "?a \<le> \<mu>G (Z 0)" by (auto intro: INF_lower) | 
| 50252 | 160 | also have "\<dots> < \<infinity>" using J by (auto simp: Z_eq mu_G_eq limP_finite proj_sets) | 
| 50088 | 161 | finally have "?a \<noteq> \<infinity>" by simp | 
| 162 | have "\<And>n. \<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" unfolding Z_eq using J J_mono | |
| 50252 | 163 | by (subst mu_G_eq) (auto simp: limP_finite proj_sets mu_G_eq) | 
| 50088 | 164 | |
| 50243 | 165 | have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite) | 
| 166 | def Utn \<equiv> "to_nat_on (\<Union>n. J n)" | |
| 167 | interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>n. J n)" for n | |
| 168 | by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J]) | |
| 169 | have inj_on_Utn: "inj_on Utn (\<Union>n. J n)" | |
| 170 | unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on) | |
| 171 | hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto | |
| 50088 | 172 | def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)" | 
| 173 |       let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
 | |
| 174 |       {
 | |
| 175 | fix n | |
| 176 | interpret finite_measure "P (J n)" by unfold_locales | |
| 177 | have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))" | |
| 178 | using J by (auto simp: P'_def mapmeasure_PiM proj_space proj_sets) | |
| 179 | also | |
| 180 | have "\<dots> = ?SUP n" | |
| 181 | proof (rule inner_regular) | |
| 182 | show "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>" | |
| 183 | unfolding P'_def | |
| 184 | by (auto simp: P'_def mapmeasure_PiF fm_measurable proj_space proj_sets) | |
| 185 | show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def) | |
| 186 | next | |
| 187 | show "fm n ` B n \<in> sets borel" | |
| 188 | unfolding borel_eq_PiF_borel | |
| 189 | by (auto simp del: J(2) simp: P'_def fm_image_measurable_finite proj_sets J) | |
| 190 | qed | |
| 191 | finally | |
| 192 | have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto | |
| 193 | } note R = this | |
| 194 | have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a | |
| 195 | \<and> compact K \<and> K \<subseteq> fm n ` B n" | |
| 196 | proof | |
| 197 | fix n | |
| 198 | have "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>" | |
| 199 | by (simp add: mapmeasure_PiF P'_def proj_space proj_sets) | |
| 200 | then interpret finite_measure "P' n" .. | |
| 201 | show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and> | |
| 202 | compact K \<and> K \<subseteq> fm n ` B n" | |
| 203 | unfolding R | |
| 204 | proof (rule ccontr) | |
| 205 | assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n) * ?a \<and> | |
| 206 | compact K' \<and> K' \<subseteq> fm n ` B n)" | |
| 207 | have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a" | |
| 208 | proof (intro SUP_least) | |
| 209 | fix K | |
| 210 |             assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
 | |
| 211 | with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a" | |
| 212 | by auto | |
| 213 | hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a" | |
| 214 | unfolding not_less[symmetric] by simp | |
| 215 | hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K" | |
| 216 | using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps) | |
| 217 | thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp | |
| 218 | qed | |
| 219 | hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp | |
| 220 | hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def . | |
| 221 | hence "0 \<le> - (2 powr (-n) * ?a)" | |
| 222 | using `?SUP _ \<noteq> \<infinity>` `?SUP _ \<noteq> - \<infinity>` | |
| 223 | by (subst (asm) ereal_add_le_add_iff) (auto simp:) | |
| 224 | moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a` | |
| 225 | by (auto simp: ereal_zero_less_0_iff) | |
| 226 | ultimately show False by simp | |
| 227 | qed | |
| 228 | qed | |
| 229 | then obtain K' where K': | |
| 230 | "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a" | |
| 231 | "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n" | |
| 232 | unfolding choice_iff by blast | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 233 | def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 234 | have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))" | 
| 50088 | 235 | unfolding K_def | 
| 236 | using compact_imp_closed[OF `compact (K' _)`] | |
| 237 | by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"]) | |
| 238 | (auto simp: borel_eq_PiF_borel[symmetric]) | |
| 239 | have K_B: "\<And>n. K n \<subseteq> B n" | |
| 240 | proof | |
| 241 | fix x n | |
| 242 | assume "x \<in> K n" hence fm_in: "fm n x \<in> fm n ` B n" | |
| 243 | using K' by (force simp: K_def) | |
| 244 | show "x \<in> B n" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50243diff
changeset | 245 | using `x \<in> K n` K_sets sets.sets_into_space J[of n] | 
| 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50243diff
changeset | 246 | by (intro inj_on_image_mem_iff[OF inj_on_fm _ fm_in, of "\<lambda>_. borel"]) auto | 
| 50088 | 247 | qed | 
| 248 | def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)" | |
| 249 | have Z': "\<And>n. Z' n \<subseteq> Z n" | |
| 250 | unfolding Z_eq unfolding Z'_def | |
| 251 | proof (rule prod_emb_mono, safe) | |
| 252 | fix n x assume "x \<in> K n" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 253 | hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" | 
| 50088 | 254 | by (simp_all add: K_def proj_space) | 
| 255 | note this(1) | |
| 256 | also have "K' n \<subseteq> fm n ` B n" by (simp add: K') | |
| 257 | finally have "fm n x \<in> fm n ` B n" . | |
| 258 | thus "x \<in> B n" | |
| 259 | proof safe | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 260 | fix y assume y: "y \<in> B n" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 261 | hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"] | 
| 50088 | 262 | by (auto simp add: proj_space proj_sets) | 
| 263 | assume "fm n x = fm n y" | |
| 264 | note inj_onD[OF inj_on_fm[OF space_borel], | |
| 265 | OF `fm n x = fm n y` `x \<in> space _` `y \<in> space _`] | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 266 | with y show "x \<in> B n" by simp | 
| 50088 | 267 | qed | 
| 268 | qed | |
| 269 |       { fix n
 | |
| 270 | have "Z' n \<in> ?G" using K' unfolding Z'_def | |
| 271 | apply (intro generatorI'[OF J(1-3)]) | |
| 272 | unfolding K_def proj_space | |
| 273 | apply (rule measurable_sets[OF fm_measurable[of _ "Collect finite"]]) | |
| 274 | apply (auto simp add: P'_def borel_eq_PiF_borel[symmetric] compact_imp_closed) | |
| 275 | done | |
| 276 | } | |
| 277 |       def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
 | |
| 278 | hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def) | |
| 279 | hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add) | |
| 280 | have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def) | |
| 281 | hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto | |
| 282 |       have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
 | |
| 283 | proof - | |
| 284 | fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact | |
| 285 |         have "Y n = (\<Inter> i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
 | |
| 286 | by (auto simp: Y_def Z'_def) | |
| 287 |         also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter> i\<in>{1..n}. emb (J n) (J i) (K i))"
 | |
| 288 | using `n \<ge> 1` | |
| 289 | by (subst prod_emb_INT) auto | |
| 290 | finally | |
| 291 | have Y_emb: | |
| 292 | "Y n = prod_emb I (\<lambda>_. borel) (J n) | |
| 293 |             (\<Inter> i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
 | |
| 294 | hence "Y n \<in> ?G" using J J_mono K_sets `n \<ge> 1` by (intro generatorI[OF _ _ _ _ Y_emb]) auto | |
| 295 | hence "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" unfolding Y_emb using J J_mono K_sets `n \<ge> 1` | |
| 50252 | 296 | by (subst mu_G_eq) (auto simp: limP_finite proj_sets mu_G_eq) | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 297 | interpret finite_measure "(limP (J n) (\<lambda>_. borel) P)" | 
| 50088 | 298 | proof | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 299 | have "emeasure (limP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^sub>E space borel) \<noteq> \<infinity>" | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 300 | using J by (subst emeasure_limP) auto | 
| 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 301 | thus "emeasure (limP (J n) (\<lambda>_. borel) P) (space (limP (J n) (\<lambda>_. borel) P)) \<noteq> \<infinity>" | 
| 50088 | 302 | by (simp add: space_PiM) | 
| 303 | qed | |
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 304 | have "\<mu>G (Z n) = limP (J n) (\<lambda>_. borel) P (B n)" | 
| 50252 | 305 | unfolding Z_eq using J by (auto simp: mu_G_eq) | 
| 50088 | 306 | moreover have "\<mu>G (Y n) = | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 307 |           limP (J n) (\<lambda>_. borel) P (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
 | 
| 50252 | 308 | unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst mu_G_eq) auto | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 309 | moreover have "\<mu>G (Z n - Y n) = limP (J n) (\<lambda>_. borel) P | 
| 50088 | 310 |           (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
 | 
| 311 | unfolding Z_eq Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1` | |
| 50252 | 312 | by (subst mu_G_eq) (auto intro!: sets.Diff) | 
| 50088 | 313 | ultimately | 
| 314 | have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)" | |
| 315 | using J J_mono K_sets `n \<ge> 1` | |
| 316 | by (simp only: emeasure_eq_measure) | |
| 317 | (auto dest!: bspec[where x=n] | |
| 318 | simp: extensional_restrict emeasure_eq_measure prod_emb_iff | |
| 319 | intro!: measure_Diff[symmetric] set_mp[OF K_B]) | |
| 320 |         also have subs: "Z n - Y n \<subseteq> (\<Union> i\<in>{1..n}. (Z i - Z' i))" using Z' Z `n \<ge> 1`
 | |
| 321 | unfolding Y_def by (force simp: decseq_def) | |
| 322 |         have "Z n - Y n \<in> ?G" "(\<Union> i\<in>{1..n}. (Z i - Z' i)) \<in> ?G"
 | |
| 323 | using `Z' _ \<in> ?G` `Z _ \<in> ?G` `Y _ \<in> ?G` by auto | |
| 324 |         hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union> i\<in>{1..n}. (Z i - Z' i))"
 | |
| 50252 | 325 |           using subs G.additive_increasing[OF positive_mu_G[OF `I \<noteq> {}`] additive_mu_G[OF `I \<noteq> {}`]]
 | 
| 50088 | 326 | unfolding increasing_def by auto | 
| 327 |         also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> ?G` `Z' _ \<in> ?G`
 | |
| 50252 | 328 |           by (intro G.subadditive[OF positive_mu_G additive_mu_G, OF `I \<noteq> {}` `I \<noteq> {}`]) auto
 | 
| 50088 | 329 |         also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
 | 
| 330 | proof (rule setsum_mono) | |
| 331 |           fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
 | |
| 332 | have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))" | |
| 333 | unfolding Z'_def Z_eq by simp | |
| 334 | also have "\<dots> = P (J i) (B i - K i)" | |
| 50252 | 335 | apply (subst mu_G_eq) using J K_sets apply auto | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 336 | apply (subst limP_finite) apply auto | 
| 50088 | 337 | done | 
| 338 | also have "\<dots> = P (J i) (B i) - P (J i) (K i)" | |
| 339 | apply (subst emeasure_Diff) using K_sets J `K _ \<subseteq> B _` apply (auto simp: proj_sets) | |
| 340 | done | |
| 341 | also have "\<dots> = P (J i) (B i) - P' i (K' i)" | |
| 342 | unfolding K_def P'_def | |
| 343 | by (auto simp: mapmeasure_PiF proj_space proj_sets borel_eq_PiF_borel[symmetric] | |
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50101diff
changeset | 344 | compact_imp_closed[OF `compact (K' _)`] space_PiM PiE_def) | 
| 50088 | 345 | also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] . | 
| 346 | finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" . | |
| 347 | qed | |
| 348 |         also have "\<dots> = (\<Sum> i\<in>{1..n}. ereal (2 powr -real i) * ereal(real ?a))"
 | |
| 349 | using `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` by (subst ereal_real') auto | |
| 350 |         also have "\<dots> = ereal (\<Sum> i\<in>{1..n}. (2 powr -real i) * (real ?a))" by simp
 | |
| 351 |         also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real i)) * real ?a)"
 | |
| 352 | by (simp add: setsum_left_distrib) | |
| 353 | also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps | |
| 354 | proof (rule mult_strict_right_mono) | |
| 355 |           have "(\<Sum>i\<in>{1..n}. 2 powr - real i) = (\<Sum>i\<in>{1..<Suc n}. (1/2) ^ i)"
 | |
| 356 | by (rule setsum_cong) | |
| 357 | (auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide) | |
| 56193 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
53374diff
changeset | 358 |           also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
 | 
| 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
53374diff
changeset | 359 |           also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
 | 
| 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
53374diff
changeset | 360 |             setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
 | 
| 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
53374diff
changeset | 361 | also have "\<dots> < 1" by (subst geometric_sum) auto | 
| 50088 | 362 | finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" . | 
| 363 | qed (auto simp: | |
| 364 | `0 < ?a` `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` ereal_less_real_iff zero_ereal_def[symmetric]) | |
| 365 | also have "\<dots> = ?a" using `0 < ?a` `?a \<noteq> \<infinity>` by (auto simp: ereal_real') | |
| 366 | also have "\<dots> \<le> \<mu>G (Z n)" by (auto intro: INF_lower) | |
| 367 | finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" . | |
| 368 | hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)" | |
| 369 | using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less) | |
| 370 | have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto | |
| 371 | also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))" | |
| 372 | apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto | |
| 373 | finally have "\<mu>G (Y n) > 0" | |
| 374 | using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric]) | |
| 50252 | 375 |         thus "Y n \<noteq> {}" using positive_mu_G `I \<noteq> {}` by (auto simp add: positive_def)
 | 
| 50088 | 376 | qed | 
| 377 |       hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
 | |
| 378 | then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force | |
| 379 |       {
 | |
| 380 | fix t and n m::nat | |
| 381 | assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp | |
| 382 | from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto | |
| 383 | also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] . | |
| 384 | finally | |
| 385 | have "fm n (restrict (y m) (J n)) \<in> K' n" | |
| 386 | unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff) | |
| 387 | moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)" | |
| 388 | using J by (simp add: fm_def) | |
| 389 | ultimately have "fm n (y m) \<in> K' n" by simp | |
| 390 | } note fm_in_K' = this | |
| 391 | interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel | |
| 392 | proof | |
| 393 | fix n show "compact (K' n)" by fact | |
| 394 | next | |
| 395 | fix n | |
| 396 | from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto | |
| 397 | also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto | |
| 398 | finally | |
| 399 | have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)" | |
| 400 | unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff) | |
| 401 |         thus "K' (Suc n) \<noteq> {}" by auto
 | |
| 402 | fix k | |
| 403 | assume "k \<in> K' (Suc n)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50243diff
changeset | 404 | with K'[of "Suc n"] sets.sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto | 
| 50088 | 405 | then obtain b where "k = fm (Suc n) b" by auto | 
| 406 | thus "domain k = domain (fm (Suc n) (y (Suc n)))" | |
| 407 | by (simp_all add: fm_def) | |
| 408 | next | |
| 409 | fix t and n m::nat | |
| 410 | assume "n \<le> m" hence "Suc n \<le> Suc m" by simp | |
| 411 | assume "t \<in> domain (fm (Suc n) (y (Suc n)))" | |
| 412 | then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto | |
| 413 | hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto | |
| 414 | have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m` | |
| 415 | by (intro fm_in_K') simp_all | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 416 | show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)" | 
| 50088 | 417 | apply (rule image_eqI[OF _ img]) | 
| 418 | using `j \<in> J (Suc n)` `j \<in> J (Suc m)` | |
| 419 | unfolding j by (subst proj_fm, auto)+ | |
| 420 | qed | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 421 | have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z" | 
| 50088 | 422 | using diagonal_tendsto .. | 
| 423 | then obtain z where z: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 424 | "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z t" | 
| 50088 | 425 | unfolding choice_iff by blast | 
| 426 |       {
 | |
| 427 | fix n :: nat assume "n \<ge> 1" | |
| 428 | have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)" | |
| 429 | by simp | |
| 430 | moreover | |
| 431 |         {
 | |
| 432 | fix t | |
| 433 | assume t: "t \<in> domain (finmap_of (Utn ` J n) z)" | |
| 434 | hence "t \<in> Utn ` J n" by simp | |
| 435 | then obtain j where j: "t = Utn j" "j \<in> J n" by auto | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 436 | have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> z t" | 
| 50088 | 437 | apply (subst (2) tendsto_iff, subst eventually_sequentially) | 
| 438 | proof safe | |
| 439 | fix e :: real assume "0 < e" | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 440 |             { fix i x
 | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 441 | assume i: "i \<ge> n" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 442 | assume "t \<in> domain (fm n x)" | 
| 50088 | 443 | hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 444 | with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t" | 
| 50243 | 445 | using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn]) | 
| 50088 | 446 | } note index_shift = this | 
| 447 | have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n" | |
| 448 | apply (rule le_SucI) | |
| 449 | apply (rule order_trans) apply simp | |
| 450 | apply (rule seq_suble[OF subseq_diagseq]) | |
| 451 | done | |
| 452 | from z | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 453 | have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" | 
| 50088 | 454 | unfolding tendsto_iff eventually_sequentially using `0 < e` by auto | 
| 455 | then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow> | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 456 | dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 457 | show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e " | 
| 50088 | 458 | proof (rule exI[where x="max N n"], safe) | 
| 459 | fix na assume "max N n \<le> na" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 460 | hence "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 461 | dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t | 
| 50088 | 462 | by (subst index_shift[OF I]) auto | 
| 463 | also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 464 | finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" . | 
| 50088 | 465 | qed | 
| 466 | qed | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 467 | hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> (finmap_of (Utn ` J n) z)\<^sub>F t" | 
| 50088 | 468 | by (simp add: tendsto_intros) | 
| 469 | } ultimately | |
| 470 | have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z" | |
| 471 | by (rule tendsto_finmap) | |
| 472 | hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) ----> finmap_of (Utn ` J n) z" | |
| 473 | by (intro lim_subseq) (simp add: subseq_def) | |
| 474 | moreover | |
| 475 | have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)" | |
| 476 | apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI) | |
| 477 | apply (rule le_trans) | |
| 478 | apply (rule le_add2) | |
| 479 | using seq_suble[OF subseq_diagseq] | |
| 480 | apply auto | |
| 481 | done | |
| 482 | moreover | |
| 483 | from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed) | |
| 484 | ultimately | |
| 485 | have "finmap_of (Utn ` J n) z \<in> K' n" | |
| 486 | unfolding closed_sequential_limits by blast | |
| 487 | also have "finmap_of (Utn ` J n) z = fm n (\<lambda>i. z (Utn i))" | |
| 50243 | 488 | unfolding finmap_eq_iff | 
| 489 | proof clarsimp | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 490 | fix i assume i: "i \<in> J n" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 491 | hence "from_nat_into (\<Union>n. J n) (Utn i) = i" | 
| 50243 | 492 | unfolding Utn_def | 
| 493 | by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 494 | with i show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)" | 
| 50243 | 495 | by (simp add: finmap_eq_iff fm_def compose_def) | 
| 496 | qed | |
| 50088 | 497 | finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" . | 
| 498 | moreover | |
| 499 | let ?J = "\<Union>n. J n" | |
| 500 | have "(?J \<inter> J n) = J n" by auto | |
| 501 | ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n" | |
| 502 | unfolding K_def by (auto simp: proj_space space_PiM) | |
| 503 | hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def | |
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50101diff
changeset | 504 | using J by (auto simp: prod_emb_def PiE_def extensional_def) | 
| 50088 | 505 | also have "\<dots> \<subseteq> Z n" using Z' by simp | 
| 506 | finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" . | |
| 507 | } note in_Z = this | |
| 508 |       hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
 | |
| 509 |       hence "(\<Inter>i. Z i) \<noteq> {}" using Z INT_decseq_offset[OF `decseq Z`] by simp
 | |
| 510 | thus False using Z by simp | |
| 511 | qed | |
| 512 | ultimately show "(\<lambda>i. \<mu>G (Z i)) ----> 0" | |
| 51351 | 513 | using LIMSEQ_INF[of "\<lambda>i. \<mu>G (Z i)"] by simp | 
| 50088 | 514 | qed | 
| 515 | then guess \<mu> .. note \<mu> = this | |
| 516 | def f \<equiv> "finmap_of J B" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 517 | show "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)" | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 518 | proof (subst emeasure_extend_measure_Pair[OF limP_def, of I "\<lambda>_. borel" \<mu>]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 519 | show "positive (sets (lim\<^sub>B I P)) \<mu>" "countably_additive (sets (lim\<^sub>B I P)) \<mu>" | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 520 | using \<mu> unfolding sets_limP sets_PiM_generator by (auto simp: measure_space_def) | 
| 50088 | 521 | next | 
| 522 |     show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> B \<in> J \<rightarrow> sets borel"
 | |
| 523 | using assms by (auto simp: f_def) | |
| 524 | next | |
| 525 | fix J and X::"'i \<Rightarrow> 'a set" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 526 | show "prod_emb I (\<lambda>_. borel) J (Pi\<^sub>E J X) \<in> Pow (I \<rightarrow>\<^sub>E space borel)" | 
| 50088 | 527 | by (auto simp: prod_emb_def) | 
| 528 |     assume JX: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> J \<rightarrow> sets borel"
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 529 | hence "emb I J (Pi\<^sub>E J X) \<in> generator" using assms | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 530 | by (intro generatorI[where J=J and X="Pi\<^sub>E J X"]) (auto intro: sets_PiM_I_finite) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 531 | hence "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))" using \<mu> by simp | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 532 | also have "\<dots> = emeasure (P J) (Pi\<^sub>E J X)" | 
| 50088 | 533 | using JX assms proj_sets | 
| 50252 | 534 | by (subst mu_G_eq) (auto simp: mu_G_eq limP_finite intro: sets_PiM_I_finite) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 535 | finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = emeasure (P J) (Pi\<^sub>E J X)" . | 
| 50088 | 536 | next | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 537 | show "emeasure (P J) (Pi\<^sub>E J B) = emeasure (limP J (\<lambda>_. borel) P) (Pi\<^sub>E J B)" | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 538 | using assms by (simp add: f_def limP_finite Pi_def) | 
| 50088 | 539 | qed | 
| 540 | qed | |
| 541 | ||
| 542 | end | |
| 543 | ||
| 50090 | 544 | 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 | 545 | hide_const (open) Pi\<^sub>F | 
| 50090 | 546 | hide_const (open) Pi' | 
| 547 | hide_const (open) Abs_finmap | |
| 548 | hide_const (open) Rep_finmap | |
| 549 | hide_const (open) finmap_of | |
| 550 | hide_const (open) proj | |
| 551 | hide_const (open) domain | |
| 50245 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50244diff
changeset | 552 | hide_const (open) basis_finmap | 
| 50090 | 553 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 554 | sublocale polish_projective \<subseteq> P!: prob_space "(lim\<^sub>B I P)" | 
| 50088 | 555 | proof | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 556 | show "emeasure (lim\<^sub>B I P) (space (lim\<^sub>B I P)) = 1" | 
| 50088 | 557 | proof cases | 
| 558 |     assume "I = {}"
 | |
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 559 |     interpret prob_space "P {}" using proj_prob_space by simp
 | 
| 50088 | 560 | show ?thesis | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 561 |       by (simp add: space_PiM_empty limP_finite emeasure_space_1 `I = {}`)
 | 
| 50088 | 562 | next | 
| 563 |     assume "I \<noteq> {}"
 | |
| 564 | then obtain i where "i \<in> I" by auto | |
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 565 |     interpret prob_space "P {i}" using proj_prob_space by simp
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 566 |     have R: "(space (lim\<^sub>B I P)) = (emb I {i} (Pi\<^sub>E {i} (\<lambda>_. space borel)))"
 | 
| 50088 | 567 | by (auto simp: prod_emb_def space_PiM) | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50101diff
changeset | 568 |     moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM PiE_def)
 | 
| 50088 | 569 | ultimately show ?thesis using `i \<in> I` | 
| 570 | apply (subst R) | |
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 571 | apply (subst emeasure_limB_emb_not_empty) | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50101diff
changeset | 572 | apply (auto simp: limP_finite emeasure_space_1 PiE_def) | 
| 50088 | 573 | done | 
| 574 | qed | |
| 575 | qed | |
| 576 | ||
| 577 | context polish_projective begin | |
| 578 | ||
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 579 | lemma emeasure_limB_emb: | 
| 50088 | 580 | assumes X: "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 581 | shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (P J) (Pi\<^sub>E J B)" | 
| 50088 | 582 | proof cases | 
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 583 |   interpret prob_space "P {}" using proj_prob_space by simp
 | 
| 50088 | 584 |   assume "J = {}"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 585 |   moreover have "emb I {} {\<lambda>x. undefined} = space (lim\<^sub>B I P)"
 | 
| 50088 | 586 | by (auto simp: space_PiM prod_emb_def) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 587 |   moreover have "{\<lambda>x. undefined} = space (lim\<^sub>B {} P)"
 | 
| 50088 | 588 | by (auto simp: space_PiM prod_emb_def) | 
| 589 | ultimately show ?thesis | |
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 590 | by (simp add: P.emeasure_space_1 limP_finite emeasure_space_1 del: space_limP) | 
| 50088 | 591 | next | 
| 592 |   assume "J \<noteq> {}" with X show ?thesis
 | |
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 593 | by (subst emeasure_limB_emb_not_empty) (auto simp: limP_finite) | 
| 50088 | 594 | qed | 
| 595 | ||
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 596 | lemma measure_limB_emb: | 
| 50088 | 597 | assumes "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 598 | shows "measure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = measure (P J) (Pi\<^sub>E J B)" | 
| 50088 | 599 | proof - | 
| 50101 
a3bede207a04
renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space
 hoelzl parents: 
50095diff
changeset | 600 | interpret prob_space "P J" using proj_prob_space assms by simp | 
| 50088 | 601 | show ?thesis | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 602 | using emeasure_limB_emb[OF assms] | 
| 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 603 | unfolding emeasure_eq_measure limP_finite[OF `finite J` `J \<subseteq> I`] P.emeasure_eq_measure | 
| 50088 | 604 | by simp | 
| 605 | qed | |
| 606 | ||
| 607 | end | |
| 608 | ||
| 609 | locale polish_product_prob_space = | |
| 610 |   product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
 | |
| 611 | ||
| 612 | sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
 | |
| 613 | proof qed | |
| 614 | ||
| 50125 
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
 hoelzl parents: 
50124diff
changeset | 615 | lemma (in polish_product_prob_space) limP_eq_PiM: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52681diff
changeset | 616 |   "I \<noteq> {} \<Longrightarrow> lim\<^sub>P I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
 | 
| 50088 | 617 | PiM I (\<lambda>_. borel)" | 
| 50095 
94d7dfa9f404
renamed to more appropriate lim_P for projective limit
 immler parents: 
50091diff
changeset | 618 | by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_limB_emb) | 
| 50088 | 619 | |
| 620 | end |