added projective limit;
authorimmler
Thu, 15 Nov 2012 11:16:58 +0100
changeset 5110332d1795cc77a
parent 51102 635d73673b5e
child 51104 1badf63e5d97
added projective limit;
proof is based on auxiliary type finmap::polish_space
src/HOL/Probability/Fin_Map.thy
src/HOL/Probability/Probability.thy
src/HOL/Probability/Projective_Limit.thy
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/Probability/Fin_Map.thy	Thu Nov 15 11:16:58 2012 +0100
     1.3 @@ -0,0 +1,1503 @@
     1.4 +(*  Title:      HOL/Probability/Projective_Family.thy
     1.5 +    Author:     Fabian Immler, TU München
     1.6 +*)
     1.7 +
     1.8 +theory Fin_Map
     1.9 +imports Finite_Product_Measure
    1.10 +begin
    1.11 +
    1.12 +section {* Finite Maps *}
    1.13 +
    1.14 +text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of
    1.15 +  projective limit. @{const extensional} functions are used for the representation in order to
    1.16 +  stay close to the developments of (finite) products @{const Pi\<^isub>E} and their sigma-algebra
    1.17 +  @{const Pi\<^isub>M}. *}
    1.18 +
    1.19 +typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^isub>F /_)" [22, 21] 21) =
    1.20 +  "{(I::'i set, f::'i \<Rightarrow> 'a). finite I \<and> f \<in> extensional I}" by auto
    1.21 +
    1.22 +subsection {* Domain and Application *}
    1.23 +
    1.24 +definition domain where "domain P = fst (Rep_finmap P)"
    1.25 +
    1.26 +lemma finite_domain[simp, intro]: "finite (domain P)"
    1.27 +  by (cases P) (auto simp: domain_def Abs_finmap_inverse)
    1.28 +
    1.29 +definition proj ("_\<^isub>F" [1000] 1000) where "proj P i = snd (Rep_finmap P) i"
    1.30 +
    1.31 +declare [[coercion proj]]
    1.32 +
    1.33 +lemma extensional_proj[simp, intro]: "(P)\<^isub>F \<in> extensional (domain P)"
    1.34 +  by (cases P) (auto simp: domain_def Abs_finmap_inverse proj_def[abs_def])
    1.35 +
    1.36 +lemma proj_undefined[simp, intro]: "i \<notin> domain P \<Longrightarrow> P i = undefined"
    1.37 +  using extensional_proj[of P] unfolding extensional_def by auto
    1.38 +
    1.39 +lemma finmap_eq_iff: "P = Q \<longleftrightarrow> (domain P = domain Q \<and> (\<forall>i\<in>domain P. P i = Q i))"
    1.40 +  by (cases P, cases Q)
    1.41 +     (auto simp add: Abs_finmap_inject extensional_def domain_def proj_def Abs_finmap_inverse
    1.42 +              intro: extensionalityI)
    1.43 +
    1.44 +subsection {* Countable Finite Maps *}
    1.45 +
    1.46 +instance finmap :: (countable, countable) countable
    1.47 +proof
    1.48 +  obtain mapper where mapper: "\<And>fm :: 'a \<Rightarrow>\<^isub>F 'b. set (mapper fm) = domain fm"
    1.49 +    by (metis finite_list[OF finite_domain])
    1.50 +  have "inj (\<lambda>fm. map (\<lambda>i. (i, (fm)\<^isub>F i)) (mapper fm))" (is "inj ?F")
    1.51 +  proof (rule inj_onI)
    1.52 +    fix f1 f2 assume "?F f1 = ?F f2"
    1.53 +    then have "map fst (?F f1) = map fst (?F f2)" by simp
    1.54 +    then have "mapper f1 = mapper f2" by (simp add: comp_def)
    1.55 +    then have "domain f1 = domain f2" by (simp add: mapper[symmetric])
    1.56 +    with `?F f1 = ?F f2` show "f1 = f2"
    1.57 +      unfolding `mapper f1 = mapper f2` map_eq_conv mapper
    1.58 +      by (simp add: finmap_eq_iff)
    1.59 +  qed
    1.60 +  then show "\<exists>to_nat :: 'a \<Rightarrow>\<^isub>F 'b \<Rightarrow> nat. inj to_nat"
    1.61 +    by (intro exI[of _ "to_nat \<circ> ?F"] inj_comp) auto
    1.62 +qed
    1.63 +
    1.64 +subsection {* Constructor of Finite Maps *}
    1.65 +
    1.66 +definition "finmap_of inds f = Abs_finmap (inds, restrict f inds)"
    1.67 +
    1.68 +lemma proj_finmap_of[simp]:
    1.69 +  assumes "finite inds"
    1.70 +  shows "(finmap_of inds f)\<^isub>F = restrict f inds"
    1.71 +  using assms
    1.72 +  by (auto simp: Abs_finmap_inverse finmap_of_def proj_def)
    1.73 +
    1.74 +lemma domain_finmap_of[simp]:
    1.75 +  assumes "finite inds"
    1.76 +  shows "domain (finmap_of inds f) = inds"
    1.77 +  using assms
    1.78 +  by (auto simp: Abs_finmap_inverse finmap_of_def domain_def)
    1.79 +
    1.80 +lemma finmap_of_eq_iff[simp]:
    1.81 +  assumes "finite i" "finite j"
    1.82 +  shows "finmap_of i m = finmap_of j n \<longleftrightarrow> i = j \<and> restrict m i = restrict n i"
    1.83 +  using assms
    1.84 +  apply (auto simp: finmap_eq_iff restrict_def) by metis
    1.85 +
    1.86 +lemma
    1.87 +  finmap_of_inj_on_extensional_finite:
    1.88 +  assumes "finite K"
    1.89 +  assumes "S \<subseteq> extensional K"
    1.90 +  shows "inj_on (finmap_of K) S"
    1.91 +proof (rule inj_onI)
    1.92 +  fix x y::"'a \<Rightarrow> 'b"
    1.93 +  assume "finmap_of K x = finmap_of K y"
    1.94 +  hence "(finmap_of K x)\<^isub>F = (finmap_of K y)\<^isub>F" by simp
    1.95 +  moreover
    1.96 +  assume "x \<in> S" "y \<in> S" hence "x \<in> extensional K" "y \<in> extensional K" using assms by auto
    1.97 +  ultimately
    1.98 +  show "x = y" using assms by (simp add: extensional_restrict)
    1.99 +qed
   1.100 +
   1.101 +lemma finmap_choice:
   1.102 +  assumes *: "\<And>i. i \<in> I \<Longrightarrow> \<exists>x. P i x" and I: "finite I"
   1.103 +  shows "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. P i (fm i))"
   1.104 +proof -
   1.105 +  have "\<exists>f. \<forall>i\<in>I. P i (f i)"
   1.106 +    unfolding bchoice_iff[symmetric] using * by auto
   1.107 +  then guess f ..
   1.108 +  with I show ?thesis
   1.109 +    by (intro exI[of _ "finmap_of I f"]) auto
   1.110 +qed
   1.111 +
   1.112 +subsection {* Product set of Finite Maps *}
   1.113 +
   1.114 +text {* This is @{term Pi} for Finite Maps, most of this is copied *}
   1.115 +
   1.116 +definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^isub>F 'a) set" where
   1.117 +  "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^isub>F i \<in> A i) } "
   1.118 +
   1.119 +syntax
   1.120 +  "_Pi'"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PI' _:_./ _)" 10)
   1.121 +
   1.122 +syntax (xsymbols)
   1.123 +  "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>' _\<in>_./ _)"   10)
   1.124 +
   1.125 +syntax (HTML output)
   1.126 +  "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>' _\<in>_./ _)"   10)
   1.127 +
   1.128 +translations
   1.129 +  "PI' x:A. B" == "CONST Pi' A (%x. B)"
   1.130 +
   1.131 +abbreviation
   1.132 +  finmapset :: "['a set, 'b set] => ('a \<Rightarrow>\<^isub>F 'b) set"
   1.133 +    (infixr "~>" 60) where
   1.134 +  "A ~> B \<equiv> Pi' A (%_. B)"
   1.135 +
   1.136 +notation (xsymbols)
   1.137 +  finmapset  (infixr "\<leadsto>" 60)
   1.138 +
   1.139 +subsubsection{*Basic Properties of @{term Pi'}*}
   1.140 +
   1.141 +lemma Pi'_I[intro!]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
   1.142 +  by (simp add: Pi'_def)
   1.143 +
   1.144 +lemma Pi'_I'[simp]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
   1.145 +  by (simp add:Pi'_def)
   1.146 +
   1.147 +lemma finmapsetI: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> f \<in> A \<leadsto> B"
   1.148 +  by (simp add: Pi_def)
   1.149 +
   1.150 +lemma Pi'_mem: "f\<in> Pi' A B \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<in> B x"
   1.151 +  by (simp add: Pi'_def)
   1.152 +
   1.153 +lemma Pi'_iff: "f \<in> Pi' I X \<longleftrightarrow> domain f = I \<and> (\<forall>i\<in>I. f i \<in> X i)"
   1.154 +  unfolding Pi'_def by auto
   1.155 +
   1.156 +lemma Pi'E [elim]:
   1.157 +  "f \<in> Pi' A B \<Longrightarrow> (f x \<in> B x \<Longrightarrow> domain f = A \<Longrightarrow> Q) \<Longrightarrow> (x \<notin> A \<Longrightarrow> Q) \<Longrightarrow> Q"
   1.158 +  by(auto simp: Pi'_def)
   1.159 +
   1.160 +lemma in_Pi'_cong:
   1.161 +  "domain f = domain g \<Longrightarrow> (\<And> w. w \<in> A \<Longrightarrow> f w = g w) \<Longrightarrow> f \<in> Pi' A B \<longleftrightarrow> g \<in> Pi' A B"
   1.162 +  by (auto simp: Pi'_def)
   1.163 +
   1.164 +lemma funcset_mem: "[|f \<in> A \<leadsto> B; x \<in> A|] ==> f x \<in> B"
   1.165 +  by (simp add: Pi'_def)
   1.166 +
   1.167 +lemma funcset_image: "f \<in> A \<leadsto> B ==> f ` A \<subseteq> B"
   1.168 +by auto
   1.169 +
   1.170 +lemma Pi'_eq_empty[simp]:
   1.171 +  assumes "finite A" shows "(Pi' A B) = {} \<longleftrightarrow> (\<exists>x\<in>A. B x = {})"
   1.172 +  using assms
   1.173 +  apply (simp add: Pi'_def, auto)
   1.174 +  apply (drule_tac x = "finmap_of A (\<lambda>u. SOME y. y \<in> B u)" in spec, auto)
   1.175 +  apply (cut_tac P= "%y. y \<in> B i" in some_eq_ex, auto)
   1.176 +  done
   1.177 +
   1.178 +lemma Pi'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> Pi' A B \<subseteq> Pi' A C"
   1.179 +  by (auto simp: Pi'_def)
   1.180 +
   1.181 +lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^isub>E A B) = proj ` Pi' A B"
   1.182 +  apply (auto simp: Pi'_def Pi_def extensional_def)
   1.183 +  apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI)
   1.184 +  apply auto
   1.185 +  done
   1.186 +
   1.187 +subsection {* Metric Space of Finite Maps *}
   1.188 +
   1.189 +instantiation finmap :: (type, metric_space) metric_space
   1.190 +begin
   1.191 +
   1.192 +definition dist_finmap where
   1.193 +  "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
   1.194 +    card ((domain P - domain Q) \<union> (domain Q - domain P))"
   1.195 +
   1.196 +lemma dist_finmap_extend:
   1.197 +  assumes "finite X"
   1.198 +  shows "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q \<union> X. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
   1.199 +    card ((domain P - domain Q) \<union> (domain Q - domain P))"
   1.200 +    unfolding dist_finmap_def add_right_cancel
   1.201 +    using assms extensional_arb[of "(P)\<^isub>F"] extensional_arb[of "(Q)\<^isub>F" "domain Q"]
   1.202 +    by (intro setsum_mono_zero_cong_left) auto
   1.203 +
   1.204 +definition open_finmap :: "('a \<Rightarrow>\<^isub>F 'b) set \<Rightarrow> bool" where
   1.205 +  "open_finmap S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
   1.206 +
   1.207 +lemma add_eq_zero_iff[simp]:
   1.208 +  fixes a b::real
   1.209 +  assumes "a \<ge> 0" "b \<ge> 0"
   1.210 +  shows "a + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
   1.211 +using assms by auto
   1.212 +
   1.213 +lemma dist_le_1_imp_domain_eq:
   1.214 +  assumes "dist P Q < 1"
   1.215 +  shows "domain P = domain Q"
   1.216 +proof -
   1.217 +  have "0 \<le> (\<Sum>i\<in>domain P \<union> domain Q. dist (P i) (Q i))"
   1.218 +    by (simp add: setsum_nonneg)
   1.219 +  with assms have "card (domain P - domain Q \<union> (domain Q - domain P)) = 0"
   1.220 +    unfolding dist_finmap_def by arith
   1.221 +  thus "domain P = domain Q" by auto
   1.222 +qed
   1.223 +
   1.224 +lemma dist_proj:
   1.225 +  shows "dist ((x)\<^isub>F i) ((y)\<^isub>F i) \<le> dist x y"
   1.226 +proof -
   1.227 +  have "dist (x i) (y i) = (\<Sum>i\<in>{i}. dist (x i) (y i))" by simp
   1.228 +  also have "\<dots> \<le> (\<Sum>i\<in>domain x \<union> domain y \<union> {i}. dist (x i) (y i))"
   1.229 +    by (intro setsum_mono2) auto
   1.230 +  also have "\<dots> \<le> dist x y" by (simp add: dist_finmap_extend[of "{i}"])
   1.231 +  finally show ?thesis by simp
   1.232 +qed
   1.233 +
   1.234 +lemma open_Pi'I:
   1.235 +  assumes open_component: "\<And>i. i \<in> I \<Longrightarrow> open (A i)"
   1.236 +  shows "open (Pi' I A)"
   1.237 +proof (subst open_finmap_def, safe)
   1.238 +  fix x assume x: "x \<in> Pi' I A"
   1.239 +  hence dim_x: "domain x = I" by (simp add: Pi'_def)
   1.240 +  hence [simp]: "finite I" unfolding dim_x[symmetric] by simp
   1.241 +  have "\<exists>ei. \<forall>i\<in>I. 0 < ei i \<and> (\<forall>y. dist y (x i) < ei i \<longrightarrow> y \<in> A i)"
   1.242 +  proof (safe intro!: bchoice)
   1.243 +    fix i assume i: "i \<in> I"
   1.244 +    moreover with open_component have "open (A i)" by simp
   1.245 +    moreover have "x i \<in> A i" using x i
   1.246 +      by (auto simp: proj_def)
   1.247 +    ultimately show "\<exists>e>0. \<forall>y. dist y (x i) < e \<longrightarrow> y \<in> A i"
   1.248 +      using x by (auto simp: open_dist Ball_def)
   1.249 +  qed
   1.250 +  then guess ei .. note ei = this
   1.251 +  def es \<equiv> "ei ` I"
   1.252 +  def e \<equiv> "if es = {} then 0.5 else min 0.5 (Min es)"
   1.253 +  from ei have "e > 0" using x
   1.254 +    by (auto simp add: e_def es_def Pi'_def Ball_def)
   1.255 +  moreover have "\<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A"
   1.256 +  proof (intro allI impI)
   1.257 +    fix y
   1.258 +    assume "dist y x < e"
   1.259 +    also have "\<dots> < 1" by (auto simp: e_def)
   1.260 +    finally have "domain y = domain x" by (rule dist_le_1_imp_domain_eq)
   1.261 +    with dim_x have dims: "domain y = domain x" "domain x = I" by auto
   1.262 +    show "y \<in> Pi' I A"
   1.263 +    proof
   1.264 +      show "domain y = I" using dims by simp
   1.265 +    next
   1.266 +      fix i
   1.267 +      assume "i \<in> I"
   1.268 +      have "dist (y i) (x i) \<le> dist y x" using dims `i \<in> I`
   1.269 +        by (auto intro: dist_proj)
   1.270 +      also have "\<dots> < e" using `dist y x < e` dims
   1.271 +        by (simp add: dist_finmap_def)
   1.272 +      also have "e \<le> Min (ei ` I)" using dims `i \<in> I`
   1.273 +        by (auto simp: e_def es_def)
   1.274 +      also have "\<dots> \<le> ei i" using `i \<in> I` by (simp add: e_def)
   1.275 +      finally have "dist (y i) (x i) < ei i" .
   1.276 +      with ei `i \<in> I` show "y i \<in> A  i" by simp
   1.277 +    qed
   1.278 +  qed
   1.279 +  ultimately
   1.280 +  show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A" by blast
   1.281 +qed
   1.282 +
   1.283 +instance
   1.284 +proof
   1.285 +  fix S::"('a \<Rightarrow>\<^isub>F 'b) set"
   1.286 +  show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
   1.287 +    unfolding open_finmap_def ..
   1.288 +next
   1.289 +  fix P Q::"'a \<Rightarrow>\<^isub>F 'b"
   1.290 +  show "dist P Q = 0 \<longleftrightarrow> P = Q"
   1.291 +    by (auto simp: finmap_eq_iff dist_finmap_def setsum_nonneg setsum_nonneg_eq_0_iff)
   1.292 +next
   1.293 +  fix P Q R::"'a \<Rightarrow>\<^isub>F 'b"
   1.294 +  let ?symdiff = "\<lambda>a b. domain a - domain b \<union> (domain b - domain a)"
   1.295 +  def E \<equiv> "domain P \<union> domain Q \<union> domain R"
   1.296 +  hence "finite E" by (simp add: E_def)
   1.297 +  have "card (?symdiff P Q) \<le> card (?symdiff P R \<union> ?symdiff Q R)"
   1.298 +    by (auto intro: card_mono)
   1.299 +  also have "\<dots> \<le> card (?symdiff P R) + card (?symdiff Q R)"
   1.300 +    by (subst card_Un_Int) auto
   1.301 +  finally have "dist P Q \<le> (\<Sum>i\<in>E. dist (P i) (R i) + dist (Q i) (R i)) +
   1.302 +    real (card (?symdiff P R) + card (?symdiff Q R))"
   1.303 +    unfolding dist_finmap_extend[OF `finite E`]
   1.304 +    by (intro add_mono) (auto simp: E_def intro: setsum_mono dist_triangle_le)
   1.305 +  also have "\<dots> \<le> dist P R + dist Q R"
   1.306 +    unfolding dist_finmap_extend[OF `finite E`] by (simp add: ac_simps E_def setsum_addf[symmetric])
   1.307 +  finally show "dist P Q \<le> dist P R + dist Q R" by simp
   1.308 +qed
   1.309 +
   1.310 +end
   1.311 +
   1.312 +lemma open_restricted_space:
   1.313 +  shows "open {m. P (domain m)}"
   1.314 +proof -
   1.315 +  have "{m. P (domain m)} = (\<Union>i \<in> Collect P. {m. domain m = i})" by auto
   1.316 +  also have "open \<dots>"
   1.317 +  proof (rule, safe, cases)
   1.318 +    fix i::"'a set"
   1.319 +    assume "finite i"
   1.320 +    hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
   1.321 +    also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
   1.322 +    finally show "open {m. domain m = i}" .
   1.323 +  next
   1.324 +    fix i::"'a set"
   1.325 +    assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
   1.326 +    also have "open \<dots>" by simp
   1.327 +    finally show "open {m. domain m = i}" .
   1.328 +  qed
   1.329 +  finally show ?thesis .
   1.330 +qed
   1.331 +
   1.332 +lemma closed_restricted_space:
   1.333 +  shows "closed {m. P (domain m)}"
   1.334 +proof -
   1.335 +  have "{m. P (domain m)} = - (\<Union>i \<in> - Collect P. {m. domain m = i})" by auto
   1.336 +  also have "closed \<dots>"
   1.337 +  proof (rule, rule, rule, cases)
   1.338 +    fix i::"'a set"
   1.339 +    assume "finite i"
   1.340 +    hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
   1.341 +    also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
   1.342 +    finally show "open {m. domain m = i}" .
   1.343 +  next
   1.344 +    fix i::"'a set"
   1.345 +    assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
   1.346 +    also have "open \<dots>" by simp
   1.347 +    finally show "open {m. domain m = i}" .
   1.348 +  qed
   1.349 +  finally show ?thesis .
   1.350 +qed
   1.351 +
   1.352 +lemma continuous_proj:
   1.353 +  shows "continuous_on s (\<lambda>x. (x)\<^isub>F i)"
   1.354 +  unfolding continuous_on_topological
   1.355 +proof safe
   1.356 +  fix x B assume "x \<in> s" "open B" "x i \<in> B"
   1.357 +  let ?A = "Pi' (domain x) (\<lambda>j. if i = j then B else UNIV)"
   1.358 +  have "open ?A" using `open B` by (auto intro: open_Pi'I)
   1.359 +  moreover have "x \<in> ?A" using `x i \<in> B` by auto
   1.360 +  moreover have "(\<forall>y\<in>s. y \<in> ?A \<longrightarrow> y i \<in> B)"
   1.361 +  proof (cases, safe)
   1.362 +    fix y assume "y \<in> s"
   1.363 +    assume "i \<notin> domain x" hence "undefined \<in> B" using `x i \<in> B`
   1.364 +      by simp
   1.365 +    moreover
   1.366 +    assume "y \<in> ?A" hence "domain y = domain x" by (simp add: Pi'_def)
   1.367 +    hence "y i = undefined" using `i \<notin> domain x` by simp
   1.368 +    ultimately
   1.369 +    show "y i \<in> B" by simp
   1.370 +  qed force
   1.371 +  ultimately
   1.372 +  show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> y i \<in> B)" by blast
   1.373 +qed
   1.374 +
   1.375 +subsection {* Complete Space of Finite Maps *}
   1.376 +
   1.377 +lemma tendsto_dist_zero:
   1.378 +  assumes "(\<lambda>i. dist (f i) g) ----> 0"
   1.379 +  shows "f ----> g"
   1.380 +  using assms by (auto simp: tendsto_iff dist_real_def)
   1.381 +
   1.382 +lemma tendsto_dist_zero':
   1.383 +  assumes "(\<lambda>i. dist (f i) g) ----> x"
   1.384 +  assumes "0 = x"
   1.385 +  shows "f ----> g"
   1.386 +  using assms tendsto_dist_zero by simp
   1.387 +
   1.388 +lemma tendsto_finmap:
   1.389 +  fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^isub>F ('a::metric_space))"
   1.390 +  assumes ind_f:  "\<And>n. domain (f n) = domain g"
   1.391 +  assumes proj_g:  "\<And>i. i \<in> domain g \<Longrightarrow> (\<lambda>n. (f n) i) ----> g i"
   1.392 +  shows "f ----> g"
   1.393 +  apply (rule tendsto_dist_zero')
   1.394 +  unfolding dist_finmap_def assms
   1.395 +  apply (rule tendsto_intros proj_g | simp)+
   1.396 +  done
   1.397 +
   1.398 +instance finmap :: (type, complete_space) complete_space
   1.399 +proof
   1.400 +  fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^isub>F 'b"
   1.401 +  assume "Cauchy P"
   1.402 +  then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
   1.403 +    by (force simp: cauchy)
   1.404 +  def d \<equiv> "domain (P Nd)"
   1.405 +  with Nd have dim: "\<And>n. n \<ge> Nd \<Longrightarrow> domain (P n) = d" using dist_le_1_imp_domain_eq by auto
   1.406 +  have [simp]: "finite d" unfolding d_def by simp
   1.407 +  def p \<equiv> "\<lambda>i n. (P n) i"
   1.408 +  def q \<equiv> "\<lambda>i. lim (p i)"
   1.409 +  def Q \<equiv> "finmap_of d q"
   1.410 +  have q: "\<And>i. i \<in> d \<Longrightarrow> q i = Q i" by (auto simp add: Q_def Abs_finmap_inverse)
   1.411 +  {
   1.412 +    fix i assume "i \<in> d"
   1.413 +    have "Cauchy (p i)" unfolding cauchy p_def
   1.414 +    proof safe
   1.415 +      fix e::real assume "0 < e"
   1.416 +      with `Cauchy P` obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
   1.417 +        by (force simp: cauchy min_def)
   1.418 +      hence "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto
   1.419 +      with dim have dim: "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = d" by (metis nat_le_linear)
   1.420 +      show "\<exists>N. \<forall>n\<ge>N. dist ((P n) i) ((P N) i) < e"
   1.421 +      proof (safe intro!: exI[where x="N"])
   1.422 +        fix n assume "N \<le> n" have "N \<le> N" by simp
   1.423 +        have "dist ((P n) i) ((P N) i) \<le> dist (P n) (P N)"
   1.424 +          using dim[OF `N \<le> n`]  dim[OF `N \<le> N`] `i \<in> d`
   1.425 +          by (auto intro!: dist_proj)
   1.426 +        also have "\<dots> < e" using N[OF `N \<le> n`] by simp
   1.427 +        finally show "dist ((P n) i) ((P N) i) < e" .
   1.428 +      qed
   1.429 +    qed
   1.430 +    hence "convergent (p i)" by (metis Cauchy_convergent_iff)
   1.431 +    hence "p i ----> q i" unfolding q_def convergent_def by (metis limI)
   1.432 +  } note p = this
   1.433 +  have "P ----> Q"
   1.434 +  proof (rule metric_LIMSEQ_I)
   1.435 +    fix e::real assume "0 < e"
   1.436 +    def e' \<equiv> "min 1 (e / (card d + 1))"
   1.437 +    hence "0 < e'" using `0 < e` by (auto simp: e'_def intro: divide_pos_pos)
   1.438 +    have "\<exists>ni. \<forall>i\<in>d. \<forall>n\<ge>ni i. dist (p i n) (q i) < e'"
   1.439 +    proof (safe intro!: bchoice)
   1.440 +      fix i assume "i \<in> d"
   1.441 +      from p[OF `i \<in> d`, THEN metric_LIMSEQ_D, OF `0 < e'`]
   1.442 +      show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e'" .
   1.443 +    qed then guess ni .. note ni = this
   1.444 +    def N \<equiv> "max Nd (Max (ni ` d))"
   1.445 +    show "\<exists>N. \<forall>n\<ge>N. dist (P n) Q < e"
   1.446 +    proof (safe intro!: exI[where x="N"])
   1.447 +      fix n assume "N \<le> n"
   1.448 +      hence "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q"
   1.449 +        using dim by (simp_all add: N_def Q_def dim_def Abs_finmap_inverse)
   1.450 +      hence "dist (P n) Q = (\<Sum>i\<in>d. dist ((P n) i) (Q i))" by (simp add: dist_finmap_def)
   1.451 +      also have "\<dots> \<le> (\<Sum>i\<in>d. e')"
   1.452 +      proof (intro setsum_mono less_imp_le)
   1.453 +        fix i assume "i \<in> d"
   1.454 +        hence "ni i \<le> Max (ni ` d)" by simp
   1.455 +        also have "\<dots> \<le> N" by (simp add: N_def)
   1.456 +        also have "\<dots> \<le> n" using `N \<le> n` .
   1.457 +        finally
   1.458 +        show "dist ((P n) i) (Q i) < e'"
   1.459 +          using ni `i \<in> d` by (auto simp: p_def q N_def)
   1.460 +      qed
   1.461 +      also have "\<dots> = card d * e'" by (simp add: real_eq_of_nat)
   1.462 +      also have "\<dots> < e" using `0 < e` by (simp add: e'_def field_simps min_def)
   1.463 +      finally show "dist (P n) Q < e" .
   1.464 +    qed
   1.465 +  qed
   1.466 +  thus "convergent P" by (auto simp: convergent_def)
   1.467 +qed
   1.468 +
   1.469 +subsection {* Polish Space of Finite Maps *}
   1.470 +
   1.471 +instantiation finmap :: (countable, polish_space) polish_space
   1.472 +begin
   1.473 +
   1.474 +definition enum_basis_finmap :: "nat \<Rightarrow> ('a \<Rightarrow>\<^isub>F 'b) set" where
   1.475 +  "enum_basis_finmap n =
   1.476 +  (let m = from_nat n::('a \<Rightarrow>\<^isub>F nat) in Pi' (domain m) (enum_basis o (m)\<^isub>F))"
   1.477 +
   1.478 +lemma range_enum_basis_eq:
   1.479 +  "range enum_basis_finmap = {Pi' I S|I S. finite I \<and> (\<forall>i \<in> I. S i \<in> range enum_basis)}"
   1.480 +proof (auto simp: enum_basis_finmap_def[abs_def])
   1.481 +  fix S::"('a \<Rightarrow> 'b set)" and I
   1.482 +  assume "\<forall>i\<in>I. S i \<in> range enum_basis"
   1.483 +  hence "\<forall>i\<in>I. \<exists>n. S i = enum_basis n" by auto
   1.484 +  then obtain n where n: "\<forall>i\<in>I. S i = enum_basis (n i)"
   1.485 +    unfolding bchoice_iff by blast
   1.486 +  assume [simp]: "finite I"
   1.487 +  have "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. n i = (fm i))"
   1.488 +    by (rule finmap_choice) auto
   1.489 +  then obtain m where "Pi' I S = Pi' (domain m) (enum_basis o m)"
   1.490 +    using n by (auto simp: Pi'_def)
   1.491 +  hence "Pi' I S = (let m = from_nat (to_nat m) in Pi' (domain m) (enum_basis \<circ> m))"
   1.492 +    by simp
   1.493 +  thus "Pi' I S \<in> range (\<lambda>n. let m = from_nat n in Pi' (domain m) (enum_basis \<circ> m))"
   1.494 +    by blast
   1.495 +qed (metis finite_domain o_apply rangeI)
   1.496 +
   1.497 +lemma in_enum_basis_finmapI:
   1.498 +  assumes "finite I" assumes "\<And>i. i \<in> I \<Longrightarrow> S i \<in> range enum_basis"
   1.499 +  shows "Pi' I S \<in> range enum_basis_finmap"
   1.500 +  using assms unfolding range_enum_basis_eq by auto
   1.501 +
   1.502 +lemma finmap_topological_basis:
   1.503 +  "topological_basis (range (enum_basis_finmap))"
   1.504 +proof (subst topological_basis_iff, safe)
   1.505 +  fix n::nat
   1.506 +  show "open (enum_basis_finmap n::('a \<Rightarrow>\<^isub>F 'b) set)" using enumerable_basis
   1.507 +    by (auto intro!: open_Pi'I simp: topological_basis_def enum_basis_finmap_def Let_def)
   1.508 +next
   1.509 +  fix O'::"('a \<Rightarrow>\<^isub>F 'b) set" and x
   1.510 +  assume "open O'" "x \<in> O'"
   1.511 +  then obtain e where e: "e > 0" "\<And>y. dist y x < e \<Longrightarrow> y \<in> O'"  unfolding open_dist by blast
   1.512 +  def e' \<equiv> "e / (card (domain x) + 1)"
   1.513 +
   1.514 +  have "\<exists>B.
   1.515 +    (\<forall>i\<in>domain x. x i \<in> enum_basis (B i) \<and> enum_basis (B i) \<subseteq> ball (x i) e')"
   1.516 +  proof (rule bchoice, safe)
   1.517 +    fix i assume "i \<in> domain x"
   1.518 +    have "open (ball (x i) e')" "x i \<in> ball (x i) e'" using e
   1.519 +      by (auto simp add: e'_def intro!: divide_pos_pos)
   1.520 +    from enumerable_basisE[OF this] guess b' .
   1.521 +    thus "\<exists>y. x i \<in> enum_basis y \<and>
   1.522 +            enum_basis y \<subseteq> ball (x i) e'" by auto
   1.523 +  qed
   1.524 +  then guess B .. note B = this
   1.525 +  def B' \<equiv> "Pi' (domain x) (\<lambda>i. enum_basis (B i)::'b set)"
   1.526 +  hence "B' \<in> range enum_basis_finmap" unfolding B'_def
   1.527 +    by (intro in_enum_basis_finmapI) auto
   1.528 +  moreover have "x \<in> B'" unfolding B'_def using B by auto
   1.529 +  moreover have "B' \<subseteq> O'"
   1.530 +  proof
   1.531 +    fix y assume "y \<in> B'" with B have "domain y = domain x" unfolding B'_def
   1.532 +      by (simp add: Pi'_def)
   1.533 +    show "y \<in> O'"
   1.534 +    proof (rule e)
   1.535 +      have "dist y x = (\<Sum>i \<in> domain x. dist (y i) (x i))"
   1.536 +        using `domain y = domain x` by (simp add: dist_finmap_def)
   1.537 +      also have "\<dots> \<le> (\<Sum>i \<in> domain x. e')"
   1.538 +      proof (rule setsum_mono)
   1.539 +        fix i assume "i \<in> domain x"
   1.540 +        with `y \<in> B'` B have "y i \<in> enum_basis (B i)"
   1.541 +          by (simp add: Pi'_def B'_def)
   1.542 +        hence "y i \<in> ball (x i) e'" using B `domain y = domain x` `i \<in> domain x`
   1.543 +          by force
   1.544 +        thus "dist (y i) (x i) \<le> e'" by (simp add: dist_commute)
   1.545 +      qed
   1.546 +      also have "\<dots> = card (domain x) * e'" by (simp add: real_eq_of_nat)
   1.547 +      also have "\<dots> < e" using e by (simp add: e'_def field_simps)
   1.548 +      finally show "dist y x < e" .
   1.549 +    qed
   1.550 +  qed
   1.551 +  ultimately
   1.552 +  show "\<exists>B'\<in>range enum_basis_finmap. x \<in> B' \<and> B' \<subseteq> O'" by blast
   1.553 +qed
   1.554 +
   1.555 +lemma range_enum_basis_finmap_imp_open:
   1.556 +  assumes "x \<in> range enum_basis_finmap"
   1.557 +  shows "open x"
   1.558 +  using finmap_topological_basis assms by (auto simp: topological_basis_def)
   1.559 +
   1.560 +lemma
   1.561 +  open_imp_ex_UNION_of_enum:
   1.562 +  fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
   1.563 +  assumes "open X" assumes "X \<noteq> {}"
   1.564 +  shows "\<exists>A::nat\<Rightarrow>'a set. \<exists>B::nat\<Rightarrow>('a \<Rightarrow> 'b set) . X = UNION UNIV (\<lambda>i. Pi' (A i) (B i)) \<and>
   1.565 +    (\<forall>n. \<forall>i\<in>A n. (B n) i \<in> range enum_basis) \<and> (\<forall>n. finite (A n))"
   1.566 +proof -
   1.567 +  from `open X` obtain B' where B': "B'\<subseteq>range enum_basis_finmap" "\<Union>B' = X"
   1.568 +    using finmap_topological_basis by (force simp add: topological_basis_def)
   1.569 +  then obtain B where B: "B' = enum_basis_finmap ` B" by (auto simp: subset_image_iff)
   1.570 +  show ?thesis
   1.571 +  proof cases
   1.572 +    assume "B = {}" with B have "B' = {}" by simp hence False using B' assms by simp
   1.573 +    thus ?thesis by simp
   1.574 +  next
   1.575 +    assume "B \<noteq> {}" then obtain b where b: "b \<in> B" by auto
   1.576 +    def NA \<equiv> "\<lambda>n::nat. if n \<in> B
   1.577 +      then domain ((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n)
   1.578 +      else domain ((from_nat::_\<Rightarrow>'a\<Rightarrow>\<^isub>F nat) b)"
   1.579 +    def NB \<equiv> "\<lambda>n::nat. if n \<in> B
   1.580 +      then (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n) i))
   1.581 +      else (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) b) i))"
   1.582 +    have "X = UNION UNIV (\<lambda>i. Pi' (NA i) (NB i))" unfolding B'(2)[symmetric] using b
   1.583 +      unfolding B
   1.584 +      by safe
   1.585 +         (auto simp add: NA_def NB_def enum_basis_finmap_def Let_def o_def split: split_if_asm)
   1.586 +    moreover
   1.587 +    have "(\<forall>n. \<forall>i\<in>NA n. (NB n) i \<in> range enum_basis)"
   1.588 +      using enumerable_basis by (auto simp: topological_basis_def NA_def NB_def)
   1.589 +    moreover have "(\<forall>n. finite (NA n))" by (simp add: NA_def)
   1.590 +    ultimately show ?thesis by auto
   1.591 +  qed
   1.592 +qed
   1.593 +
   1.594 +lemma
   1.595 +  open_imp_ex_UNION:
   1.596 +  fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
   1.597 +  assumes "open X" assumes "X \<noteq> {}"
   1.598 +  shows "\<exists>A::nat\<Rightarrow>'a set. \<exists>B::nat\<Rightarrow>('a \<Rightarrow> 'b set) . X = UNION UNIV (\<lambda>i. Pi' (A i) (B i)) \<and>
   1.599 +    (\<forall>n. \<forall>i\<in>A n. open ((B n) i)) \<and> (\<forall>n. finite (A n))"
   1.600 +  using open_imp_ex_UNION_of_enum[OF assms]
   1.601 +  apply auto
   1.602 +  apply (rule_tac x = A in exI)
   1.603 +  apply (rule_tac x = B in exI)
   1.604 +  apply (auto simp: open_enum_basis)
   1.605 +  done
   1.606 +
   1.607 +lemma
   1.608 +  open_basisE:
   1.609 +  assumes "open X" assumes "X \<noteq> {}"
   1.610 +  obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
   1.611 +  "X = UNION UNIV (\<lambda>i. Pi' (A i) (B i))" "\<And>n i. i\<in>A n \<Longrightarrow> open ((B n) i)" "\<And>n. finite (A n)"
   1.612 +using open_imp_ex_UNION[OF assms] by auto
   1.613 +
   1.614 +lemma
   1.615 +  open_basis_of_enumE:
   1.616 +  assumes "open X" assumes "X \<noteq> {}"
   1.617 +  obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
   1.618 +  "X = UNION UNIV (\<lambda>i. Pi' (A i) (B i))" "\<And>n i. i\<in>A n \<Longrightarrow> (B n) i \<in> range enum_basis"
   1.619 +  "\<And>n. finite (A n)"
   1.620 +using open_imp_ex_UNION_of_enum[OF assms] by auto
   1.621 +
   1.622 +instance proof qed (blast intro: finmap_topological_basis)
   1.623 +
   1.624 +end
   1.625 +
   1.626 +subsection {* Product Measurable Space of Finite Maps *}
   1.627 +
   1.628 +definition "PiF I M \<equiv>
   1.629 +  sigma
   1.630 +    (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
   1.631 +    {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
   1.632 +
   1.633 +abbreviation
   1.634 +  "Pi\<^isub>F I M \<equiv> PiF I M"
   1.635 +
   1.636 +syntax
   1.637 +  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3PIF _:_./ _)" 10)
   1.638 +
   1.639 +syntax (xsymbols)
   1.640 +  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>F _\<in>_./ _)"  10)
   1.641 +
   1.642 +syntax (HTML output)
   1.643 +  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>F _\<in>_./ _)"  10)
   1.644 +
   1.645 +translations
   1.646 +  "PIF x:I. M" == "CONST PiF I (%x. M)"
   1.647 +
   1.648 +lemma PiF_gen_subset: "{(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))} \<subseteq>
   1.649 +    Pow (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
   1.650 +  by (auto simp: Pi'_def) (blast dest: sets_into_space)
   1.651 +
   1.652 +lemma space_PiF: "space (PiF I M) = (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
   1.653 +  unfolding PiF_def using PiF_gen_subset by (rule space_measure_of)
   1.654 +
   1.655 +lemma sets_PiF:
   1.656 +  "sets (PiF I M) = sigma_sets (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
   1.657 +    {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
   1.658 +  unfolding PiF_def using PiF_gen_subset by (rule sets_measure_of)
   1.659 +
   1.660 +lemma sets_PiF_singleton:
   1.661 +  "sets (PiF {I} M) = sigma_sets (\<Pi>' j\<in>I. space (M j))
   1.662 +    {(\<Pi>' j\<in>I. X j) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
   1.663 +  unfolding sets_PiF by simp
   1.664 +
   1.665 +lemma in_sets_PiFI:
   1.666 +  assumes "X = (Pi' J S)" "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
   1.667 +  shows "X \<in> sets (PiF I M)"
   1.668 +  unfolding sets_PiF
   1.669 +  using assms by blast
   1.670 +
   1.671 +lemma product_in_sets_PiFI:
   1.672 +  assumes "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
   1.673 +  shows "(Pi' J S) \<in> sets (PiF I M)"
   1.674 +  unfolding sets_PiF
   1.675 +  using assms by blast
   1.676 +
   1.677 +lemma singleton_space_subset_in_sets:
   1.678 +  fixes J
   1.679 +  assumes "J \<in> I"
   1.680 +  assumes "finite J"
   1.681 +  shows "space (PiF {J} M) \<in> sets (PiF I M)"
   1.682 +  using assms
   1.683 +  by (intro in_sets_PiFI[where J=J and S="\<lambda>i. space (M i)"])
   1.684 +      (auto simp: product_def space_PiF)
   1.685 +
   1.686 +lemma singleton_subspace_set_in_sets:
   1.687 +  assumes A: "A \<in> sets (PiF {J} M)"
   1.688 +  assumes "finite J"
   1.689 +  assumes "J \<in> I"
   1.690 +  shows "A \<in> sets (PiF I M)"
   1.691 +  using A[unfolded sets_PiF]
   1.692 +  apply (induct A)
   1.693 +  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
   1.694 +  using assms
   1.695 +  by (auto intro: in_sets_PiFI intro!: singleton_space_subset_in_sets)
   1.696 +
   1.697 +lemma
   1.698 +  finite_measurable_singletonI:
   1.699 +  assumes "finite I"
   1.700 +  assumes "\<And>J. J \<in> I \<Longrightarrow> finite J"
   1.701 +  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
   1.702 +  shows "A \<in> measurable (PiF I M) N"
   1.703 +  unfolding measurable_def
   1.704 +proof safe
   1.705 +  fix y assume "y \<in> sets N"
   1.706 +  have "A -` y \<inter> space (PiF I M) = (\<Union>J\<in>I. A -` y \<inter> space (PiF {J} M))"
   1.707 +    by (auto simp: space_PiF)
   1.708 +  also have "\<dots> \<in> sets (PiF I M)"
   1.709 +  proof
   1.710 +    show "finite I" by fact
   1.711 +    fix J assume "J \<in> I"
   1.712 +    with assms have "finite J" by simp
   1.713 +    show "A -` y \<inter> space (PiF {J} M) \<in> sets (PiF I M)"
   1.714 +      by (rule singleton_subspace_set_in_sets[OF measurable_sets[OF assms(3)]]) fact+
   1.715 +  qed
   1.716 +  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
   1.717 +next
   1.718 +  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
   1.719 +    using MN[of "domain x"]
   1.720 +    by (auto simp: space_PiF measurable_space Pi'_def)
   1.721 +qed
   1.722 +
   1.723 +lemma
   1.724 +  countable_finite_comprehension:
   1.725 +  fixes f :: "'a::countable set \<Rightarrow> _"
   1.726 +  assumes "\<And>s. P s \<Longrightarrow> finite s"
   1.727 +  assumes "\<And>s. P s \<Longrightarrow> f s \<in> sets M"
   1.728 +  shows "\<Union>{f s|s. P s} \<in> sets M"
   1.729 +proof -
   1.730 +  have "\<Union>{f s|s. P s} = (\<Union>n::nat. let s = set (from_nat n) in if P s then f s else {})"
   1.731 +  proof safe
   1.732 +    fix x X s assume "x \<in> f s" "P s"
   1.733 +    moreover with assms obtain l where "s = set l" using finite_list by blast
   1.734 +    ultimately show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using `P s`
   1.735 +      by (auto intro!: exI[where x="to_nat l"])
   1.736 +  next
   1.737 +    fix x n assume "x \<in> (let s = set (from_nat n) in if P s then f s else {})"
   1.738 +    thus "x \<in> \<Union>{f s|s. P s}" using assms by (auto simp: Let_def split: split_if_asm)
   1.739 +  qed
   1.740 +  hence "\<Union>{f s|s. P s} = (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" by simp
   1.741 +  also have "\<dots> \<in> sets M" using assms by (auto simp: Let_def)
   1.742 +  finally show ?thesis .
   1.743 +qed
   1.744 +
   1.745 +lemma space_subset_in_sets:
   1.746 +  fixes J::"'a::countable set set"
   1.747 +  assumes "J \<subseteq> I"
   1.748 +  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
   1.749 +  shows "space (PiF J M) \<in> sets (PiF I M)"
   1.750 +proof -
   1.751 +  have "space (PiF J M) = \<Union>{space (PiF {j} M)|j. j \<in> J}"
   1.752 +    unfolding space_PiF by blast
   1.753 +  also have "\<dots> \<in> sets (PiF I M)" using assms
   1.754 +    by (intro countable_finite_comprehension) (auto simp: singleton_space_subset_in_sets)
   1.755 +  finally show ?thesis .
   1.756 +qed
   1.757 +
   1.758 +lemma subspace_set_in_sets:
   1.759 +  fixes J::"'a::countable set set"
   1.760 +  assumes A: "A \<in> sets (PiF J M)"
   1.761 +  assumes "J \<subseteq> I"
   1.762 +  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
   1.763 +  shows "A \<in> sets (PiF I M)"
   1.764 +  using A[unfolded sets_PiF]
   1.765 +  apply (induct A)
   1.766 +  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
   1.767 +  using assms
   1.768 +  by (auto intro: in_sets_PiFI intro!: space_subset_in_sets)
   1.769 +
   1.770 +lemma
   1.771 +  countable_measurable_PiFI:
   1.772 +  fixes I::"'a::countable set set"
   1.773 +  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> finite J \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
   1.774 +  shows "A \<in> measurable (PiF I M) N"
   1.775 +  unfolding measurable_def
   1.776 +proof safe
   1.777 +  fix y assume "y \<in> sets N"
   1.778 +  have "A -` y = (\<Union>{A -` y \<inter> {x. domain x = J}|J. finite J})" by auto
   1.779 +  hence "A -` y \<inter> space (PiF I M) = (\<Union>n. A -` y \<inter> space (PiF ({set (from_nat n)}\<inter>I) M))"
   1.780 +    apply (auto simp: space_PiF Pi'_def)
   1.781 +  proof -
   1.782 +    case goal1
   1.783 +    from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto
   1.784 +    thus ?case
   1.785 +      apply (intro exI[where x="to_nat xs"])
   1.786 +      apply auto
   1.787 +      done
   1.788 +  qed
   1.789 +  also have "\<dots> \<in> sets (PiF I M)"
   1.790 +    apply (intro Int countable_nat_UN subsetI, safe)
   1.791 +    apply (case_tac "set (from_nat i) \<in> I")
   1.792 +    apply simp_all
   1.793 +    apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]])
   1.794 +    using assms `y \<in> sets N`
   1.795 +    apply (auto simp: space_PiF)
   1.796 +    done
   1.797 +  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
   1.798 +next
   1.799 +  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
   1.800 +    using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def)
   1.801 +qed
   1.802 +
   1.803 +lemma measurable_PiF:
   1.804 +  assumes f: "\<And>x. x \<in> space N \<Longrightarrow> domain (f x) \<in> I \<and> (\<forall>i\<in>domain (f x). (f x) i \<in> space (M i))"
   1.805 +  assumes S: "\<And>J S. J \<in> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> S i \<in> sets (M i)) \<Longrightarrow>
   1.806 +    f -` (Pi' J S) \<inter> space N \<in> sets N"
   1.807 +  shows "f \<in> measurable N (PiF I M)"
   1.808 +  unfolding PiF_def
   1.809 +  using PiF_gen_subset
   1.810 +  apply (rule measurable_measure_of)
   1.811 +  using f apply force
   1.812 +  apply (insert S, auto)
   1.813 +  done
   1.814 +
   1.815 +lemma
   1.816 +  restrict_sets_measurable:
   1.817 +  assumes A: "A \<in> sets (PiF I M)" and "J \<subseteq> I"
   1.818 +  shows "A \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
   1.819 +  using A[unfolded sets_PiF]
   1.820 +  apply (induct A)
   1.821 +  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
   1.822 +proof -
   1.823 +  fix a assume "a \<in> {Pi' J X |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
   1.824 +  then obtain K S where S: "a = Pi' K S" "K \<in> I" "(\<forall>i\<in>K. S i \<in> sets (M i))"
   1.825 +    by auto
   1.826 +  show "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
   1.827 +  proof cases
   1.828 +    assume "K \<in> J"
   1.829 +    hence "a \<inter> {m. domain m \<in> J} \<in> {Pi' K X |X K. K \<in> J \<and> X \<in> (\<Pi> j\<in>K. sets (M j))}" using S
   1.830 +      by (auto intro!: exI[where x=K] exI[where x=S] simp: Pi'_def)
   1.831 +    also have "\<dots> \<subseteq> sets (PiF J M)" unfolding sets_PiF by auto
   1.832 +    finally show ?thesis .
   1.833 +  next
   1.834 +    assume "K \<notin> J"
   1.835 +    hence "a \<inter> {m. domain m \<in> J} = {}" using S by (auto simp: Pi'_def)
   1.836 +    also have "\<dots> \<in> sets (PiF J M)" by simp
   1.837 +    finally show ?thesis .
   1.838 +  qed
   1.839 +next
   1.840 +  show "{} \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" by simp
   1.841 +next
   1.842 +  fix a :: "nat \<Rightarrow> _"
   1.843 +  assume a: "(\<And>i. a i \<inter> {m. domain m \<in> J} \<in> sets (PiF J M))"
   1.844 +  have "UNION UNIV a \<inter> {m. domain m \<in> J} = (\<Union>i. (a i \<inter> {m. domain m \<in> J}))"
   1.845 +    by simp
   1.846 +  also have "\<dots> \<in> sets (PiF J M)" using a by (intro countable_nat_UN) auto
   1.847 +  finally show "UNION UNIV a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
   1.848 +next
   1.849 +  fix a assume a: "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
   1.850 +  have "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} = (space (PiF J M) - (a \<inter> {m. domain m \<in> J}))"
   1.851 +    using `J \<subseteq> I` by (auto simp: space_PiF Pi'_def)
   1.852 +  also have "\<dots> \<in> sets (PiF J M)" using a by auto
   1.853 +  finally show "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
   1.854 +qed
   1.855 +
   1.856 +lemma measurable_finmap_of:
   1.857 +  assumes f: "\<And>i. (\<exists>x \<in> space N. i \<in> J x) \<Longrightarrow> (\<lambda>x. f x i) \<in> measurable N (M i)"
   1.858 +  assumes J: "\<And>x. x \<in> space N \<Longrightarrow> J x \<in> I" "\<And>x. x \<in> space N \<Longrightarrow> finite (J x)"
   1.859 +  assumes JN: "\<And>S. {x. J x = S} \<inter> space N \<in> sets N"
   1.860 +  shows "(\<lambda>x. finmap_of (J x) (f x)) \<in> measurable N (PiF I M)"
   1.861 +proof (rule measurable_PiF)
   1.862 +  fix x assume "x \<in> space N"
   1.863 +  with J[of x] measurable_space[OF f]
   1.864 +  show "domain (finmap_of (J x) (f x)) \<in> I \<and>
   1.865 +        (\<forall>i\<in>domain (finmap_of (J x) (f x)). (finmap_of (J x) (f x)) i \<in> space (M i))"
   1.866 +    by auto
   1.867 +next
   1.868 +  fix K S assume "K \<in> I" and *: "\<And>i. i \<in> K \<Longrightarrow> S i \<in> sets (M i)"
   1.869 +  with J have eq: "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N =
   1.870 +    (if \<exists>x \<in> space N. K = J x \<and> finite K then if K = {} then {x \<in> space N. J x = K}
   1.871 +      else (\<Inter>i\<in>K. (\<lambda>x. f x i) -` S i \<inter> {x \<in> space N. J x = K}) else {})"
   1.872 +    by (auto simp: Pi'_def)
   1.873 +  have r: "{x \<in> space N. J x = K} = space N \<inter> ({x. J x = K} \<inter> space N)" by auto
   1.874 +  show "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N \<in> sets N"
   1.875 +    unfolding eq r
   1.876 +    apply (simp del: INT_simps add: )
   1.877 +    apply (intro conjI impI finite_INT JN Int[OF top])
   1.878 +    apply simp apply assumption
   1.879 +    apply (subst Int_assoc[symmetric])
   1.880 +    apply (rule Int)
   1.881 +    apply (intro measurable_sets[OF f] *) apply force apply assumption
   1.882 +    apply (intro JN)
   1.883 +    done
   1.884 +qed
   1.885 +
   1.886 +lemma measurable_PiM_finmap_of:
   1.887 +  assumes "finite J"
   1.888 +  shows "finmap_of J \<in> measurable (Pi\<^isub>M J M) (PiF {J} M)"
   1.889 +  apply (rule measurable_finmap_of)
   1.890 +  apply (rule measurable_component_singleton)
   1.891 +  apply simp
   1.892 +  apply rule
   1.893 +  apply (rule `finite J`)
   1.894 +  apply simp
   1.895 +  done
   1.896 +
   1.897 +lemma proj_measurable_singleton:
   1.898 +  assumes "A \<in> sets (M i)" "finite I"
   1.899 +  shows "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
   1.900 +proof cases
   1.901 +  assume "i \<in> I"
   1.902 +  hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
   1.903 +    Pi' I (\<lambda>x. if x = i then A else space (M x))"
   1.904 +    using sets_into_space[OF ] `A \<in> sets (M i)` assms
   1.905 +    by (auto simp: space_PiF Pi'_def)
   1.906 +  thus ?thesis  using assms `A \<in> sets (M i)`
   1.907 +    by (intro in_sets_PiFI) auto
   1.908 +next
   1.909 +  assume "i \<notin> I"
   1.910 +  hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
   1.911 +    (if undefined \<in> A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def)
   1.912 +  thus ?thesis by simp
   1.913 +qed
   1.914 +
   1.915 +lemma measurable_proj_singleton:
   1.916 +  fixes I
   1.917 +  assumes "finite I" "i \<in> I"
   1.918 +  shows "(\<lambda>x. (x)\<^isub>F i) \<in> measurable (PiF {I} M) (M i)"
   1.919 +proof (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
   1.920 +qed (insert `i \<in> I`, auto simp: space_PiF)
   1.921 +
   1.922 +lemma measurable_proj_countable:
   1.923 +  fixes I::"'a::countable set set"
   1.924 +  assumes "y \<in> space (M i)"
   1.925 +  shows "(\<lambda>x. if i \<in> domain x then (x)\<^isub>F i else y) \<in> measurable (PiF I M) (M i)"
   1.926 +proof (rule countable_measurable_PiFI)
   1.927 +  fix J assume "J \<in> I" "finite J"
   1.928 +  show "(\<lambda>x. if i \<in> domain x then x i else y) \<in> measurable (PiF {J} M) (M i)"
   1.929 +    unfolding measurable_def
   1.930 +  proof safe
   1.931 +    fix z assume "z \<in> sets (M i)"
   1.932 +    have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
   1.933 +      (\<lambda>x. if i \<in> J then (x)\<^isub>F i else y) -` z \<inter> space (PiF {J} M)"
   1.934 +      by (auto simp: space_PiF Pi'_def)
   1.935 +    also have "\<dots> \<in> sets (PiF {J} M)" using `z \<in> sets (M i)` `finite J`
   1.936 +      by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
   1.937 +    finally show "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) \<in>
   1.938 +      sets (PiF {J} M)" .
   1.939 +  qed (insert `y \<in> space (M i)`, auto simp: space_PiF Pi'_def)
   1.940 +qed
   1.941 +
   1.942 +lemma measurable_restrict_proj:
   1.943 +  assumes "J \<in> II" "finite J"
   1.944 +  shows "finmap_of J \<in> measurable (PiM J M) (PiF II M)"
   1.945 +  using assms
   1.946 +  by (intro measurable_finmap_of measurable_component_singleton) auto
   1.947 +
   1.948 +lemma
   1.949 +  measurable_proj_PiM:
   1.950 +  fixes J K ::"'a::countable set" and I::"'a set set"
   1.951 +  assumes "finite J" "J \<in> I"
   1.952 +  assumes "x \<in> space (PiM J M)"
   1.953 +  shows "proj \<in>
   1.954 +    measurable (PiF {J} M) (PiM J M)"
   1.955 +proof (rule measurable_PiM_single)
   1.956 +  show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^isub>E i \<in> J. space (M i))"
   1.957 +    using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def)
   1.958 +next
   1.959 +  fix A i assume A: "i \<in> J" "A \<in> sets (M i)"
   1.960 +  show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} \<in> sets (PiF {J} M)"
   1.961 +  proof
   1.962 +    have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} =
   1.963 +      (\<lambda>\<omega>. (\<omega>)\<^isub>F i) -` A \<inter> space (PiF {J} M)" by auto
   1.964 +    also have "\<dots> \<in> sets (PiF {J} M)"
   1.965 +      using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM)
   1.966 +    finally show ?thesis .
   1.967 +  qed simp
   1.968 +qed
   1.969 +
   1.970 +lemma sets_subspaceI:
   1.971 +  assumes "A \<inter> space M \<in> sets M"
   1.972 +  assumes "B \<in> sets M"
   1.973 +  shows "A \<inter> B \<in> sets M" using assms
   1.974 +proof -
   1.975 +  have "A \<inter> B = (A \<inter> space M) \<inter> B"
   1.976 +    using assms sets_into_space by auto
   1.977 +  thus ?thesis using assms by auto
   1.978 +qed
   1.979 +
   1.980 +lemma space_PiF_singleton_eq_product:
   1.981 +  assumes "finite I"
   1.982 +  shows "space (PiF {I} M) = (\<Pi>' i\<in>I. space (M i))"
   1.983 +  by (auto simp: product_def space_PiF assms)
   1.984 +
   1.985 +text {* adapted from @{thm sets_PiM_single} *}
   1.986 +
   1.987 +lemma sets_PiF_single:
   1.988 +  assumes "finite I" "I \<noteq> {}"
   1.989 +  shows "sets (PiF {I} M) =
   1.990 +    sigma_sets (\<Pi>' i\<in>I. space (M i))
   1.991 +      {{f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
   1.992 +    (is "_ = sigma_sets ?\<Omega> ?R")
   1.993 +  unfolding sets_PiF_singleton
   1.994 +proof (rule sigma_sets_eqI)
   1.995 +  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
   1.996 +  fix A assume "A \<in> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
   1.997 +  then obtain X where X: "A = Pi' I X" "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
   1.998 +  show "A \<in> sigma_sets ?\<Omega> ?R"
   1.999 +  proof -
  1.1000 +    from `I \<noteq> {}` X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
  1.1001 +      using sets_into_space
  1.1002 +      by (auto simp: space_PiF product_def) blast
  1.1003 +    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
  1.1004 +      using X `I \<noteq> {}` assms by (intro R.finite_INT) (auto simp: space_PiF)
  1.1005 +    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
  1.1006 +  qed
  1.1007 +next
  1.1008 +  fix A assume "A \<in> ?R"
  1.1009 +  then obtain i B where A: "A = {f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
  1.1010 +    by auto
  1.1011 +  then have "A = (\<Pi>' j \<in> I. if j = i then B else space (M j))"
  1.1012 +    using sets_into_space[OF A(3)]
  1.1013 +    apply (auto simp: Pi'_iff split: split_if_asm)
  1.1014 +    apply blast
  1.1015 +    done
  1.1016 +  also have "\<dots> \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
  1.1017 +    using A
  1.1018 +    by (intro sigma_sets.Basic )
  1.1019 +       (auto intro: exI[where x="\<lambda>j. if j = i then B else space (M j)"])
  1.1020 +  finally show "A \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" .
  1.1021 +qed
  1.1022 +
  1.1023 +text {* adapted from @{thm PiE_cong} *}
  1.1024 +
  1.1025 +lemma Pi'_cong:
  1.1026 +  assumes "finite I"
  1.1027 +  assumes "\<And>i. i \<in> I \<Longrightarrow> f i = g i"
  1.1028 +  shows "Pi' I f = Pi' I g"
  1.1029 +using assms by (auto simp: Pi'_def)
  1.1030 +
  1.1031 +text {* adapted from @{thm Pi_UN} *}
  1.1032 +
  1.1033 +lemma Pi'_UN:
  1.1034 +  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
  1.1035 +  assumes "finite I"
  1.1036 +  assumes mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
  1.1037 +  shows "(\<Union>n. Pi' I (A n)) = Pi' I (\<lambda>i. \<Union>n. A n i)"
  1.1038 +proof (intro set_eqI iffI)
  1.1039 +  fix f assume "f \<in> Pi' I (\<lambda>i. \<Union>n. A n i)"
  1.1040 +  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: `finite I` Pi'_def)
  1.1041 +  from bchoice[OF this(1)] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
  1.1042 +  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
  1.1043 +    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
  1.1044 +  have "f \<in> Pi' I (\<lambda>i. A k i)"
  1.1045 +  proof
  1.1046 +    fix i assume "i \<in> I"
  1.1047 +    from mono[OF this, of "n i" k] k[OF this] n[OF this] `domain f = I` `i \<in> I`
  1.1048 +    show "f i \<in> A k i " by (auto simp: `finite I`)
  1.1049 +  qed (simp add: `domain f = I` `finite I`)
  1.1050 +  then show "f \<in> (\<Union>n. Pi' I (A n))" by auto
  1.1051 +qed (auto simp: Pi'_def `finite I`)
  1.1052 +
  1.1053 +text {* adapted from @{thm sigma_prod_algebra_sigma_eq} *}
  1.1054 +
  1.1055 +lemma sigma_fprod_algebra_sigma_eq:
  1.1056 +  fixes E :: "'i \<Rightarrow> 'a set set"
  1.1057 +  assumes [simp]: "finite I" "I \<noteq> {}"
  1.1058 +  assumes S_mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
  1.1059 +    and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
  1.1060 +    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
  1.1061 +  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
  1.1062 +    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
  1.1063 +  defines "P == { Pi' I F | F. \<forall>i\<in>I. F i \<in> E i }"
  1.1064 +  shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P"
  1.1065 +proof
  1.1066 +  let ?P = "sigma (space (Pi\<^isub>F {I} M)) P"
  1.1067 +  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>F {I} M))"
  1.1068 +    using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq)
  1.1069 +  then have space_P: "space ?P = (\<Pi>' i\<in>I. space (M i))"
  1.1070 +    by (simp add: space_PiF)
  1.1071 +  have "sets (PiF {I} M) =
  1.1072 +      sigma_sets (space ?P) {{f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
  1.1073 +    using sets_PiF_single[of I M] by (simp add: space_P)
  1.1074 +  also have "\<dots> \<subseteq> sets (sigma (space (PiF {I} M)) P)"
  1.1075 +  proof (safe intro!: sigma_sets_subset)
  1.1076 +    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
  1.1077 +    have "(\<lambda>x. (x)\<^isub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
  1.1078 +    proof (subst measurable_iff_measure_of)
  1.1079 +      show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
  1.1080 +      from space_P `i \<in> I` show "(\<lambda>x. (x)\<^isub>F i) \<in> space ?P \<rightarrow> space (M i)"
  1.1081 +        by auto
  1.1082 +      show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  1.1083 +      proof
  1.1084 +        fix A assume A: "A \<in> E i"
  1.1085 +        then have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
  1.1086 +          using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
  1.1087 +        also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
  1.1088 +          by (intro Pi'_cong) (simp_all add: S_union)
  1.1089 +        also have "\<dots> = (\<Union>n. \<Pi>' j\<in>I. if i = j then A else S j n)"
  1.1090 +          using S_mono
  1.1091 +          by (subst Pi'_UN[symmetric, OF `finite I`]) (auto simp: incseq_def)
  1.1092 +        also have "\<dots> \<in> sets ?P"
  1.1093 +        proof (safe intro!: countable_UN)
  1.1094 +          fix n show "(\<Pi>' j\<in>I. if i = j then A else S j n) \<in> sets ?P"
  1.1095 +            using A S_in_E
  1.1096 +            by (simp add: P_closed)
  1.1097 +               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j n"])
  1.1098 +        qed
  1.1099 +        finally show "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  1.1100 +          using P_closed by simp
  1.1101 +      qed
  1.1102 +    qed
  1.1103 +    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
  1.1104 +    have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  1.1105 +      by (simp add: E_generates)
  1.1106 +    also have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
  1.1107 +      using P_closed by (auto simp: space_PiF)
  1.1108 +    finally show "\<dots> \<in> sets ?P" .
  1.1109 +  qed
  1.1110 +  finally show "sets (PiF {I} M) \<subseteq> sigma_sets (space (PiF {I} M)) P"
  1.1111 +    by (simp add: P_closed)
  1.1112 +  show "sigma_sets (space (PiF {I} M)) P \<subseteq> sets (PiF {I} M)"
  1.1113 +    using `finite I` `I \<noteq> {}`
  1.1114 +    by (auto intro!: sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def)
  1.1115 +qed
  1.1116 +
  1.1117 +lemma enumerable_sigma_fprod_algebra_sigma_eq:
  1.1118 +  assumes "I \<noteq> {}"
  1.1119 +  assumes [simp]: "finite I"
  1.1120 +  shows "sets (PiF {I} (\<lambda>_. borel)) = sigma_sets (space (PiF {I} (\<lambda>_. borel)))
  1.1121 +    {Pi' I F |F. (\<forall>i\<in>I. F i \<in> range enum_basis)}"
  1.1122 +proof -
  1.1123 +  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  1.1124 +  show ?thesis
  1.1125 +  proof (rule sigma_fprod_algebra_sigma_eq)
  1.1126 +    show "finite I" by simp
  1.1127 +    show "I \<noteq> {}" by fact
  1.1128 +    show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
  1.1129 +      using S by simp_all
  1.1130 +    show "range enum_basis \<subseteq> Pow (space borel)" by simp
  1.1131 +    show "sets borel = sigma_sets (space borel) (range enum_basis)"
  1.1132 +      by (simp add: borel_eq_enum_basis)
  1.1133 +  qed
  1.1134 +qed
  1.1135 +
  1.1136 +text {* adapted from @{thm enumerable_sigma_fprod_algebra_sigma_eq} *}
  1.1137 +
  1.1138 +lemma enumerable_sigma_prod_algebra_sigma_eq:
  1.1139 +  assumes "I \<noteq> {}"
  1.1140 +  assumes [simp]: "finite I"
  1.1141 +  shows "sets (PiM I (\<lambda>_. borel)) = sigma_sets (space (PiM I (\<lambda>_. borel)))
  1.1142 +    {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range enum_basis}"
  1.1143 +proof -
  1.1144 +  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  1.1145 +  show ?thesis
  1.1146 +  proof (rule sigma_prod_algebra_sigma_eq)
  1.1147 +    show "finite I" by simp note[[show_types]]
  1.1148 +    fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
  1.1149 +      using S by simp_all
  1.1150 +    show "range enum_basis \<subseteq> Pow (space borel)" by simp
  1.1151 +    show "sets borel = sigma_sets (space borel) (range enum_basis)"
  1.1152 +      by (simp add: borel_eq_enum_basis)
  1.1153 +  qed
  1.1154 +qed
  1.1155 +
  1.1156 +lemma product_open_generates_sets_PiF_single:
  1.1157 +  assumes "I \<noteq> {}"
  1.1158 +  assumes [simp]: "finite I"
  1.1159 +  shows "sets (PiF {I} (\<lambda>_. borel::'b::enumerable_basis measure)) =
  1.1160 +    sigma_sets (space (PiF {I} (\<lambda>_. borel))) {Pi' I F |F. (\<forall>i\<in>I. F i \<in> Collect open)}"
  1.1161 +proof -
  1.1162 +  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  1.1163 +  show ?thesis
  1.1164 +  proof (rule sigma_fprod_algebra_sigma_eq)
  1.1165 +    show "finite I" by simp
  1.1166 +    show "I \<noteq> {}" by fact
  1.1167 +    show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
  1.1168 +      using S by (auto simp: open_enum_basis)
  1.1169 +    show "Collect open \<subseteq> Pow (space borel)" by simp
  1.1170 +    show "sets borel = sigma_sets (space borel) (Collect open)"
  1.1171 +      by (simp add: borel_def)
  1.1172 +  qed
  1.1173 +qed
  1.1174 +
  1.1175 +lemma product_open_generates_sets_PiM:
  1.1176 +  assumes "I \<noteq> {}"
  1.1177 +  assumes [simp]: "finite I"
  1.1178 +  shows "sets (PiM I (\<lambda>_. borel::'b::enumerable_basis measure)) =
  1.1179 +    sigma_sets (space (PiM I (\<lambda>_. borel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> Collect open}"
  1.1180 +proof -
  1.1181 +  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  1.1182 +  show ?thesis
  1.1183 +  proof (rule sigma_prod_algebra_sigma_eq)
  1.1184 +    show "finite I" by simp note[[show_types]]
  1.1185 +    fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
  1.1186 +      using S by (auto simp: open_enum_basis)
  1.1187 +    show "Collect open \<subseteq> Pow (space borel)" by simp
  1.1188 +    show "sets borel = sigma_sets (space borel) (Collect open)"
  1.1189 +      by (simp add: borel_def)
  1.1190 +  qed
  1.1191 +qed
  1.1192 +
  1.1193 +lemma finmap_UNIV[simp]: "(\<Union>J\<in>Collect finite. J \<leadsto> UNIV) = UNIV" by auto
  1.1194 +
  1.1195 +lemma borel_eq_PiF_borel:
  1.1196 +  shows "(borel :: ('i::countable \<Rightarrow>\<^isub>F 'a::polish_space) measure) =
  1.1197 +  PiF (Collect finite) (\<lambda>_. borel :: 'a measure)"
  1.1198 +proof (rule measure_eqI)
  1.1199 +  have C: "Collect finite \<noteq> {}" by auto
  1.1200 +  show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) = sets (PiF (Collect finite) (\<lambda>_. borel))"
  1.1201 +  proof
  1.1202 +    show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) \<subseteq> sets (PiF (Collect finite) (\<lambda>_. borel))"
  1.1203 +      apply (simp add: borel_def sets_PiF)
  1.1204 +    proof (rule sigma_sets_mono, safe, cases)
  1.1205 +      fix X::"('i \<Rightarrow>\<^isub>F 'a) set" assume "open X" "X \<noteq> {}"
  1.1206 +      from open_basisE[OF this] guess NA NB . note N = this
  1.1207 +      hence "X = (\<Union>i. Pi' (NA i) (NB i))" by simp
  1.1208 +      also have "\<dots> \<in>
  1.1209 +        sigma_sets UNIV {Pi' J S |S J. finite J \<and> S \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
  1.1210 +        using N by (intro Union sigma_sets.Basic) blast
  1.1211 +      finally show "X \<in> sigma_sets UNIV
  1.1212 +        {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}" .
  1.1213 +    qed (auto simp: Empty)
  1.1214 +  next
  1.1215 +    show "sets (PiF (Collect finite) (\<lambda>_. borel)) \<subseteq> sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure)"
  1.1216 +    proof
  1.1217 +      fix x assume x: "x \<in> sets (PiF (Collect finite::'i set set) (\<lambda>_. borel::'a measure))"
  1.1218 +      hence x_sp: "x \<subseteq> space (PiF (Collect finite) (\<lambda>_. borel))" by (rule sets_into_space)
  1.1219 +      let ?x = "\<lambda>J. x \<inter> {x. domain x = J}"
  1.1220 +      have "x = \<Union>{?x J |J. finite J}" by auto
  1.1221 +      also have "\<dots> \<in> sets borel"
  1.1222 +      proof (rule countable_finite_comprehension, assumption)
  1.1223 +        fix J::"'i set" assume "finite J"
  1.1224 +        { assume ef: "J = {}"
  1.1225 +          { assume e: "?x J = {}"
  1.1226 +            hence "?x J \<in> sets borel" by simp
  1.1227 +          } moreover {
  1.1228 +            assume "?x J \<noteq> {}"
  1.1229 +            then obtain f where "f \<in> x" "domain f = {}" using ef by auto
  1.1230 +            hence "?x J = {f}" using `J = {}`
  1.1231 +              by (auto simp: finmap_eq_iff)
  1.1232 +            also have "{f} \<in> sets borel" by simp
  1.1233 +            finally have "?x J \<in> sets borel" .
  1.1234 +          } ultimately have "?x J \<in> sets borel" by blast
  1.1235 +        } moreover {
  1.1236 +          assume "J \<noteq> ({}::'i set)"
  1.1237 +          from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'a set" . note S = this
  1.1238 +          have "(?x J) = x \<inter> {m. domain m \<in> {J}}" by auto
  1.1239 +          also have "\<dots> \<in> sets (PiF {J} (\<lambda>_. borel))"
  1.1240 +            using x by (rule restrict_sets_measurable) (auto simp: `finite J`)
  1.1241 +          also have "\<dots> = sigma_sets (space (PiF {J} (\<lambda>_. borel)))
  1.1242 +            {Pi' (J) F |F. (\<forall>j\<in>J. F j \<in> range enum_basis)}"
  1.1243 +            (is "_ = sigma_sets _ ?P")
  1.1244 +            by (rule enumerable_sigma_fprod_algebra_sigma_eq[OF `J \<noteq> {}` `finite J`])
  1.1245 +          also have "\<dots> \<subseteq> sets borel"
  1.1246 +          proof
  1.1247 +            fix x
  1.1248 +            assume "x \<in> sigma_sets (space (PiF {J} (\<lambda>_. borel))) ?P"
  1.1249 +            thus "x \<in> sets borel"
  1.1250 +            proof (rule sigma_sets.induct, safe)
  1.1251 +              fix F::"'i \<Rightarrow> 'a set"
  1.1252 +              assume "\<forall>j\<in>J. F j \<in> range enum_basis"
  1.1253 +              hence "Pi' J F \<in> range enum_basis_finmap"
  1.1254 +                unfolding range_enum_basis_eq
  1.1255 +                by (auto simp: `finite J` intro!: exI[where x=J] exI[where x=F])
  1.1256 +              hence "open (Pi' (J) F)" by (rule range_enum_basis_finmap_imp_open)
  1.1257 +              thus "Pi' (J) F \<in> sets borel" by simp
  1.1258 +            next
  1.1259 +              fix a::"('i \<Rightarrow>\<^isub>F 'a) set"
  1.1260 +              have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) =
  1.1261 +                Pi' (J) (\<lambda>_. UNIV)"
  1.1262 +                by (auto simp: space_PiF product_def)
  1.1263 +              moreover have "open (Pi' (J::'i set) (\<lambda>_. UNIV::'a set))"
  1.1264 +                by (intro open_Pi'I) auto
  1.1265 +              ultimately
  1.1266 +              have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) \<in> sets borel"
  1.1267 +                by simp
  1.1268 +              moreover
  1.1269 +              assume "a \<in> sets borel"
  1.1270 +              ultimately show "space (PiF {J} (\<lambda>_. borel)) - a \<in> sets borel" ..
  1.1271 +            qed auto
  1.1272 +          qed
  1.1273 +          finally have "(?x J) \<in> sets borel" .
  1.1274 +        } ultimately show "(?x J) \<in> sets borel" by blast
  1.1275 +      qed
  1.1276 +      finally show "x \<in> sets (borel)" .
  1.1277 +    qed
  1.1278 +  qed
  1.1279 +qed (simp add: emeasure_sigma borel_def PiF_def)
  1.1280 +
  1.1281 +subsection {* Isomorphism between Functions and Finite Maps *}
  1.1282 +
  1.1283 +lemma
  1.1284 +  measurable_compose:
  1.1285 +  fixes f::"'a \<Rightarrow> 'b"
  1.1286 +  assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
  1.1287 +  assumes "finite J"
  1.1288 +  shows "(\<lambda>m. compose J m f) \<in> measurable (PiM (f ` J) (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
  1.1289 +proof (rule measurable_PiM)
  1.1290 +  show "(\<lambda>m. compose J m f)
  1.1291 +    \<in> space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<rightarrow>
  1.1292 +      (J \<rightarrow> space M) \<inter> extensional J"
  1.1293 +  proof safe
  1.1294 +    fix x and i
  1.1295 +    assume x: "x \<in> space (PiM (f ` J) (\<lambda>_. M))" "i \<in> J"
  1.1296 +    with inj show  "compose J x f i \<in> space M"
  1.1297 +      by (auto simp: space_PiM compose_def)
  1.1298 +  next
  1.1299 +    fix x assume "x \<in> space (PiM (f ` J) (\<lambda>_. M))"
  1.1300 +    show "(compose J x f) \<in> extensional J" by (rule compose_extensional)
  1.1301 +  qed
  1.1302 +next
  1.1303 +  fix S X
  1.1304 +  have inv: "\<And>j. j \<in> f ` J \<Longrightarrow> f (f' j) = j" using assms by auto
  1.1305 +  assume S: "S \<noteq> {} \<or> J = {}" "finite S" "S \<subseteq> J" and P: "\<And>i. i \<in> S \<Longrightarrow> X i \<in> sets M"
  1.1306 +  have "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
  1.1307 +    space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) = prod_emb (f ` J) (\<lambda>_. M) (f ` S) (Pi\<^isub>E (f ` S) (\<lambda>b. X (f' b)))"
  1.1308 +    using assms inv S sets_into_space[OF P]
  1.1309 +    by (force simp: prod_emb_iff compose_def space_PiM extensional_def Pi_def intro: imageI)
  1.1310 +  also have "\<dots> \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))"
  1.1311 +  proof
  1.1312 +    from S show "f ` S \<subseteq> f `  J" by auto
  1.1313 +    show "(\<Pi>\<^isub>E b\<in>f ` S. X (f' b)) \<in> sets (Pi\<^isub>M (f ` S) (\<lambda>_. M))"
  1.1314 +    proof (rule sets_PiM_I_finite)
  1.1315 +      show "finite (f ` S)" using S by simp
  1.1316 +      fix i assume "i \<in> f ` S" hence "f' i \<in> S" using S assms by auto
  1.1317 +      thus "X (f' i) \<in> sets M" by (rule P)
  1.1318 +    qed
  1.1319 +  qed
  1.1320 +  finally show "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
  1.1321 +    space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))" .
  1.1322 +qed
  1.1323 +
  1.1324 +lemma
  1.1325 +  measurable_compose_inv:
  1.1326 +  fixes f::"'a \<Rightarrow> 'b"
  1.1327 +  assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
  1.1328 +  assumes "finite J"
  1.1329 +  shows "(\<lambda>m. compose (f ` J) m f') \<in> measurable (PiM J (\<lambda>_. M)) (PiM (f ` J) (\<lambda>_. M))"
  1.1330 +proof -
  1.1331 +  have "(\<lambda>m. compose (f ` J) m f') \<in> measurable (Pi\<^isub>M (f' ` f ` J) (\<lambda>_. M)) (Pi\<^isub>M (f ` J) (\<lambda>_. M))"
  1.1332 +    using assms by (auto intro: measurable_compose)
  1.1333 +  moreover
  1.1334 +  from inj have "f' ` f ` J = J" by (metis (hide_lams, mono_tags) image_iff set_eqI)
  1.1335 +  ultimately show ?thesis by simp
  1.1336 +qed
  1.1337 +
  1.1338 +locale function_to_finmap =
  1.1339 +  fixes J::"'a set" and f :: "'a \<Rightarrow> 'b::countable" and f'
  1.1340 +  assumes [simp]: "finite J"
  1.1341 +  assumes inv: "i \<in> J \<Longrightarrow> f' (f i) = i"
  1.1342 +begin
  1.1343 +
  1.1344 +text {* to measure finmaps *}
  1.1345 +
  1.1346 +definition "fm = (finmap_of (f ` J)) o (\<lambda>g. compose (f ` J) g f')"
  1.1347 +
  1.1348 +lemma domain_fm[simp]: "domain (fm x) = f ` J"
  1.1349 +  unfolding fm_def by simp
  1.1350 +
  1.1351 +lemma fm_restrict[simp]: "fm (restrict y J) = fm y"
  1.1352 +  unfolding fm_def by (auto simp: compose_def inv intro: restrict_ext)
  1.1353 +
  1.1354 +lemma fm_product:
  1.1355 +  assumes "\<And>i. space (M i) = UNIV"
  1.1356 +  shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^isub>M J M) = (\<Pi>\<^isub>E j \<in> J. S (f j))"
  1.1357 +  using assms
  1.1358 +  by (auto simp: inv fm_def compose_def space_PiM Pi'_def)
  1.1359 +
  1.1360 +lemma fm_measurable:
  1.1361 +  assumes "f ` J \<in> N"
  1.1362 +  shows "fm \<in> measurable (Pi\<^isub>M J (\<lambda>_. M)) (Pi\<^isub>F N (\<lambda>_. M))"
  1.1363 +  unfolding fm_def
  1.1364 +proof (rule measurable_comp, rule measurable_compose_inv)
  1.1365 +  show "finmap_of (f ` J) \<in> measurable (Pi\<^isub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
  1.1366 +    using assms by (intro measurable_finmap_of measurable_component_singleton) auto
  1.1367 +qed (simp_all add: inv)
  1.1368 +
  1.1369 +lemma proj_fm:
  1.1370 +  assumes "x \<in> J"
  1.1371 +  shows "fm m (f x) = m x"
  1.1372 +  using assms by (auto simp: fm_def compose_def o_def inv)
  1.1373 +
  1.1374 +lemma inj_on_compose_f': "inj_on (\<lambda>g. compose (f ` J) g f') (extensional J)"
  1.1375 +proof (rule inj_on_inverseI)
  1.1376 +  fix x::"'a \<Rightarrow> 'c" assume "x \<in> extensional J"
  1.1377 +  thus "(\<lambda>x. compose J x f) (compose (f ` J) x f') = x"
  1.1378 +    by (auto simp: compose_def inv extensional_def)
  1.1379 +qed
  1.1380 +
  1.1381 +lemma inj_on_fm:
  1.1382 +  assumes "\<And>i. space (M i) = UNIV"
  1.1383 +  shows "inj_on fm (space (Pi\<^isub>M J M))"
  1.1384 +  using assms
  1.1385 +  apply (auto simp: fm_def space_PiM)
  1.1386 +  apply (rule comp_inj_on)
  1.1387 +  apply (rule inj_on_compose_f')
  1.1388 +  apply (rule finmap_of_inj_on_extensional_finite)
  1.1389 +  apply simp
  1.1390 +  apply (auto)
  1.1391 +  done
  1.1392 +
  1.1393 +text {* to measure functions *}
  1.1394 +
  1.1395 +definition "mf = (\<lambda>g. compose J g f) o proj"
  1.1396 +
  1.1397 +lemma
  1.1398 +  assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))" "finite J"
  1.1399 +  shows "proj (finmap_of J x) = x"
  1.1400 +  using assms by (auto simp: space_PiM extensional_def)
  1.1401 +
  1.1402 +lemma
  1.1403 +  assumes "x \<in> space (Pi\<^isub>F {J} (\<lambda>_. M))"
  1.1404 +  shows "finmap_of J (proj x) = x"
  1.1405 +  using assms by (auto simp: space_PiF Pi'_def finmap_eq_iff)
  1.1406 +
  1.1407 +lemma mf_fm:
  1.1408 +  assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))"
  1.1409 +  shows "mf (fm x) = x"
  1.1410 +proof -
  1.1411 +  have "mf (fm x) \<in> extensional J"
  1.1412 +    by (auto simp: mf_def extensional_def compose_def)
  1.1413 +  moreover
  1.1414 +  have "x \<in> extensional J" using assms sets_into_space
  1.1415 +    by (force simp: space_PiM)
  1.1416 +  moreover
  1.1417 +  { fix i assume "i \<in> J"
  1.1418 +    hence "mf (fm x) i = x i"
  1.1419 +      by (auto simp: inv mf_def compose_def fm_def)
  1.1420 +  }
  1.1421 +  ultimately
  1.1422 +  show ?thesis by (rule extensionalityI)
  1.1423 +qed
  1.1424 +
  1.1425 +lemma mf_measurable:
  1.1426 +  assumes "space M = UNIV"
  1.1427 +  shows "mf \<in> measurable (PiF {f ` J} (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
  1.1428 +  unfolding mf_def
  1.1429 +proof (rule measurable_comp, rule measurable_proj_PiM)
  1.1430 +  show "(\<lambda>g. compose J g f) \<in>
  1.1431 +    measurable (Pi\<^isub>M (f ` J) (\<lambda>x. M)) (Pi\<^isub>M J (\<lambda>_. M))"
  1.1432 +    by (rule measurable_compose, rule inv) auto
  1.1433 +qed (auto simp add: space_PiM extensional_def assms)
  1.1434 +
  1.1435 +lemma fm_image_measurable:
  1.1436 +  assumes "space M = UNIV"
  1.1437 +  assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M))"
  1.1438 +  shows "fm ` X \<in> sets (PiF {f ` J} (\<lambda>_. M))"
  1.1439 +proof -
  1.1440 +  have "fm ` X = (mf) -` X \<inter> space (PiF {f ` J} (\<lambda>_. M))"
  1.1441 +  proof safe
  1.1442 +    fix x assume "x \<in> X"
  1.1443 +    with mf_fm[of x] sets_into_space[OF assms(2)] show "fm x \<in> mf -` X" by auto
  1.1444 +    show "fm x \<in> space (PiF {f ` J} (\<lambda>_. M))" by (simp add: space_PiF assms)
  1.1445 +  next
  1.1446 +    fix y x
  1.1447 +    assume x: "mf y \<in> X"
  1.1448 +    assume y: "y \<in> space (PiF {f ` J} (\<lambda>_. M))"
  1.1449 +    thus "y \<in> fm ` X"
  1.1450 +      by (intro image_eqI[OF _ x], unfold finmap_eq_iff)
  1.1451 +         (auto simp: space_PiF fm_def mf_def compose_def inv Pi'_def)
  1.1452 +  qed
  1.1453 +  also have "\<dots> \<in> sets (PiF {f ` J} (\<lambda>_. M))"
  1.1454 +    using assms
  1.1455 +    by (intro measurable_sets[OF mf_measurable]) auto
  1.1456 +  finally show ?thesis .
  1.1457 +qed
  1.1458 +
  1.1459 +lemma fm_image_measurable_finite:
  1.1460 +  assumes "space M = UNIV"
  1.1461 +  assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M::'c measure))"
  1.1462 +  shows "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. M::'c measure))"
  1.1463 +  using fm_image_measurable[OF assms]
  1.1464 +  by (rule subspace_set_in_sets) (auto simp: finite_subset)
  1.1465 +
  1.1466 +text {* measure on finmaps *}
  1.1467 +
  1.1468 +definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"
  1.1469 +
  1.1470 +lemma sets_mapmeasure[simp]: "sets (mapmeasure M N) = sets (PiF (Collect finite) N)"
  1.1471 +  unfolding mapmeasure_def by simp
  1.1472 +
  1.1473 +lemma space_mapmeasure[simp]: "space (mapmeasure M N) = space (PiF (Collect finite) N)"
  1.1474 +  unfolding mapmeasure_def by simp
  1.1475 +
  1.1476 +lemma mapmeasure_PiF:
  1.1477 +  assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
  1.1478 +  assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
  1.1479 +  assumes "space N = UNIV"
  1.1480 +  assumes "X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
  1.1481 +  shows "emeasure (mapmeasure M (\<lambda>_. N)) X = emeasure M ((fm -` X \<inter> extensional J))"
  1.1482 +  using assms
  1.1483 +  by (auto simp: measurable_eqI[OF s1 refl s2 refl] mapmeasure_def emeasure_distr
  1.1484 +    fm_measurable space_PiM)
  1.1485 +
  1.1486 +lemma mapmeasure_PiM:
  1.1487 +  fixes N::"'c measure"
  1.1488 +  assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
  1.1489 +  assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
  1.1490 +  assumes N: "space N = UNIV"
  1.1491 +  assumes X: "X \<in> sets M"
  1.1492 +  shows "emeasure M X = emeasure (mapmeasure M (\<lambda>_. N)) (fm ` X)"
  1.1493 +  unfolding mapmeasure_def
  1.1494 +proof (subst emeasure_distr, subst measurable_eqI[OF s1 refl s2 refl], rule fm_measurable)
  1.1495 +  have "X \<subseteq> space (Pi\<^isub>M J (\<lambda>_. N))" using assms by (simp add: sets_into_space)
  1.1496 +  from assms inj_on_fm[of "\<lambda>_. N"] set_mp[OF this] have "fm -` fm ` X \<inter> space (Pi\<^isub>M J (\<lambda>_. N)) = X"
  1.1497 +    by (auto simp: vimage_image_eq inj_on_def)
  1.1498 +  thus "emeasure M X = emeasure M (fm -` fm ` X \<inter> space M)" using s1
  1.1499 +    by simp
  1.1500 +  show "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
  1.1501 +    by (rule fm_image_measurable_finite[OF N X[simplified s2]])
  1.1502 +qed simp
  1.1503 +
  1.1504 +end
  1.1505 +
  1.1506 +end
     2.1 --- a/src/HOL/Probability/Probability.thy	Thu Nov 15 10:49:58 2012 +0100
     2.2 +++ b/src/HOL/Probability/Probability.thy	Thu Nov 15 11:16:58 2012 +0100
     2.3 @@ -3,7 +3,7 @@
     2.4    Complete_Measure
     2.5    Probability_Measure
     2.6    Infinite_Product_Measure
     2.7 -  Regularity
     2.8 +  Projective_Limit
     2.9    Independent_Family
    2.10    Information
    2.11  begin
     3.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     3.2 +++ b/src/HOL/Probability/Projective_Limit.thy	Thu Nov 15 11:16:58 2012 +0100
     3.3 @@ -0,0 +1,690 @@
     3.4 +(*  Title:      HOL/Probability/Projective_Family.thy
     3.5 +    Author:     Fabian Immler, TU München
     3.6 +*)
     3.7 +
     3.8 +header {* Projective Limit *}
     3.9 +
    3.10 +theory Projective_Limit
    3.11 +  imports
    3.12 +    Caratheodory
    3.13 +    Fin_Map
    3.14 +    Regularity
    3.15 +    Projective_Family
    3.16 +    Infinite_Product_Measure
    3.17 +begin
    3.18 +
    3.19 +subsection {* Enumeration of Countable Union of Finite Sets *}
    3.20 +
    3.21 +locale finite_set_sequence =
    3.22 +  fixes Js::"nat \<Rightarrow> 'a set"
    3.23 +  assumes finite_seq[simp]: "finite (Js n)"
    3.24 +begin
    3.25 +
    3.26 +text {* Enumerate finite set *}
    3.27 +
    3.28 +definition "enum_finite_max J = (SOME n. \<exists> f. J = f ` {i. i < n} \<and> inj_on f {i. i < n})"
    3.29 +
    3.30 +definition enum_finite where
    3.31 +  "enum_finite J =
    3.32 +    (SOME f. J = f ` {i::nat. i < enum_finite_max J} \<and> inj_on f {i. i < enum_finite_max J})"
    3.33 +
    3.34 +lemma enum_finite_max:
    3.35 +  assumes "finite J"
    3.36 +  shows "\<exists>f::nat\<Rightarrow>_. J = f ` {i. i < enum_finite_max J} \<and> inj_on f {i. i < enum_finite_max J}"
    3.37 +    unfolding enum_finite_max_def
    3.38 +    by (rule someI_ex) (rule finite_imp_nat_seg_image_inj_on[OF `finite J`])
    3.39 +
    3.40 +lemma enum_finite:
    3.41 +  assumes "finite J"
    3.42 +  shows "J = enum_finite J ` {i::nat. i < enum_finite_max J} \<and>
    3.43 +    inj_on (enum_finite J) {i::nat. i < enum_finite_max J}"
    3.44 +  unfolding enum_finite_def
    3.45 +  by (rule someI_ex[of "\<lambda>f. J = f ` {i::nat. i < enum_finite_max J} \<and>
    3.46 +    inj_on f {i. i < enum_finite_max J}"]) (rule enum_finite_max[OF `finite J`])
    3.47 +
    3.48 +lemma in_set_enum_exist:
    3.49 +  assumes "finite A"
    3.50 +  assumes "y \<in> A"
    3.51 +  shows "\<exists>i. y = enum_finite A i"
    3.52 +  using assms enum_finite by auto
    3.53 +
    3.54 +definition set_of_Un where "set_of_Un j = (LEAST n. j \<in> Js n)"
    3.55 +
    3.56 +definition index_in_set where "index_in_set J j = (SOME n. j = enum_finite J n)"
    3.57 +
    3.58 +definition Un_to_nat where
    3.59 +  "Un_to_nat j = to_nat (set_of_Un j, index_in_set (Js (set_of_Un j)) j)"
    3.60 +
    3.61 +lemma inj_on_Un_to_nat:
    3.62 +  shows "inj_on Un_to_nat (\<Union>n::nat. Js n)"
    3.63 +proof (rule inj_onI)
    3.64 +  fix x y
    3.65 +  assume "x \<in> (\<Union>n. Js n)" "y \<in> (\<Union>n. Js n)"
    3.66 +  then obtain ix iy where ix: "x \<in> Js ix" and iy: "y \<in> Js iy" by blast
    3.67 +  assume "Un_to_nat x = Un_to_nat y"
    3.68 +  hence "set_of_Un x = set_of_Un y"
    3.69 +    "index_in_set (Js (set_of_Un y)) y = index_in_set (Js (set_of_Un x)) x"
    3.70 +    by (auto simp: Un_to_nat_def)
    3.71 +  moreover
    3.72 +  {
    3.73 +    fix x assume "x \<in> Js (set_of_Un x)"
    3.74 +    have "x = enum_finite (Js (set_of_Un x)) (index_in_set (Js (set_of_Un x)) x)"
    3.75 +      unfolding index_in_set_def
    3.76 +      apply (rule someI_ex)
    3.77 +      using `x \<in> Js (set_of_Un x)` finite_seq
    3.78 +      apply (auto intro!: in_set_enum_exist)
    3.79 +      done
    3.80 +  } note H = this
    3.81 +  moreover
    3.82 +  have "y \<in> Js (set_of_Un y)" unfolding set_of_Un_def using iy by (rule LeastI)
    3.83 +  note H[OF this]
    3.84 +  moreover
    3.85 +  have "x \<in> Js (set_of_Un x)" unfolding set_of_Un_def using ix by (rule LeastI)
    3.86 +  note H[OF this]
    3.87 +  ultimately show "x = y" by simp
    3.88 +qed
    3.89 +
    3.90 +lemma inj_Un[simp]:
    3.91 +  shows "inj_on (Un_to_nat) (Js n)"
    3.92 +  by (intro subset_inj_on[OF inj_on_Un_to_nat]) (auto simp: assms)
    3.93 +
    3.94 +lemma Un_to_nat_injectiveD:
    3.95 +  assumes "Un_to_nat x = Un_to_nat y"
    3.96 +  assumes "x \<in> Js i" "y \<in> Js j"
    3.97 +  shows "x = y"
    3.98 +  using assms
    3.99 +  by (intro inj_onD[OF inj_on_Un_to_nat]) auto
   3.100 +
   3.101 +end
   3.102 +
   3.103 +subsection {* Sequences of Finite Maps in Compact Sets *}
   3.104 +
   3.105 +locale finmap_seqs_into_compact =
   3.106 +  fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a)" and M
   3.107 +  assumes compact: "\<And>n. compact (K n)"
   3.108 +  assumes f_in_K: "\<And>n. K n \<noteq> {}"
   3.109 +  assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)"
   3.110 +  assumes proj_in_K:
   3.111 +    "\<And>t n m. m \<ge> n \<Longrightarrow> t \<in> domain (f n) \<Longrightarrow> (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n"
   3.112 +begin
   3.113 +
   3.114 +lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n)"
   3.115 +  using proj_in_K f_in_K
   3.116 +proof cases
   3.117 +  obtain k where "k \<in> K (Suc 0)" using f_in_K by auto
   3.118 +  assume "\<forall>n. t \<notin> domain (f n)"
   3.119 +  thus ?thesis
   3.120 +    by (auto intro!: exI[where x=1] image_eqI[OF _ `k \<in> K (Suc 0)`]
   3.121 +      simp: domain_K[OF `k \<in> K (Suc 0)`])
   3.122 +qed blast
   3.123 +
   3.124 +lemma proj_in_KE:
   3.125 +  obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n"
   3.126 +  using proj_in_K' by blast
   3.127 +
   3.128 +lemma compact_projset:
   3.129 +  shows "compact ((\<lambda>k. (k)\<^isub>F i) ` K n)"
   3.130 +  using continuous_proj compact by (rule compact_continuous_image)
   3.131 +
   3.132 +end
   3.133 +
   3.134 +lemma compactE':
   3.135 +  assumes "compact S" "\<forall>n\<ge>m. f n \<in> S"
   3.136 +  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
   3.137 +proof atomize_elim
   3.138 +  have "subseq (op + m)" by (simp add: subseq_def)
   3.139 +  have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> S" using assms by auto
   3.140 +  from compactE[OF `compact S` this] guess l r .
   3.141 +  hence "l \<in> S" "subseq ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) ----> l"
   3.142 +    using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def)
   3.143 +  thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) ----> l" by blast
   3.144 +qed
   3.145 +
   3.146 +sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^isub>F n) ----> l)"
   3.147 +proof
   3.148 +  fix n s
   3.149 +  assume "subseq s"
   3.150 +  from proj_in_KE[of n] guess n0 . note n0 = this
   3.151 +  have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0"
   3.152 +  proof safe
   3.153 +    fix i assume "n0 \<le> i"
   3.154 +    also have "\<dots> \<le> s i" by (rule seq_suble) fact
   3.155 +    finally have "n0 \<le> s i" .
   3.156 +    with n0 show "((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0 "
   3.157 +      by auto
   3.158 +  qed
   3.159 +  from compactE'[OF compact_projset this] guess ls rs .
   3.160 +  thus "\<exists>r'. subseq r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^isub>F n) ----> l)" by (auto simp: o_def)
   3.161 +qed
   3.162 +
   3.163 +lemma (in finmap_seqs_into_compact)
   3.164 +  diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^isub>F n) ----> l"
   3.165 +proof -
   3.166 +  have "\<And>i n0. (f o seqseq i) i = f (diagseq i)" unfolding diagseq_def by simp
   3.167 +  from reducer_reduces obtain l where l: "(\<lambda>i. ((f \<circ> seqseq (Suc n)) i)\<^isub>F n) ----> l"
   3.168 +    unfolding seqseq_reducer
   3.169 +  by auto
   3.170 +  have "(\<lambda>i. (f (diagseq (i + Suc n)))\<^isub>F n) =
   3.171 +    (\<lambda>i. ((f o (diagseq o (op + (Suc n)))) i)\<^isub>F n)" by (simp add: add_commute)
   3.172 +  also have "\<dots> =
   3.173 +    (\<lambda>i. ((f o ((seqseq (Suc n) o (\<lambda>x. fold_reduce (Suc n) x (Suc n + x))))) i)\<^isub>F n)"
   3.174 +    unfolding diagseq_seqseq by simp
   3.175 +  also have "\<dots> = (\<lambda>i. ((f o ((seqseq (Suc n)))) i)\<^isub>F n) o (\<lambda>x. fold_reduce (Suc n) x (Suc n + x))"
   3.176 +    by (simp add: o_def)
   3.177 +  also have "\<dots> ----> l"
   3.178 +  proof (rule LIMSEQ_subseq_LIMSEQ[OF _ subseq_diagonal_rest], rule tendstoI)
   3.179 +    fix e::real assume "0 < e"
   3.180 +    from tendstoD[OF l `0 < e`]
   3.181 +    show "eventually (\<lambda>x. dist (((f \<circ> seqseq (Suc n)) x)\<^isub>F n) l < e)
   3.182 +      sequentially" .
   3.183 +  qed
   3.184 +  finally show ?thesis by (intro exI) (rule LIMSEQ_offset)
   3.185 +qed
   3.186 +
   3.187 +subsection {* Daniell-Kolmogorov Theorem *}
   3.188 +
   3.189 +text {* Existence of Projective Limit *}
   3.190 +
   3.191 +locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure"
   3.192 +  for I::"'i set" and P
   3.193 +begin
   3.194 +
   3.195 +abbreviation "PiB \<equiv> (\<lambda>J P. PiP J (\<lambda>_. borel) P)"
   3.196 +
   3.197 +lemma
   3.198 +  emeasure_PiB_emb_not_empty:
   3.199 +  assumes "I \<noteq> {}"
   3.200 +  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
   3.201 +  shows "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (PiB J P) (Pi\<^isub>E J B)"
   3.202 +proof -
   3.203 +  let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space borel"
   3.204 +  let ?G = generator
   3.205 +  interpret G!: algebra ?\<Omega> generator by (intro  algebra_generator) fact
   3.206 +  note \<mu>G_mono =
   3.207 +    G.additive_increasing[OF positive_\<mu>G[OF `I \<noteq> {}`] additive_\<mu>G[OF `I \<noteq> {}`], THEN increasingD]
   3.208 +  have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
   3.209 +  proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G,
   3.210 +      OF `I \<noteq> {}`, OF `I \<noteq> {}`])
   3.211 +    fix A assume "A \<in> ?G"
   3.212 +    with generatorE guess J X . note JX = this
   3.213 +    interpret prob_space "P J" using prob_space[OF `finite J`] .
   3.214 +    show "\<mu>G A \<noteq> \<infinity>" using JX by (simp add: PiP_finite)
   3.215 +  next
   3.216 +    fix Z assume Z: "range Z \<subseteq> ?G" "decseq Z" "(\<Inter>i. Z i) = {}"
   3.217 +    then have "decseq (\<lambda>i. \<mu>G (Z i))"
   3.218 +      by (auto intro!: \<mu>G_mono simp: decseq_def)
   3.219 +    moreover
   3.220 +    have "(INF i. \<mu>G (Z i)) = 0"
   3.221 +    proof (rule ccontr)
   3.222 +      assume "(INF i. \<mu>G (Z i)) \<noteq> 0" (is "?a \<noteq> 0")
   3.223 +      moreover have "0 \<le> ?a"
   3.224 +        using Z positive_\<mu>G[OF `I \<noteq> {}`] by (auto intro!: INF_greatest simp: positive_def)
   3.225 +      ultimately have "0 < ?a" by auto
   3.226 +      hence "?a \<noteq> -\<infinity>" by auto
   3.227 +      have "\<forall>n. \<exists>J B. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> B \<in> sets (Pi\<^isub>M J (\<lambda>_. borel)) \<and>
   3.228 +        Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (PiB J P) B"
   3.229 +        using Z by (intro allI generator_Ex) auto
   3.230 +      then obtain J' B' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I"
   3.231 +          "\<And>n. B' n \<in> sets (\<Pi>\<^isub>M i\<in>J' n. borel)"
   3.232 +        and Z_emb: "\<And>n. Z n = emb I (J' n) (B' n)"
   3.233 +        unfolding choice_iff by blast
   3.234 +      moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
   3.235 +      moreover def B \<equiv> "\<lambda>n. emb (J n) (J' n) (B' n)"
   3.236 +      ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I"
   3.237 +        "\<And>n. B n \<in> sets (\<Pi>\<^isub>M i\<in>J n. borel)"
   3.238 +        by auto
   3.239 +      have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
   3.240 +        unfolding J_def by force
   3.241 +      have "\<forall>n. \<exists>j. j \<in> J n" using J by blast
   3.242 +      then obtain j where j: "\<And>n. j n \<in> J n"
   3.243 +        unfolding choice_iff by blast
   3.244 +      note [simp] = `\<And>n. finite (J n)`
   3.245 +      from J  Z_emb have Z_eq: "\<And>n. Z n = emb I (J n) (B n)" "\<And>n. Z n \<in> ?G"
   3.246 +        unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto)
   3.247 +      interpret prob_space "P (J i)" for i using prob_space by simp
   3.248 +      have "?a \<le> \<mu>G (Z 0)" by (auto intro: INF_lower)
   3.249 +      also have "\<dots> < \<infinity>" using J by (auto simp: Z_eq \<mu>G_eq PiP_finite proj_sets)
   3.250 +      finally have "?a \<noteq> \<infinity>" by simp
   3.251 +      have "\<And>n. \<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" unfolding Z_eq using J J_mono
   3.252 +        by (subst \<mu>G_eq) (auto simp: PiP_finite proj_sets \<mu>G_eq)
   3.253 +
   3.254 +      interpret finite_set_sequence J by unfold_locales simp
   3.255 +      def Utn \<equiv> Un_to_nat
   3.256 +      interpret function_to_finmap "J n" Utn "inv_into (J n) Utn" for n
   3.257 +        by unfold_locales (auto simp: Utn_def)
   3.258 +      def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
   3.259 +      let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
   3.260 +      {
   3.261 +        fix n
   3.262 +        interpret finite_measure "P (J n)" by unfold_locales
   3.263 +        have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
   3.264 +          using J by (auto simp: P'_def mapmeasure_PiM proj_space proj_sets)
   3.265 +        also
   3.266 +        have "\<dots> = ?SUP n"
   3.267 +        proof (rule inner_regular)
   3.268 +          show "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
   3.269 +            unfolding P'_def
   3.270 +            by (auto simp: P'_def mapmeasure_PiF fm_measurable proj_space proj_sets)
   3.271 +          show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
   3.272 +        next
   3.273 +          show "fm n ` B n \<in> sets borel"
   3.274 +            unfolding borel_eq_PiF_borel
   3.275 +            by (auto simp del: J(2) simp: P'_def fm_image_measurable_finite proj_sets J)
   3.276 +        qed
   3.277 +        finally
   3.278 +        have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
   3.279 +      } note R = this
   3.280 +      have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a
   3.281 +        \<and> compact K \<and> K \<subseteq> fm n ` B n"
   3.282 +      proof
   3.283 +        fix n
   3.284 +        have "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
   3.285 +          by (simp add: mapmeasure_PiF P'_def proj_space proj_sets)
   3.286 +        then interpret finite_measure "P' n" ..
   3.287 +        show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
   3.288 +            compact K \<and> K \<subseteq> fm n ` B n"
   3.289 +          unfolding R
   3.290 +        proof (rule ccontr)
   3.291 +          assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n)  * ?a \<and>
   3.292 +            compact K' \<and> K' \<subseteq> fm n ` B n)"
   3.293 +          have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
   3.294 +          proof (intro SUP_least)
   3.295 +            fix K
   3.296 +            assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
   3.297 +            with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
   3.298 +              by auto
   3.299 +            hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
   3.300 +              unfolding not_less[symmetric] by simp
   3.301 +            hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
   3.302 +              using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps)
   3.303 +            thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
   3.304 +          qed
   3.305 +          hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
   3.306 +          hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
   3.307 +          hence "0 \<le> - (2 powr (-n) * ?a)"
   3.308 +            using `?SUP _ \<noteq> \<infinity>` `?SUP _ \<noteq> - \<infinity>`
   3.309 +            by (subst (asm) ereal_add_le_add_iff) (auto simp:)
   3.310 +          moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a`
   3.311 +            by (auto simp: ereal_zero_less_0_iff)
   3.312 +          ultimately show False by simp
   3.313 +        qed
   3.314 +      qed
   3.315 +      then obtain K' where K':
   3.316 +        "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
   3.317 +        "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
   3.318 +        unfolding choice_iff by blast
   3.319 +      def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
   3.320 +      have K_sets: "\<And>n. K n \<in> sets (Pi\<^isub>M (J n) (\<lambda>_. borel))"
   3.321 +        unfolding K_def
   3.322 +        using compact_imp_closed[OF `compact (K' _)`]
   3.323 +        by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
   3.324 +           (auto simp: borel_eq_PiF_borel[symmetric])
   3.325 +      have K_B: "\<And>n. K n \<subseteq> B n"
   3.326 +      proof
   3.327 +        fix x n
   3.328 +        assume "x \<in> K n" hence fm_in: "fm n x \<in> fm n ` B n"
   3.329 +          using K' by (force simp: K_def)
   3.330 +        show "x \<in> B n"
   3.331 +          apply (rule inj_on_image_mem_iff[OF inj_on_fm _ fm_in])
   3.332 +          using `x \<in> K n` K_sets J[of n] sets_into_space
   3.333 +          apply (auto simp: proj_space)
   3.334 +          using J[of n] sets_into_space apply auto
   3.335 +          done
   3.336 +      qed
   3.337 +      def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)"
   3.338 +      have Z': "\<And>n. Z' n \<subseteq> Z n"
   3.339 +        unfolding Z_eq unfolding Z'_def
   3.340 +      proof (rule prod_emb_mono, safe)
   3.341 +        fix n x assume "x \<in> K n"
   3.342 +        hence "fm n x \<in> K' n" "x \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
   3.343 +          by (simp_all add: K_def proj_space)
   3.344 +        note this(1)
   3.345 +        also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
   3.346 +        finally have "fm n x \<in> fm n ` B n" .
   3.347 +        thus "x \<in> B n"
   3.348 +        proof safe
   3.349 +          fix y assume "y \<in> B n"
   3.350 +          moreover
   3.351 +          hence "y \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))" using J sets_into_space[of "B n" "P (J n)"]
   3.352 +            by (auto simp add: proj_space proj_sets)
   3.353 +          assume "fm n x = fm n y"
   3.354 +          note inj_onD[OF inj_on_fm[OF space_borel],
   3.355 +            OF `fm n x = fm n y` `x \<in> space _` `y \<in> space _`]
   3.356 +          ultimately show "x \<in> B n" by simp
   3.357 +        qed
   3.358 +      qed
   3.359 +      { fix n
   3.360 +        have "Z' n \<in> ?G" using K' unfolding Z'_def
   3.361 +          apply (intro generatorI'[OF J(1-3)])
   3.362 +          unfolding K_def proj_space
   3.363 +          apply (rule measurable_sets[OF fm_measurable[of _ "Collect finite"]])
   3.364 +          apply (auto simp add: P'_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
   3.365 +          done
   3.366 +      }
   3.367 +      def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
   3.368 +      hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
   3.369 +      hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
   3.370 +      have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
   3.371 +      hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto
   3.372 +      have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
   3.373 +      proof -
   3.374 +        fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact
   3.375 +        have "Y n = (\<Inter> i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
   3.376 +          by (auto simp: Y_def Z'_def)
   3.377 +        also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter> i\<in>{1..n}. emb (J n) (J i) (K i))"
   3.378 +          using `n \<ge> 1`
   3.379 +          by (subst prod_emb_INT) auto
   3.380 +        finally
   3.381 +        have Y_emb:
   3.382 +          "Y n = prod_emb I (\<lambda>_. borel) (J n)
   3.383 +            (\<Inter> i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
   3.384 +        hence "Y n \<in> ?G" using J J_mono K_sets `n \<ge> 1` by (intro generatorI[OF _ _ _ _ Y_emb]) auto
   3.385 +        hence "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" unfolding Y_emb using J J_mono K_sets `n \<ge> 1`
   3.386 +          by (subst \<mu>G_eq) (auto simp: PiP_finite proj_sets \<mu>G_eq)
   3.387 +        interpret finite_measure "(PiP (J n) (\<lambda>_. borel) P)"
   3.388 +        proof
   3.389 +          have "emeasure (PiP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^isub>E space borel) \<noteq> \<infinity>"
   3.390 +            using J by (subst emeasure_PiP) auto
   3.391 +          thus  "emeasure (PiP (J n) (\<lambda>_. borel) P) (space (PiP (J n) (\<lambda>_. borel) P)) \<noteq> \<infinity>"
   3.392 +             by (simp add: space_PiM)
   3.393 +        qed
   3.394 +        have "\<mu>G (Z n) = PiP (J n) (\<lambda>_. borel) P (B n)"
   3.395 +          unfolding Z_eq using J by (auto simp: \<mu>G_eq)
   3.396 +        moreover have "\<mu>G (Y n) =
   3.397 +          PiP (J n) (\<lambda>_. borel) P (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
   3.398 +          unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst \<mu>G_eq) auto
   3.399 +        moreover have "\<mu>G (Z n - Y n) = PiP (J n) (\<lambda>_. borel) P
   3.400 +          (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
   3.401 +          unfolding Z_eq Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1`
   3.402 +          by (subst \<mu>G_eq) (auto intro!: Diff)
   3.403 +        ultimately
   3.404 +        have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
   3.405 +          using J J_mono K_sets `n \<ge> 1`
   3.406 +          by (simp only: emeasure_eq_measure)
   3.407 +            (auto dest!: bspec[where x=n]
   3.408 +            simp: extensional_restrict emeasure_eq_measure prod_emb_iff
   3.409 +            intro!: measure_Diff[symmetric] set_mp[OF K_B])
   3.410 +        also have subs: "Z n - Y n \<subseteq> (\<Union> i\<in>{1..n}. (Z i - Z' i))" using Z' Z `n \<ge> 1`
   3.411 +          unfolding Y_def by (force simp: decseq_def)
   3.412 +        have "Z n - Y n \<in> ?G" "(\<Union> i\<in>{1..n}. (Z i - Z' i)) \<in> ?G"
   3.413 +          using `Z' _ \<in> ?G` `Z _ \<in> ?G` `Y _ \<in> ?G` by auto
   3.414 +        hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union> i\<in>{1..n}. (Z i - Z' i))"
   3.415 +          using subs G.additive_increasing[OF positive_\<mu>G[OF `I \<noteq> {}`] additive_\<mu>G[OF `I \<noteq> {}`]]
   3.416 +          unfolding increasing_def by auto
   3.417 +        also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> ?G` `Z' _ \<in> ?G`
   3.418 +          by (intro G.subadditive[OF positive_\<mu>G additive_\<mu>G, OF `I \<noteq> {}` `I \<noteq> {}`]) auto
   3.419 +        also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
   3.420 +        proof (rule setsum_mono)
   3.421 +          fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
   3.422 +          have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
   3.423 +            unfolding Z'_def Z_eq by simp
   3.424 +          also have "\<dots> = P (J i) (B i - K i)"
   3.425 +            apply (subst \<mu>G_eq) using J K_sets apply auto
   3.426 +            apply (subst PiP_finite) apply auto
   3.427 +            done
   3.428 +          also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
   3.429 +            apply (subst emeasure_Diff) using K_sets J `K _ \<subseteq> B _` apply (auto simp: proj_sets)
   3.430 +            done
   3.431 +          also have "\<dots> = P (J i) (B i) - P' i (K' i)"
   3.432 +            unfolding K_def P'_def
   3.433 +            by (auto simp: mapmeasure_PiF proj_space proj_sets borel_eq_PiF_borel[symmetric]
   3.434 +              compact_imp_closed[OF `compact (K' _)`] space_PiM)
   3.435 +          also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
   3.436 +          finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
   3.437 +        qed
   3.438 +        also have "\<dots> = (\<Sum> i\<in>{1..n}. ereal (2 powr -real i) * ereal(real ?a))"
   3.439 +          using `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` by (subst ereal_real') auto
   3.440 +        also have "\<dots> = ereal (\<Sum> i\<in>{1..n}. (2 powr -real i) * (real ?a))" by simp
   3.441 +        also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real i)) * real ?a)"
   3.442 +          by (simp add: setsum_left_distrib)
   3.443 +        also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps
   3.444 +        proof (rule mult_strict_right_mono)
   3.445 +          have "(\<Sum>i\<in>{1..n}. 2 powr - real i) = (\<Sum>i\<in>{1..<Suc n}. (1/2) ^ i)"
   3.446 +            by (rule setsum_cong)
   3.447 +               (auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide)
   3.448 +          also have "{1..<Suc n} = {0..<Suc n} - {0}" by auto
   3.449 +          also have "setsum (op ^ (1 / 2::real)) ({0..<Suc n} - {0}) =
   3.450 +            setsum (op ^ (1 / 2)) ({0..<Suc n}) - 1" by (auto simp: setsum_diff1)
   3.451 +          also have "\<dots> < 1" by (subst sumr_geometric) auto
   3.452 +          finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" .
   3.453 +        qed (auto simp:
   3.454 +          `0 < ?a` `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` ereal_less_real_iff zero_ereal_def[symmetric])
   3.455 +        also have "\<dots> = ?a" using `0 < ?a` `?a \<noteq> \<infinity>` by (auto simp: ereal_real')
   3.456 +        also have "\<dots> \<le> \<mu>G (Z n)" by (auto intro: INF_lower)
   3.457 +        finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
   3.458 +        hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
   3.459 +          using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less)
   3.460 +        have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
   3.461 +        also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
   3.462 +          apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
   3.463 +        finally have "\<mu>G (Y n) > 0"
   3.464 +          using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric])
   3.465 +        thus "Y n \<noteq> {}" using positive_\<mu>G `I \<noteq> {}` by (auto simp add: positive_def)
   3.466 +      qed
   3.467 +      hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
   3.468 +      then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force
   3.469 +      {
   3.470 +        fix t and n m::nat
   3.471 +        assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
   3.472 +        from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto
   3.473 +        also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] .
   3.474 +        finally
   3.475 +        have "fm n (restrict (y m) (J n)) \<in> K' n"
   3.476 +          unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
   3.477 +        moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
   3.478 +          using J by (simp add: fm_def)
   3.479 +        ultimately have "fm n (y m) \<in> K' n" by simp
   3.480 +      } note fm_in_K' = this
   3.481 +      interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel
   3.482 +      proof
   3.483 +        fix n show "compact (K' n)" by fact
   3.484 +      next
   3.485 +        fix n
   3.486 +        from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto
   3.487 +        also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto
   3.488 +        finally
   3.489 +        have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"
   3.490 +          unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
   3.491 +        thus "K' (Suc n) \<noteq> {}" by auto
   3.492 +        fix k
   3.493 +        assume "k \<in> K' (Suc n)"
   3.494 +        with K'[of "Suc n"] sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto
   3.495 +        then obtain b where "k = fm (Suc n) b" by auto
   3.496 +        thus "domain k = domain (fm (Suc n) (y (Suc n)))"
   3.497 +          by (simp_all add: fm_def)
   3.498 +      next
   3.499 +        fix t and n m::nat
   3.500 +        assume "n \<le> m" hence "Suc n \<le> Suc m" by simp
   3.501 +        assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
   3.502 +        then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
   3.503 +        hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
   3.504 +        have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
   3.505 +          by (intro fm_in_K') simp_all
   3.506 +        show "(fm (Suc m) (y (Suc m)))\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K' (Suc n)"
   3.507 +          apply (rule image_eqI[OF _ img])
   3.508 +          using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
   3.509 +          unfolding j by (subst proj_fm, auto)+
   3.510 +      qed
   3.511 +      have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z"
   3.512 +        using diagonal_tendsto ..
   3.513 +      then obtain z where z:
   3.514 +        "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
   3.515 +        unfolding choice_iff by blast
   3.516 +      {
   3.517 +        fix n :: nat assume "n \<ge> 1"
   3.518 +        have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
   3.519 +          by simp
   3.520 +        moreover
   3.521 +        {
   3.522 +          fix t
   3.523 +          assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
   3.524 +          hence "t \<in> Utn ` J n" by simp
   3.525 +          then obtain j where j: "t = Utn j" "j \<in> J n" by auto
   3.526 +          have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
   3.527 +            apply (subst (2) tendsto_iff, subst eventually_sequentially)
   3.528 +          proof safe
   3.529 +            fix e :: real assume "0 < e"
   3.530 +            { fix i x assume "i \<ge> n" "t \<in> domain (fm n x)"
   3.531 +              moreover
   3.532 +              hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
   3.533 +              ultimately have "(fm i x)\<^isub>F t = (fm n x)\<^isub>F t"
   3.534 +                using j by (auto simp: proj_fm dest!:
   3.535 +                  Un_to_nat_injectiveD[simplified Utn_def[symmetric]])
   3.536 +            } note index_shift = this
   3.537 +            have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
   3.538 +              apply (rule le_SucI)
   3.539 +              apply (rule order_trans) apply simp
   3.540 +              apply (rule seq_suble[OF subseq_diagseq])
   3.541 +              done
   3.542 +            from z
   3.543 +            have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e"
   3.544 +              unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
   3.545 +            then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
   3.546 +              dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e" by auto
   3.547 +            show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e "
   3.548 +            proof (rule exI[where x="max N n"], safe)
   3.549 +              fix na assume "max N n \<le> na"
   3.550 +              hence  "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) =
   3.551 +                      dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^isub>F t) (z t)" using t
   3.552 +                by (subst index_shift[OF I]) auto
   3.553 +              also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp
   3.554 +              finally show "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e" .
   3.555 +            qed
   3.556 +          qed
   3.557 +          hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> (finmap_of (Utn ` J n) z)\<^isub>F t"
   3.558 +            by (simp add: tendsto_intros)
   3.559 +        } ultimately
   3.560 +        have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
   3.561 +          by (rule tendsto_finmap)
   3.562 +        hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) ----> finmap_of (Utn ` J n) z"
   3.563 +          by (intro lim_subseq) (simp add: subseq_def)
   3.564 +        moreover
   3.565 +        have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
   3.566 +          apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI)
   3.567 +          apply (rule le_trans)
   3.568 +          apply (rule le_add2)
   3.569 +          using seq_suble[OF subseq_diagseq]
   3.570 +          apply auto
   3.571 +          done
   3.572 +        moreover
   3.573 +        from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed)
   3.574 +        ultimately
   3.575 +        have "finmap_of (Utn ` J n) z \<in> K' n"
   3.576 +          unfolding closed_sequential_limits by blast
   3.577 +        also have "finmap_of (Utn ` J n) z  = fm n (\<lambda>i. z (Utn i))"
   3.578 +          by (auto simp: finmap_eq_iff fm_def compose_def f_inv_into_f)
   3.579 +        finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
   3.580 +        moreover
   3.581 +        let ?J = "\<Union>n. J n"
   3.582 +        have "(?J \<inter> J n) = J n" by auto
   3.583 +        ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"
   3.584 +          unfolding K_def by (auto simp: proj_space space_PiM)
   3.585 +        hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def
   3.586 +          using J by (auto simp: prod_emb_def extensional_def)
   3.587 +        also have "\<dots> \<subseteq> Z n" using Z' by simp
   3.588 +        finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .
   3.589 +      } note in_Z = this
   3.590 +      hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
   3.591 +      hence "(\<Inter>i. Z i) \<noteq> {}" using Z INT_decseq_offset[OF `decseq Z`] by simp
   3.592 +      thus False using Z by simp
   3.593 +    qed
   3.594 +    ultimately show "(\<lambda>i. \<mu>G (Z i)) ----> 0"
   3.595 +      using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (Z i)"] by simp
   3.596 +  qed
   3.597 +  then guess \<mu> .. note \<mu> = this
   3.598 +  def f \<equiv> "finmap_of J B"
   3.599 +  show "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (PiB J P) (Pi\<^isub>E J B)"
   3.600 +  proof (subst emeasure_extend_measure_Pair[OF PiP_def, of I "\<lambda>_. borel" \<mu>])
   3.601 +    show "positive (sets (PiB I P)) \<mu>" "countably_additive (sets (PiB I P)) \<mu>"
   3.602 +      using \<mu> unfolding sets_PiP sets_PiM_generator by (auto simp: measure_space_def)
   3.603 +  next
   3.604 +    show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> B \<in> J \<rightarrow> sets borel"
   3.605 +      using assms by (auto simp: f_def)
   3.606 +  next
   3.607 +    fix J and X::"'i \<Rightarrow> 'a set"
   3.608 +    show "prod_emb I (\<lambda>_. borel) J (Pi\<^isub>E J X) \<in> Pow ((I \<rightarrow> space borel) \<inter> extensional I)"
   3.609 +      by (auto simp: prod_emb_def)
   3.610 +    assume JX: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> J \<rightarrow> sets borel"
   3.611 +    hence "emb I J (Pi\<^isub>E J X) \<in> generator" using assms
   3.612 +      by (intro generatorI[where J=J and X="Pi\<^isub>E J X"]) (auto intro: sets_PiM_I_finite)
   3.613 +    hence "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))" using \<mu> by simp
   3.614 +    also have "\<dots> = emeasure (P J) (Pi\<^isub>E J X)"
   3.615 +      using JX assms proj_sets
   3.616 +      by (subst \<mu>G_eq) (auto simp: \<mu>G_eq PiP_finite intro: sets_PiM_I_finite)
   3.617 +    finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = emeasure (P J) (Pi\<^isub>E J X)" .
   3.618 +  next
   3.619 +    show "emeasure (P J) (Pi\<^isub>E J B) = emeasure (PiP J (\<lambda>_. borel) P) (Pi\<^isub>E J B)"
   3.620 +      using assms by (simp add: f_def PiP_finite Pi_def)
   3.621 +  qed
   3.622 +qed
   3.623 +
   3.624 +end
   3.625 +
   3.626 +sublocale polish_projective \<subseteq> P!: prob_space "(PiB I P)"
   3.627 +proof
   3.628 +  show "emeasure (PiB I P) (space (PiB I P)) = 1"
   3.629 +  proof cases
   3.630 +    assume "I = {}"
   3.631 +    interpret prob_space "P {}" using prob_space by simp
   3.632 +    show ?thesis
   3.633 +      by (simp add: space_PiM_empty PiP_finite emeasure_space_1 `I = {}`)
   3.634 +  next
   3.635 +    assume "I \<noteq> {}"
   3.636 +    then obtain i where "i \<in> I" by auto
   3.637 +    interpret prob_space "P {i}" using prob_space by simp
   3.638 +    have R: "(space (PiB I P)) = (emb I {i} (Pi\<^isub>E {i} (\<lambda>_. space borel)))"
   3.639 +      by (auto simp: prod_emb_def space_PiM)
   3.640 +    moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM)
   3.641 +    ultimately show ?thesis using `i \<in> I`
   3.642 +      apply (subst R)
   3.643 +      apply (subst emeasure_PiB_emb_not_empty)
   3.644 +      apply (auto simp: PiP_finite emeasure_space_1)
   3.645 +      done
   3.646 +  qed
   3.647 +qed
   3.648 +
   3.649 +context polish_projective begin
   3.650 +
   3.651 +lemma emeasure_PiB_emb:
   3.652 +  assumes X: "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
   3.653 +  shows "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (P J) (Pi\<^isub>E J B)"
   3.654 +proof cases
   3.655 +  interpret prob_space "P {}" using prob_space by simp
   3.656 +  assume "J = {}"
   3.657 +  moreover have "emb I {} {\<lambda>x. undefined} = space (PiB I P)"
   3.658 +    by (auto simp: space_PiM prod_emb_def)
   3.659 +  moreover have "{\<lambda>x. undefined} = space (PiB {} P)"
   3.660 +    by (auto simp: space_PiM prod_emb_def)
   3.661 +  ultimately show ?thesis
   3.662 +    by (simp add: P.emeasure_space_1 PiP_finite emeasure_space_1 del: space_PiP)
   3.663 +next
   3.664 +  assume "J \<noteq> {}" with X show ?thesis
   3.665 +    by (subst emeasure_PiB_emb_not_empty) (auto simp: PiP_finite)
   3.666 +qed
   3.667 +
   3.668 +lemma measure_PiB_emb:
   3.669 +  assumes "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
   3.670 +  shows "measure (PiB I P) (emb I J (Pi\<^isub>E J B)) = measure (P J) (Pi\<^isub>E J B)"
   3.671 +proof -
   3.672 +  interpret prob_space "P J" using prob_space assms by simp
   3.673 +  show ?thesis
   3.674 +    using emeasure_PiB_emb[OF assms]
   3.675 +    unfolding emeasure_eq_measure PiP_finite[OF `finite J` `J \<subseteq> I`] P.emeasure_eq_measure
   3.676 +    by simp
   3.677 +qed
   3.678 +
   3.679 +end
   3.680 +
   3.681 +locale polish_product_prob_space =
   3.682 +  product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
   3.683 +
   3.684 +sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
   3.685 +proof qed
   3.686 +
   3.687 +lemma (in polish_product_prob_space)
   3.688 +  PiP_eq_PiM:
   3.689 +  "I \<noteq> {} \<Longrightarrow> PiP I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
   3.690 +    PiM I (\<lambda>_. borel)"
   3.691 +  by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_PiB_emb)
   3.692 +
   3.693 +end