--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Fin_Map.thy Thu Nov 15 11:16:58 2012 +0100
@@ -0,0 +1,1503 @@
+(* Title: HOL/Probability/Projective_Family.thy
+ Author: Fabian Immler, TU München
+*)
+
+theory Fin_Map
+imports Finite_Product_Measure
+begin
+
+section {* Finite Maps *}
+
+text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of
+ projective limit. @{const extensional} functions are used for the representation in order to
+ stay close to the developments of (finite) products @{const Pi\<^isub>E} and their sigma-algebra
+ @{const Pi\<^isub>M}. *}
+
+typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^isub>F /_)" [22, 21] 21) =
+ "{(I::'i set, f::'i \<Rightarrow> 'a). finite I \<and> f \<in> extensional I}" by auto
+
+subsection {* Domain and Application *}
+
+definition domain where "domain P = fst (Rep_finmap P)"
+
+lemma finite_domain[simp, intro]: "finite (domain P)"
+ by (cases P) (auto simp: domain_def Abs_finmap_inverse)
+
+definition proj ("_\<^isub>F" [1000] 1000) where "proj P i = snd (Rep_finmap P) i"
+
+declare [[coercion proj]]
+
+lemma extensional_proj[simp, intro]: "(P)\<^isub>F \<in> extensional (domain P)"
+ by (cases P) (auto simp: domain_def Abs_finmap_inverse proj_def[abs_def])
+
+lemma proj_undefined[simp, intro]: "i \<notin> domain P \<Longrightarrow> P i = undefined"
+ using extensional_proj[of P] unfolding extensional_def by auto
+
+lemma finmap_eq_iff: "P = Q \<longleftrightarrow> (domain P = domain Q \<and> (\<forall>i\<in>domain P. P i = Q i))"
+ by (cases P, cases Q)
+ (auto simp add: Abs_finmap_inject extensional_def domain_def proj_def Abs_finmap_inverse
+ intro: extensionalityI)
+
+subsection {* Countable Finite Maps *}
+
+instance finmap :: (countable, countable) countable
+proof
+ obtain mapper where mapper: "\<And>fm :: 'a \<Rightarrow>\<^isub>F 'b. set (mapper fm) = domain fm"
+ by (metis finite_list[OF finite_domain])
+ have "inj (\<lambda>fm. map (\<lambda>i. (i, (fm)\<^isub>F i)) (mapper fm))" (is "inj ?F")
+ proof (rule inj_onI)
+ fix f1 f2 assume "?F f1 = ?F f2"
+ then have "map fst (?F f1) = map fst (?F f2)" by simp
+ then have "mapper f1 = mapper f2" by (simp add: comp_def)
+ then have "domain f1 = domain f2" by (simp add: mapper[symmetric])
+ with `?F f1 = ?F f2` show "f1 = f2"
+ unfolding `mapper f1 = mapper f2` map_eq_conv mapper
+ by (simp add: finmap_eq_iff)
+ qed
+ then show "\<exists>to_nat :: 'a \<Rightarrow>\<^isub>F 'b \<Rightarrow> nat. inj to_nat"
+ by (intro exI[of _ "to_nat \<circ> ?F"] inj_comp) auto
+qed
+
+subsection {* Constructor of Finite Maps *}
+
+definition "finmap_of inds f = Abs_finmap (inds, restrict f inds)"
+
+lemma proj_finmap_of[simp]:
+ assumes "finite inds"
+ shows "(finmap_of inds f)\<^isub>F = restrict f inds"
+ using assms
+ by (auto simp: Abs_finmap_inverse finmap_of_def proj_def)
+
+lemma domain_finmap_of[simp]:
+ assumes "finite inds"
+ shows "domain (finmap_of inds f) = inds"
+ using assms
+ by (auto simp: Abs_finmap_inverse finmap_of_def domain_def)
+
+lemma finmap_of_eq_iff[simp]:
+ assumes "finite i" "finite j"
+ shows "finmap_of i m = finmap_of j n \<longleftrightarrow> i = j \<and> restrict m i = restrict n i"
+ using assms
+ apply (auto simp: finmap_eq_iff restrict_def) by metis
+
+lemma
+ finmap_of_inj_on_extensional_finite:
+ assumes "finite K"
+ assumes "S \<subseteq> extensional K"
+ shows "inj_on (finmap_of K) S"
+proof (rule inj_onI)
+ fix x y::"'a \<Rightarrow> 'b"
+ assume "finmap_of K x = finmap_of K y"
+ hence "(finmap_of K x)\<^isub>F = (finmap_of K y)\<^isub>F" by simp
+ moreover
+ assume "x \<in> S" "y \<in> S" hence "x \<in> extensional K" "y \<in> extensional K" using assms by auto
+ ultimately
+ show "x = y" using assms by (simp add: extensional_restrict)
+qed
+
+lemma finmap_choice:
+ assumes *: "\<And>i. i \<in> I \<Longrightarrow> \<exists>x. P i x" and I: "finite I"
+ shows "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. P i (fm i))"
+proof -
+ have "\<exists>f. \<forall>i\<in>I. P i (f i)"
+ unfolding bchoice_iff[symmetric] using * by auto
+ then guess f ..
+ with I show ?thesis
+ by (intro exI[of _ "finmap_of I f"]) auto
+qed
+
+subsection {* Product set of Finite Maps *}
+
+text {* This is @{term Pi} for Finite Maps, most of this is copied *}
+
+definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^isub>F 'a) set" where
+ "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^isub>F i \<in> A i) } "
+
+syntax
+ "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PI' _:_./ _)" 10)
+
+syntax (xsymbols)
+ "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>' _\<in>_./ _)" 10)
+
+syntax (HTML output)
+ "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>' _\<in>_./ _)" 10)
+
+translations
+ "PI' x:A. B" == "CONST Pi' A (%x. B)"
+
+abbreviation
+ finmapset :: "['a set, 'b set] => ('a \<Rightarrow>\<^isub>F 'b) set"
+ (infixr "~>" 60) where
+ "A ~> B \<equiv> Pi' A (%_. B)"
+
+notation (xsymbols)
+ finmapset (infixr "\<leadsto>" 60)
+
+subsubsection{*Basic Properties of @{term Pi'}*}
+
+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"
+ by (simp add: Pi'_def)
+
+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"
+ by (simp add:Pi'_def)
+
+lemma finmapsetI: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> f \<in> A \<leadsto> B"
+ by (simp add: Pi_def)
+
+lemma Pi'_mem: "f\<in> Pi' A B \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<in> B x"
+ by (simp add: Pi'_def)
+
+lemma Pi'_iff: "f \<in> Pi' I X \<longleftrightarrow> domain f = I \<and> (\<forall>i\<in>I. f i \<in> X i)"
+ unfolding Pi'_def by auto
+
+lemma Pi'E [elim]:
+ "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"
+ by(auto simp: Pi'_def)
+
+lemma in_Pi'_cong:
+ "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"
+ by (auto simp: Pi'_def)
+
+lemma funcset_mem: "[|f \<in> A \<leadsto> B; x \<in> A|] ==> f x \<in> B"
+ by (simp add: Pi'_def)
+
+lemma funcset_image: "f \<in> A \<leadsto> B ==> f ` A \<subseteq> B"
+by auto
+
+lemma Pi'_eq_empty[simp]:
+ assumes "finite A" shows "(Pi' A B) = {} \<longleftrightarrow> (\<exists>x\<in>A. B x = {})"
+ using assms
+ apply (simp add: Pi'_def, auto)
+ apply (drule_tac x = "finmap_of A (\<lambda>u. SOME y. y \<in> B u)" in spec, auto)
+ apply (cut_tac P= "%y. y \<in> B i" in some_eq_ex, auto)
+ done
+
+lemma Pi'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> Pi' A B \<subseteq> Pi' A C"
+ by (auto simp: Pi'_def)
+
+lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^isub>E A B) = proj ` Pi' A B"
+ apply (auto simp: Pi'_def Pi_def extensional_def)
+ apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI)
+ apply auto
+ done
+
+subsection {* Metric Space of Finite Maps *}
+
+instantiation finmap :: (type, metric_space) metric_space
+begin
+
+definition dist_finmap where
+ "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
+ card ((domain P - domain Q) \<union> (domain Q - domain P))"
+
+lemma dist_finmap_extend:
+ assumes "finite X"
+ shows "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q \<union> X. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
+ card ((domain P - domain Q) \<union> (domain Q - domain P))"
+ unfolding dist_finmap_def add_right_cancel
+ using assms extensional_arb[of "(P)\<^isub>F"] extensional_arb[of "(Q)\<^isub>F" "domain Q"]
+ by (intro setsum_mono_zero_cong_left) auto
+
+definition open_finmap :: "('a \<Rightarrow>\<^isub>F 'b) set \<Rightarrow> bool" where
+ "open_finmap S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
+
+lemma add_eq_zero_iff[simp]:
+ fixes a b::real
+ assumes "a \<ge> 0" "b \<ge> 0"
+ shows "a + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
+using assms by auto
+
+lemma dist_le_1_imp_domain_eq:
+ assumes "dist P Q < 1"
+ shows "domain P = domain Q"
+proof -
+ have "0 \<le> (\<Sum>i\<in>domain P \<union> domain Q. dist (P i) (Q i))"
+ by (simp add: setsum_nonneg)
+ with assms have "card (domain P - domain Q \<union> (domain Q - domain P)) = 0"
+ unfolding dist_finmap_def by arith
+ thus "domain P = domain Q" by auto
+qed
+
+lemma dist_proj:
+ shows "dist ((x)\<^isub>F i) ((y)\<^isub>F i) \<le> dist x y"
+proof -
+ have "dist (x i) (y i) = (\<Sum>i\<in>{i}. dist (x i) (y i))" by simp
+ also have "\<dots> \<le> (\<Sum>i\<in>domain x \<union> domain y \<union> {i}. dist (x i) (y i))"
+ by (intro setsum_mono2) auto
+ also have "\<dots> \<le> dist x y" by (simp add: dist_finmap_extend[of "{i}"])
+ finally show ?thesis by simp
+qed
+
+lemma open_Pi'I:
+ assumes open_component: "\<And>i. i \<in> I \<Longrightarrow> open (A i)"
+ shows "open (Pi' I A)"
+proof (subst open_finmap_def, safe)
+ fix x assume x: "x \<in> Pi' I A"
+ hence dim_x: "domain x = I" by (simp add: Pi'_def)
+ hence [simp]: "finite I" unfolding dim_x[symmetric] by simp
+ have "\<exists>ei. \<forall>i\<in>I. 0 < ei i \<and> (\<forall>y. dist y (x i) < ei i \<longrightarrow> y \<in> A i)"
+ proof (safe intro!: bchoice)
+ fix i assume i: "i \<in> I"
+ moreover with open_component have "open (A i)" by simp
+ moreover have "x i \<in> A i" using x i
+ by (auto simp: proj_def)
+ ultimately show "\<exists>e>0. \<forall>y. dist y (x i) < e \<longrightarrow> y \<in> A i"
+ using x by (auto simp: open_dist Ball_def)
+ qed
+ then guess ei .. note ei = this
+ def es \<equiv> "ei ` I"
+ def e \<equiv> "if es = {} then 0.5 else min 0.5 (Min es)"
+ from ei have "e > 0" using x
+ by (auto simp add: e_def es_def Pi'_def Ball_def)
+ moreover have "\<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A"
+ proof (intro allI impI)
+ fix y
+ assume "dist y x < e"
+ also have "\<dots> < 1" by (auto simp: e_def)
+ finally have "domain y = domain x" by (rule dist_le_1_imp_domain_eq)
+ with dim_x have dims: "domain y = domain x" "domain x = I" by auto
+ show "y \<in> Pi' I A"
+ proof
+ show "domain y = I" using dims by simp
+ next
+ fix i
+ assume "i \<in> I"
+ have "dist (y i) (x i) \<le> dist y x" using dims `i \<in> I`
+ by (auto intro: dist_proj)
+ also have "\<dots> < e" using `dist y x < e` dims
+ by (simp add: dist_finmap_def)
+ also have "e \<le> Min (ei ` I)" using dims `i \<in> I`
+ by (auto simp: e_def es_def)
+ also have "\<dots> \<le> ei i" using `i \<in> I` by (simp add: e_def)
+ finally have "dist (y i) (x i) < ei i" .
+ with ei `i \<in> I` show "y i \<in> A i" by simp
+ qed
+ qed
+ ultimately
+ show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A" by blast
+qed
+
+instance
+proof
+ fix S::"('a \<Rightarrow>\<^isub>F 'b) set"
+ show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
+ unfolding open_finmap_def ..
+next
+ fix P Q::"'a \<Rightarrow>\<^isub>F 'b"
+ show "dist P Q = 0 \<longleftrightarrow> P = Q"
+ by (auto simp: finmap_eq_iff dist_finmap_def setsum_nonneg setsum_nonneg_eq_0_iff)
+next
+ fix P Q R::"'a \<Rightarrow>\<^isub>F 'b"
+ let ?symdiff = "\<lambda>a b. domain a - domain b \<union> (domain b - domain a)"
+ def E \<equiv> "domain P \<union> domain Q \<union> domain R"
+ hence "finite E" by (simp add: E_def)
+ have "card (?symdiff P Q) \<le> card (?symdiff P R \<union> ?symdiff Q R)"
+ by (auto intro: card_mono)
+ also have "\<dots> \<le> card (?symdiff P R) + card (?symdiff Q R)"
+ by (subst card_Un_Int) auto
+ finally have "dist P Q \<le> (\<Sum>i\<in>E. dist (P i) (R i) + dist (Q i) (R i)) +
+ real (card (?symdiff P R) + card (?symdiff Q R))"
+ unfolding dist_finmap_extend[OF `finite E`]
+ by (intro add_mono) (auto simp: E_def intro: setsum_mono dist_triangle_le)
+ also have "\<dots> \<le> dist P R + dist Q R"
+ unfolding dist_finmap_extend[OF `finite E`] by (simp add: ac_simps E_def setsum_addf[symmetric])
+ finally show "dist P Q \<le> dist P R + dist Q R" by simp
+qed
+
+end
+
+lemma open_restricted_space:
+ shows "open {m. P (domain m)}"
+proof -
+ have "{m. P (domain m)} = (\<Union>i \<in> Collect P. {m. domain m = i})" by auto
+ also have "open \<dots>"
+ proof (rule, safe, cases)
+ fix i::"'a set"
+ assume "finite i"
+ hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
+ also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
+ finally show "open {m. domain m = i}" .
+ next
+ fix i::"'a set"
+ assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
+ also have "open \<dots>" by simp
+ finally show "open {m. domain m = i}" .
+ qed
+ finally show ?thesis .
+qed
+
+lemma closed_restricted_space:
+ shows "closed {m. P (domain m)}"
+proof -
+ have "{m. P (domain m)} = - (\<Union>i \<in> - Collect P. {m. domain m = i})" by auto
+ also have "closed \<dots>"
+ proof (rule, rule, rule, cases)
+ fix i::"'a set"
+ assume "finite i"
+ hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
+ also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
+ finally show "open {m. domain m = i}" .
+ next
+ fix i::"'a set"
+ assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
+ also have "open \<dots>" by simp
+ finally show "open {m. domain m = i}" .
+ qed
+ finally show ?thesis .
+qed
+
+lemma continuous_proj:
+ shows "continuous_on s (\<lambda>x. (x)\<^isub>F i)"
+ unfolding continuous_on_topological
+proof safe
+ fix x B assume "x \<in> s" "open B" "x i \<in> B"
+ let ?A = "Pi' (domain x) (\<lambda>j. if i = j then B else UNIV)"
+ have "open ?A" using `open B` by (auto intro: open_Pi'I)
+ moreover have "x \<in> ?A" using `x i \<in> B` by auto
+ moreover have "(\<forall>y\<in>s. y \<in> ?A \<longrightarrow> y i \<in> B)"
+ proof (cases, safe)
+ fix y assume "y \<in> s"
+ assume "i \<notin> domain x" hence "undefined \<in> B" using `x i \<in> B`
+ by simp
+ moreover
+ assume "y \<in> ?A" hence "domain y = domain x" by (simp add: Pi'_def)
+ hence "y i = undefined" using `i \<notin> domain x` by simp
+ ultimately
+ show "y i \<in> B" by simp
+ qed force
+ ultimately
+ show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> y i \<in> B)" by blast
+qed
+
+subsection {* Complete Space of Finite Maps *}
+
+lemma tendsto_dist_zero:
+ assumes "(\<lambda>i. dist (f i) g) ----> 0"
+ shows "f ----> g"
+ using assms by (auto simp: tendsto_iff dist_real_def)
+
+lemma tendsto_dist_zero':
+ assumes "(\<lambda>i. dist (f i) g) ----> x"
+ assumes "0 = x"
+ shows "f ----> g"
+ using assms tendsto_dist_zero by simp
+
+lemma tendsto_finmap:
+ fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^isub>F ('a::metric_space))"
+ assumes ind_f: "\<And>n. domain (f n) = domain g"
+ assumes proj_g: "\<And>i. i \<in> domain g \<Longrightarrow> (\<lambda>n. (f n) i) ----> g i"
+ shows "f ----> g"
+ apply (rule tendsto_dist_zero')
+ unfolding dist_finmap_def assms
+ apply (rule tendsto_intros proj_g | simp)+
+ done
+
+instance finmap :: (type, complete_space) complete_space
+proof
+ fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^isub>F 'b"
+ assume "Cauchy P"
+ then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
+ by (force simp: cauchy)
+ def d \<equiv> "domain (P Nd)"
+ with Nd have dim: "\<And>n. n \<ge> Nd \<Longrightarrow> domain (P n) = d" using dist_le_1_imp_domain_eq by auto
+ have [simp]: "finite d" unfolding d_def by simp
+ def p \<equiv> "\<lambda>i n. (P n) i"
+ def q \<equiv> "\<lambda>i. lim (p i)"
+ def Q \<equiv> "finmap_of d q"
+ have q: "\<And>i. i \<in> d \<Longrightarrow> q i = Q i" by (auto simp add: Q_def Abs_finmap_inverse)
+ {
+ fix i assume "i \<in> d"
+ have "Cauchy (p i)" unfolding cauchy p_def
+ proof safe
+ fix e::real assume "0 < e"
+ with `Cauchy P` obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
+ by (force simp: cauchy min_def)
+ hence "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto
+ with dim have dim: "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = d" by (metis nat_le_linear)
+ show "\<exists>N. \<forall>n\<ge>N. dist ((P n) i) ((P N) i) < e"
+ proof (safe intro!: exI[where x="N"])
+ fix n assume "N \<le> n" have "N \<le> N" by simp
+ have "dist ((P n) i) ((P N) i) \<le> dist (P n) (P N)"
+ using dim[OF `N \<le> n`] dim[OF `N \<le> N`] `i \<in> d`
+ by (auto intro!: dist_proj)
+ also have "\<dots> < e" using N[OF `N \<le> n`] by simp
+ finally show "dist ((P n) i) ((P N) i) < e" .
+ qed
+ qed
+ hence "convergent (p i)" by (metis Cauchy_convergent_iff)
+ hence "p i ----> q i" unfolding q_def convergent_def by (metis limI)
+ } note p = this
+ have "P ----> Q"
+ proof (rule metric_LIMSEQ_I)
+ fix e::real assume "0 < e"
+ def e' \<equiv> "min 1 (e / (card d + 1))"
+ hence "0 < e'" using `0 < e` by (auto simp: e'_def intro: divide_pos_pos)
+ have "\<exists>ni. \<forall>i\<in>d. \<forall>n\<ge>ni i. dist (p i n) (q i) < e'"
+ proof (safe intro!: bchoice)
+ fix i assume "i \<in> d"
+ from p[OF `i \<in> d`, THEN metric_LIMSEQ_D, OF `0 < e'`]
+ show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e'" .
+ qed then guess ni .. note ni = this
+ def N \<equiv> "max Nd (Max (ni ` d))"
+ show "\<exists>N. \<forall>n\<ge>N. dist (P n) Q < e"
+ proof (safe intro!: exI[where x="N"])
+ fix n assume "N \<le> n"
+ hence "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q"
+ using dim by (simp_all add: N_def Q_def dim_def Abs_finmap_inverse)
+ hence "dist (P n) Q = (\<Sum>i\<in>d. dist ((P n) i) (Q i))" by (simp add: dist_finmap_def)
+ also have "\<dots> \<le> (\<Sum>i\<in>d. e')"
+ proof (intro setsum_mono less_imp_le)
+ fix i assume "i \<in> d"
+ hence "ni i \<le> Max (ni ` d)" by simp
+ also have "\<dots> \<le> N" by (simp add: N_def)
+ also have "\<dots> \<le> n" using `N \<le> n` .
+ finally
+ show "dist ((P n) i) (Q i) < e'"
+ using ni `i \<in> d` by (auto simp: p_def q N_def)
+ qed
+ also have "\<dots> = card d * e'" by (simp add: real_eq_of_nat)
+ also have "\<dots> < e" using `0 < e` by (simp add: e'_def field_simps min_def)
+ finally show "dist (P n) Q < e" .
+ qed
+ qed
+ thus "convergent P" by (auto simp: convergent_def)
+qed
+
+subsection {* Polish Space of Finite Maps *}
+
+instantiation finmap :: (countable, polish_space) polish_space
+begin
+
+definition enum_basis_finmap :: "nat \<Rightarrow> ('a \<Rightarrow>\<^isub>F 'b) set" where
+ "enum_basis_finmap n =
+ (let m = from_nat n::('a \<Rightarrow>\<^isub>F nat) in Pi' (domain m) (enum_basis o (m)\<^isub>F))"
+
+lemma range_enum_basis_eq:
+ "range enum_basis_finmap = {Pi' I S|I S. finite I \<and> (\<forall>i \<in> I. S i \<in> range enum_basis)}"
+proof (auto simp: enum_basis_finmap_def[abs_def])
+ fix S::"('a \<Rightarrow> 'b set)" and I
+ assume "\<forall>i\<in>I. S i \<in> range enum_basis"
+ hence "\<forall>i\<in>I. \<exists>n. S i = enum_basis n" by auto
+ then obtain n where n: "\<forall>i\<in>I. S i = enum_basis (n i)"
+ unfolding bchoice_iff by blast
+ assume [simp]: "finite I"
+ have "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. n i = (fm i))"
+ by (rule finmap_choice) auto
+ then obtain m where "Pi' I S = Pi' (domain m) (enum_basis o m)"
+ using n by (auto simp: Pi'_def)
+ hence "Pi' I S = (let m = from_nat (to_nat m) in Pi' (domain m) (enum_basis \<circ> m))"
+ by simp
+ thus "Pi' I S \<in> range (\<lambda>n. let m = from_nat n in Pi' (domain m) (enum_basis \<circ> m))"
+ by blast
+qed (metis finite_domain o_apply rangeI)
+
+lemma in_enum_basis_finmapI:
+ assumes "finite I" assumes "\<And>i. i \<in> I \<Longrightarrow> S i \<in> range enum_basis"
+ shows "Pi' I S \<in> range enum_basis_finmap"
+ using assms unfolding range_enum_basis_eq by auto
+
+lemma finmap_topological_basis:
+ "topological_basis (range (enum_basis_finmap))"
+proof (subst topological_basis_iff, safe)
+ fix n::nat
+ show "open (enum_basis_finmap n::('a \<Rightarrow>\<^isub>F 'b) set)" using enumerable_basis
+ by (auto intro!: open_Pi'I simp: topological_basis_def enum_basis_finmap_def Let_def)
+next
+ fix O'::"('a \<Rightarrow>\<^isub>F 'b) set" and x
+ assume "open O'" "x \<in> O'"
+ then obtain e where e: "e > 0" "\<And>y. dist y x < e \<Longrightarrow> y \<in> O'" unfolding open_dist by blast
+ def e' \<equiv> "e / (card (domain x) + 1)"
+
+ have "\<exists>B.
+ (\<forall>i\<in>domain x. x i \<in> enum_basis (B i) \<and> enum_basis (B i) \<subseteq> ball (x i) e')"
+ proof (rule bchoice, safe)
+ fix i assume "i \<in> domain x"
+ have "open (ball (x i) e')" "x i \<in> ball (x i) e'" using e
+ by (auto simp add: e'_def intro!: divide_pos_pos)
+ from enumerable_basisE[OF this] guess b' .
+ thus "\<exists>y. x i \<in> enum_basis y \<and>
+ enum_basis y \<subseteq> ball (x i) e'" by auto
+ qed
+ then guess B .. note B = this
+ def B' \<equiv> "Pi' (domain x) (\<lambda>i. enum_basis (B i)::'b set)"
+ hence "B' \<in> range enum_basis_finmap" unfolding B'_def
+ by (intro in_enum_basis_finmapI) auto
+ moreover have "x \<in> B'" unfolding B'_def using B by auto
+ moreover have "B' \<subseteq> O'"
+ proof
+ fix y assume "y \<in> B'" with B have "domain y = domain x" unfolding B'_def
+ by (simp add: Pi'_def)
+ show "y \<in> O'"
+ proof (rule e)
+ have "dist y x = (\<Sum>i \<in> domain x. dist (y i) (x i))"
+ using `domain y = domain x` by (simp add: dist_finmap_def)
+ also have "\<dots> \<le> (\<Sum>i \<in> domain x. e')"
+ proof (rule setsum_mono)
+ fix i assume "i \<in> domain x"
+ with `y \<in> B'` B have "y i \<in> enum_basis (B i)"
+ by (simp add: Pi'_def B'_def)
+ hence "y i \<in> ball (x i) e'" using B `domain y = domain x` `i \<in> domain x`
+ by force
+ thus "dist (y i) (x i) \<le> e'" by (simp add: dist_commute)
+ qed
+ also have "\<dots> = card (domain x) * e'" by (simp add: real_eq_of_nat)
+ also have "\<dots> < e" using e by (simp add: e'_def field_simps)
+ finally show "dist y x < e" .
+ qed
+ qed
+ ultimately
+ show "\<exists>B'\<in>range enum_basis_finmap. x \<in> B' \<and> B' \<subseteq> O'" by blast
+qed
+
+lemma range_enum_basis_finmap_imp_open:
+ assumes "x \<in> range enum_basis_finmap"
+ shows "open x"
+ using finmap_topological_basis assms by (auto simp: topological_basis_def)
+
+lemma
+ open_imp_ex_UNION_of_enum:
+ fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
+ assumes "open X" assumes "X \<noteq> {}"
+ 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>
+ (\<forall>n. \<forall>i\<in>A n. (B n) i \<in> range enum_basis) \<and> (\<forall>n. finite (A n))"
+proof -
+ from `open X` obtain B' where B': "B'\<subseteq>range enum_basis_finmap" "\<Union>B' = X"
+ using finmap_topological_basis by (force simp add: topological_basis_def)
+ then obtain B where B: "B' = enum_basis_finmap ` B" by (auto simp: subset_image_iff)
+ show ?thesis
+ proof cases
+ assume "B = {}" with B have "B' = {}" by simp hence False using B' assms by simp
+ thus ?thesis by simp
+ next
+ assume "B \<noteq> {}" then obtain b where b: "b \<in> B" by auto
+ def NA \<equiv> "\<lambda>n::nat. if n \<in> B
+ then domain ((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n)
+ else domain ((from_nat::_\<Rightarrow>'a\<Rightarrow>\<^isub>F nat) b)"
+ def NB \<equiv> "\<lambda>n::nat. if n \<in> B
+ then (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n) i))
+ else (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) b) i))"
+ have "X = UNION UNIV (\<lambda>i. Pi' (NA i) (NB i))" unfolding B'(2)[symmetric] using b
+ unfolding B
+ by safe
+ (auto simp add: NA_def NB_def enum_basis_finmap_def Let_def o_def split: split_if_asm)
+ moreover
+ have "(\<forall>n. \<forall>i\<in>NA n. (NB n) i \<in> range enum_basis)"
+ using enumerable_basis by (auto simp: topological_basis_def NA_def NB_def)
+ moreover have "(\<forall>n. finite (NA n))" by (simp add: NA_def)
+ ultimately show ?thesis by auto
+ qed
+qed
+
+lemma
+ open_imp_ex_UNION:
+ fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
+ assumes "open X" assumes "X \<noteq> {}"
+ 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>
+ (\<forall>n. \<forall>i\<in>A n. open ((B n) i)) \<and> (\<forall>n. finite (A n))"
+ using open_imp_ex_UNION_of_enum[OF assms]
+ apply auto
+ apply (rule_tac x = A in exI)
+ apply (rule_tac x = B in exI)
+ apply (auto simp: open_enum_basis)
+ done
+
+lemma
+ open_basisE:
+ assumes "open X" assumes "X \<noteq> {}"
+ obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
+ "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)"
+using open_imp_ex_UNION[OF assms] by auto
+
+lemma
+ open_basis_of_enumE:
+ assumes "open X" assumes "X \<noteq> {}"
+ obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
+ "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"
+ "\<And>n. finite (A n)"
+using open_imp_ex_UNION_of_enum[OF assms] by auto
+
+instance proof qed (blast intro: finmap_topological_basis)
+
+end
+
+subsection {* Product Measurable Space of Finite Maps *}
+
+definition "PiF I M \<equiv>
+ sigma
+ (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
+ {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
+
+abbreviation
+ "Pi\<^isub>F I M \<equiv> PiF I M"
+
+syntax
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3PIF _:_./ _)" 10)
+
+syntax (xsymbols)
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>F _\<in>_./ _)" 10)
+
+syntax (HTML output)
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>F _\<in>_./ _)" 10)
+
+translations
+ "PIF x:I. M" == "CONST PiF I (%x. M)"
+
+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>
+ Pow (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
+ by (auto simp: Pi'_def) (blast dest: sets_into_space)
+
+lemma space_PiF: "space (PiF I M) = (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
+ unfolding PiF_def using PiF_gen_subset by (rule space_measure_of)
+
+lemma sets_PiF:
+ "sets (PiF I M) = sigma_sets (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
+ {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
+ unfolding PiF_def using PiF_gen_subset by (rule sets_measure_of)
+
+lemma sets_PiF_singleton:
+ "sets (PiF {I} M) = sigma_sets (\<Pi>' j\<in>I. space (M j))
+ {(\<Pi>' j\<in>I. X j) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
+ unfolding sets_PiF by simp
+
+lemma in_sets_PiFI:
+ assumes "X = (Pi' J S)" "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
+ shows "X \<in> sets (PiF I M)"
+ unfolding sets_PiF
+ using assms by blast
+
+lemma product_in_sets_PiFI:
+ assumes "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
+ shows "(Pi' J S) \<in> sets (PiF I M)"
+ unfolding sets_PiF
+ using assms by blast
+
+lemma singleton_space_subset_in_sets:
+ fixes J
+ assumes "J \<in> I"
+ assumes "finite J"
+ shows "space (PiF {J} M) \<in> sets (PiF I M)"
+ using assms
+ by (intro in_sets_PiFI[where J=J and S="\<lambda>i. space (M i)"])
+ (auto simp: product_def space_PiF)
+
+lemma singleton_subspace_set_in_sets:
+ assumes A: "A \<in> sets (PiF {J} M)"
+ assumes "finite J"
+ assumes "J \<in> I"
+ shows "A \<in> sets (PiF I M)"
+ using A[unfolded sets_PiF]
+ apply (induct A)
+ unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
+ using assms
+ by (auto intro: in_sets_PiFI intro!: singleton_space_subset_in_sets)
+
+lemma
+ finite_measurable_singletonI:
+ assumes "finite I"
+ assumes "\<And>J. J \<in> I \<Longrightarrow> finite J"
+ assumes MN: "\<And>J. J \<in> I \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
+ shows "A \<in> measurable (PiF I M) N"
+ unfolding measurable_def
+proof safe
+ fix y assume "y \<in> sets N"
+ have "A -` y \<inter> space (PiF I M) = (\<Union>J\<in>I. A -` y \<inter> space (PiF {J} M))"
+ by (auto simp: space_PiF)
+ also have "\<dots> \<in> sets (PiF I M)"
+ proof
+ show "finite I" by fact
+ fix J assume "J \<in> I"
+ with assms have "finite J" by simp
+ show "A -` y \<inter> space (PiF {J} M) \<in> sets (PiF I M)"
+ by (rule singleton_subspace_set_in_sets[OF measurable_sets[OF assms(3)]]) fact+
+ qed
+ finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
+next
+ fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
+ using MN[of "domain x"]
+ by (auto simp: space_PiF measurable_space Pi'_def)
+qed
+
+lemma
+ countable_finite_comprehension:
+ fixes f :: "'a::countable set \<Rightarrow> _"
+ assumes "\<And>s. P s \<Longrightarrow> finite s"
+ assumes "\<And>s. P s \<Longrightarrow> f s \<in> sets M"
+ shows "\<Union>{f s|s. P s} \<in> sets M"
+proof -
+ have "\<Union>{f s|s. P s} = (\<Union>n::nat. let s = set (from_nat n) in if P s then f s else {})"
+ proof safe
+ fix x X s assume "x \<in> f s" "P s"
+ moreover with assms obtain l where "s = set l" using finite_list by blast
+ ultimately show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using `P s`
+ by (auto intro!: exI[where x="to_nat l"])
+ next
+ fix x n assume "x \<in> (let s = set (from_nat n) in if P s then f s else {})"
+ thus "x \<in> \<Union>{f s|s. P s}" using assms by (auto simp: Let_def split: split_if_asm)
+ qed
+ 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
+ also have "\<dots> \<in> sets M" using assms by (auto simp: Let_def)
+ finally show ?thesis .
+qed
+
+lemma space_subset_in_sets:
+ fixes J::"'a::countable set set"
+ assumes "J \<subseteq> I"
+ assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
+ shows "space (PiF J M) \<in> sets (PiF I M)"
+proof -
+ have "space (PiF J M) = \<Union>{space (PiF {j} M)|j. j \<in> J}"
+ unfolding space_PiF by blast
+ also have "\<dots> \<in> sets (PiF I M)" using assms
+ by (intro countable_finite_comprehension) (auto simp: singleton_space_subset_in_sets)
+ finally show ?thesis .
+qed
+
+lemma subspace_set_in_sets:
+ fixes J::"'a::countable set set"
+ assumes A: "A \<in> sets (PiF J M)"
+ assumes "J \<subseteq> I"
+ assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
+ shows "A \<in> sets (PiF I M)"
+ using A[unfolded sets_PiF]
+ apply (induct A)
+ unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
+ using assms
+ by (auto intro: in_sets_PiFI intro!: space_subset_in_sets)
+
+lemma
+ countable_measurable_PiFI:
+ fixes I::"'a::countable set set"
+ assumes MN: "\<And>J. J \<in> I \<Longrightarrow> finite J \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
+ shows "A \<in> measurable (PiF I M) N"
+ unfolding measurable_def
+proof safe
+ fix y assume "y \<in> sets N"
+ have "A -` y = (\<Union>{A -` y \<inter> {x. domain x = J}|J. finite J})" by auto
+ hence "A -` y \<inter> space (PiF I M) = (\<Union>n. A -` y \<inter> space (PiF ({set (from_nat n)}\<inter>I) M))"
+ apply (auto simp: space_PiF Pi'_def)
+ proof -
+ case goal1
+ from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto
+ thus ?case
+ apply (intro exI[where x="to_nat xs"])
+ apply auto
+ done
+ qed
+ also have "\<dots> \<in> sets (PiF I M)"
+ apply (intro Int countable_nat_UN subsetI, safe)
+ apply (case_tac "set (from_nat i) \<in> I")
+ apply simp_all
+ apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]])
+ using assms `y \<in> sets N`
+ apply (auto simp: space_PiF)
+ done
+ finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
+next
+ fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
+ using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def)
+qed
+
+lemma measurable_PiF:
+ 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))"
+ assumes S: "\<And>J S. J \<in> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> S i \<in> sets (M i)) \<Longrightarrow>
+ f -` (Pi' J S) \<inter> space N \<in> sets N"
+ shows "f \<in> measurable N (PiF I M)"
+ unfolding PiF_def
+ using PiF_gen_subset
+ apply (rule measurable_measure_of)
+ using f apply force
+ apply (insert S, auto)
+ done
+
+lemma
+ restrict_sets_measurable:
+ assumes A: "A \<in> sets (PiF I M)" and "J \<subseteq> I"
+ shows "A \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
+ using A[unfolded sets_PiF]
+ apply (induct A)
+ unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
+proof -
+ fix a assume "a \<in> {Pi' J X |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
+ then obtain K S where S: "a = Pi' K S" "K \<in> I" "(\<forall>i\<in>K. S i \<in> sets (M i))"
+ by auto
+ show "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
+ proof cases
+ assume "K \<in> J"
+ 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
+ by (auto intro!: exI[where x=K] exI[where x=S] simp: Pi'_def)
+ also have "\<dots> \<subseteq> sets (PiF J M)" unfolding sets_PiF by auto
+ finally show ?thesis .
+ next
+ assume "K \<notin> J"
+ hence "a \<inter> {m. domain m \<in> J} = {}" using S by (auto simp: Pi'_def)
+ also have "\<dots> \<in> sets (PiF J M)" by simp
+ finally show ?thesis .
+ qed
+next
+ show "{} \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" by simp
+next
+ fix a :: "nat \<Rightarrow> _"
+ assume a: "(\<And>i. a i \<inter> {m. domain m \<in> J} \<in> sets (PiF J M))"
+ have "UNION UNIV a \<inter> {m. domain m \<in> J} = (\<Union>i. (a i \<inter> {m. domain m \<in> J}))"
+ by simp
+ also have "\<dots> \<in> sets (PiF J M)" using a by (intro countable_nat_UN) auto
+ finally show "UNION UNIV a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
+next
+ fix a assume a: "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
+ have "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} = (space (PiF J M) - (a \<inter> {m. domain m \<in> J}))"
+ using `J \<subseteq> I` by (auto simp: space_PiF Pi'_def)
+ also have "\<dots> \<in> sets (PiF J M)" using a by auto
+ finally show "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
+qed
+
+lemma measurable_finmap_of:
+ assumes f: "\<And>i. (\<exists>x \<in> space N. i \<in> J x) \<Longrightarrow> (\<lambda>x. f x i) \<in> measurable N (M i)"
+ assumes J: "\<And>x. x \<in> space N \<Longrightarrow> J x \<in> I" "\<And>x. x \<in> space N \<Longrightarrow> finite (J x)"
+ assumes JN: "\<And>S. {x. J x = S} \<inter> space N \<in> sets N"
+ shows "(\<lambda>x. finmap_of (J x) (f x)) \<in> measurable N (PiF I M)"
+proof (rule measurable_PiF)
+ fix x assume "x \<in> space N"
+ with J[of x] measurable_space[OF f]
+ show "domain (finmap_of (J x) (f x)) \<in> I \<and>
+ (\<forall>i\<in>domain (finmap_of (J x) (f x)). (finmap_of (J x) (f x)) i \<in> space (M i))"
+ by auto
+next
+ fix K S assume "K \<in> I" and *: "\<And>i. i \<in> K \<Longrightarrow> S i \<in> sets (M i)"
+ with J have eq: "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N =
+ (if \<exists>x \<in> space N. K = J x \<and> finite K then if K = {} then {x \<in> space N. J x = K}
+ else (\<Inter>i\<in>K. (\<lambda>x. f x i) -` S i \<inter> {x \<in> space N. J x = K}) else {})"
+ by (auto simp: Pi'_def)
+ have r: "{x \<in> space N. J x = K} = space N \<inter> ({x. J x = K} \<inter> space N)" by auto
+ show "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N \<in> sets N"
+ unfolding eq r
+ apply (simp del: INT_simps add: )
+ apply (intro conjI impI finite_INT JN Int[OF top])
+ apply simp apply assumption
+ apply (subst Int_assoc[symmetric])
+ apply (rule Int)
+ apply (intro measurable_sets[OF f] *) apply force apply assumption
+ apply (intro JN)
+ done
+qed
+
+lemma measurable_PiM_finmap_of:
+ assumes "finite J"
+ shows "finmap_of J \<in> measurable (Pi\<^isub>M J M) (PiF {J} M)"
+ apply (rule measurable_finmap_of)
+ apply (rule measurable_component_singleton)
+ apply simp
+ apply rule
+ apply (rule `finite J`)
+ apply simp
+ done
+
+lemma proj_measurable_singleton:
+ assumes "A \<in> sets (M i)" "finite I"
+ shows "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
+proof cases
+ assume "i \<in> I"
+ hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
+ Pi' I (\<lambda>x. if x = i then A else space (M x))"
+ using sets_into_space[OF ] `A \<in> sets (M i)` assms
+ by (auto simp: space_PiF Pi'_def)
+ thus ?thesis using assms `A \<in> sets (M i)`
+ by (intro in_sets_PiFI) auto
+next
+ assume "i \<notin> I"
+ hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
+ (if undefined \<in> A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def)
+ thus ?thesis by simp
+qed
+
+lemma measurable_proj_singleton:
+ fixes I
+ assumes "finite I" "i \<in> I"
+ shows "(\<lambda>x. (x)\<^isub>F i) \<in> measurable (PiF {I} M) (M i)"
+proof (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
+qed (insert `i \<in> I`, auto simp: space_PiF)
+
+lemma measurable_proj_countable:
+ fixes I::"'a::countable set set"
+ assumes "y \<in> space (M i)"
+ shows "(\<lambda>x. if i \<in> domain x then (x)\<^isub>F i else y) \<in> measurable (PiF I M) (M i)"
+proof (rule countable_measurable_PiFI)
+ fix J assume "J \<in> I" "finite J"
+ show "(\<lambda>x. if i \<in> domain x then x i else y) \<in> measurable (PiF {J} M) (M i)"
+ unfolding measurable_def
+ proof safe
+ fix z assume "z \<in> sets (M i)"
+ have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
+ (\<lambda>x. if i \<in> J then (x)\<^isub>F i else y) -` z \<inter> space (PiF {J} M)"
+ by (auto simp: space_PiF Pi'_def)
+ also have "\<dots> \<in> sets (PiF {J} M)" using `z \<in> sets (M i)` `finite J`
+ by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
+ finally show "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) \<in>
+ sets (PiF {J} M)" .
+ qed (insert `y \<in> space (M i)`, auto simp: space_PiF Pi'_def)
+qed
+
+lemma measurable_restrict_proj:
+ assumes "J \<in> II" "finite J"
+ shows "finmap_of J \<in> measurable (PiM J M) (PiF II M)"
+ using assms
+ by (intro measurable_finmap_of measurable_component_singleton) auto
+
+lemma
+ measurable_proj_PiM:
+ fixes J K ::"'a::countable set" and I::"'a set set"
+ assumes "finite J" "J \<in> I"
+ assumes "x \<in> space (PiM J M)"
+ shows "proj \<in>
+ measurable (PiF {J} M) (PiM J M)"
+proof (rule measurable_PiM_single)
+ show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^isub>E i \<in> J. space (M i))"
+ using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def)
+next
+ fix A i assume A: "i \<in> J" "A \<in> sets (M i)"
+ show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} \<in> sets (PiF {J} M)"
+ proof
+ have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} =
+ (\<lambda>\<omega>. (\<omega>)\<^isub>F i) -` A \<inter> space (PiF {J} M)" by auto
+ also have "\<dots> \<in> sets (PiF {J} M)"
+ using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM)
+ finally show ?thesis .
+ qed simp
+qed
+
+lemma sets_subspaceI:
+ assumes "A \<inter> space M \<in> sets M"
+ assumes "B \<in> sets M"
+ shows "A \<inter> B \<in> sets M" using assms
+proof -
+ have "A \<inter> B = (A \<inter> space M) \<inter> B"
+ using assms sets_into_space by auto
+ thus ?thesis using assms by auto
+qed
+
+lemma space_PiF_singleton_eq_product:
+ assumes "finite I"
+ shows "space (PiF {I} M) = (\<Pi>' i\<in>I. space (M i))"
+ by (auto simp: product_def space_PiF assms)
+
+text {* adapted from @{thm sets_PiM_single} *}
+
+lemma sets_PiF_single:
+ assumes "finite I" "I \<noteq> {}"
+ shows "sets (PiF {I} M) =
+ sigma_sets (\<Pi>' i\<in>I. space (M i))
+ {{f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
+ (is "_ = sigma_sets ?\<Omega> ?R")
+ unfolding sets_PiF_singleton
+proof (rule sigma_sets_eqI)
+ interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
+ fix A assume "A \<in> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
+ then obtain X where X: "A = Pi' I X" "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
+ show "A \<in> sigma_sets ?\<Omega> ?R"
+ proof -
+ from `I \<noteq> {}` X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
+ using sets_into_space
+ by (auto simp: space_PiF product_def) blast
+ also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
+ using X `I \<noteq> {}` assms by (intro R.finite_INT) (auto simp: space_PiF)
+ finally show "A \<in> sigma_sets ?\<Omega> ?R" .
+ qed
+next
+ fix A assume "A \<in> ?R"
+ 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)"
+ by auto
+ then have "A = (\<Pi>' j \<in> I. if j = i then B else space (M j))"
+ using sets_into_space[OF A(3)]
+ apply (auto simp: Pi'_iff split: split_if_asm)
+ apply blast
+ done
+ also have "\<dots> \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
+ using A
+ by (intro sigma_sets.Basic )
+ (auto intro: exI[where x="\<lambda>j. if j = i then B else space (M j)"])
+ finally show "A \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" .
+qed
+
+text {* adapted from @{thm PiE_cong} *}
+
+lemma Pi'_cong:
+ assumes "finite I"
+ assumes "\<And>i. i \<in> I \<Longrightarrow> f i = g i"
+ shows "Pi' I f = Pi' I g"
+using assms by (auto simp: Pi'_def)
+
+text {* adapted from @{thm Pi_UN} *}
+
+lemma Pi'_UN:
+ fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
+ assumes "finite I"
+ assumes mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
+ shows "(\<Union>n. Pi' I (A n)) = Pi' I (\<lambda>i. \<Union>n. A n i)"
+proof (intro set_eqI iffI)
+ fix f assume "f \<in> Pi' I (\<lambda>i. \<Union>n. A n i)"
+ then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: `finite I` Pi'_def)
+ from bchoice[OF this(1)] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
+ obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
+ using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
+ have "f \<in> Pi' I (\<lambda>i. A k i)"
+ proof
+ fix i assume "i \<in> I"
+ from mono[OF this, of "n i" k] k[OF this] n[OF this] `domain f = I` `i \<in> I`
+ show "f i \<in> A k i " by (auto simp: `finite I`)
+ qed (simp add: `domain f = I` `finite I`)
+ then show "f \<in> (\<Union>n. Pi' I (A n))" by auto
+qed (auto simp: Pi'_def `finite I`)
+
+text {* adapted from @{thm sigma_prod_algebra_sigma_eq} *}
+
+lemma sigma_fprod_algebra_sigma_eq:
+ fixes E :: "'i \<Rightarrow> 'a set set"
+ assumes [simp]: "finite I" "I \<noteq> {}"
+ assumes S_mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
+ and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
+ and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
+ assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
+ and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
+ defines "P == { Pi' I F | F. \<forall>i\<in>I. F i \<in> E i }"
+ shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P"
+proof
+ let ?P = "sigma (space (Pi\<^isub>F {I} M)) P"
+ have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>F {I} M))"
+ using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq)
+ then have space_P: "space ?P = (\<Pi>' i\<in>I. space (M i))"
+ by (simp add: space_PiF)
+ have "sets (PiF {I} M) =
+ 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)}"
+ using sets_PiF_single[of I M] by (simp add: space_P)
+ also have "\<dots> \<subseteq> sets (sigma (space (PiF {I} M)) P)"
+ proof (safe intro!: sigma_sets_subset)
+ fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
+ have "(\<lambda>x. (x)\<^isub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
+ proof (subst measurable_iff_measure_of)
+ show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
+ from space_P `i \<in> I` show "(\<lambda>x. (x)\<^isub>F i) \<in> space ?P \<rightarrow> space (M i)"
+ by auto
+ show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ proof
+ fix A assume A: "A \<in> E i"
+ 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))"
+ using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
+ also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
+ by (intro Pi'_cong) (simp_all add: S_union)
+ also have "\<dots> = (\<Union>n. \<Pi>' j\<in>I. if i = j then A else S j n)"
+ using S_mono
+ by (subst Pi'_UN[symmetric, OF `finite I`]) (auto simp: incseq_def)
+ also have "\<dots> \<in> sets ?P"
+ proof (safe intro!: countable_UN)
+ fix n show "(\<Pi>' j\<in>I. if i = j then A else S j n) \<in> sets ?P"
+ using A S_in_E
+ by (simp add: P_closed)
+ (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j n"])
+ qed
+ finally show "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ using P_closed by simp
+ qed
+ qed
+ from measurable_sets[OF this, of A] A `i \<in> I` E_closed
+ have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ by (simp add: E_generates)
+ 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}"
+ using P_closed by (auto simp: space_PiF)
+ finally show "\<dots> \<in> sets ?P" .
+ qed
+ finally show "sets (PiF {I} M) \<subseteq> sigma_sets (space (PiF {I} M)) P"
+ by (simp add: P_closed)
+ show "sigma_sets (space (PiF {I} M)) P \<subseteq> sets (PiF {I} M)"
+ using `finite I` `I \<noteq> {}`
+ by (auto intro!: sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def)
+qed
+
+lemma enumerable_sigma_fprod_algebra_sigma_eq:
+ assumes "I \<noteq> {}"
+ assumes [simp]: "finite I"
+ shows "sets (PiF {I} (\<lambda>_. borel)) = sigma_sets (space (PiF {I} (\<lambda>_. borel)))
+ {Pi' I F |F. (\<forall>i\<in>I. F i \<in> range enum_basis)}"
+proof -
+ from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
+ show ?thesis
+ proof (rule sigma_fprod_algebra_sigma_eq)
+ show "finite I" by simp
+ show "I \<noteq> {}" by fact
+ show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
+ using S by simp_all
+ show "range enum_basis \<subseteq> Pow (space borel)" by simp
+ show "sets borel = sigma_sets (space borel) (range enum_basis)"
+ by (simp add: borel_eq_enum_basis)
+ qed
+qed
+
+text {* adapted from @{thm enumerable_sigma_fprod_algebra_sigma_eq} *}
+
+lemma enumerable_sigma_prod_algebra_sigma_eq:
+ assumes "I \<noteq> {}"
+ assumes [simp]: "finite I"
+ shows "sets (PiM I (\<lambda>_. borel)) = sigma_sets (space (PiM I (\<lambda>_. borel)))
+ {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range enum_basis}"
+proof -
+ from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
+ show ?thesis
+ proof (rule sigma_prod_algebra_sigma_eq)
+ show "finite I" by simp note[[show_types]]
+ fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
+ using S by simp_all
+ show "range enum_basis \<subseteq> Pow (space borel)" by simp
+ show "sets borel = sigma_sets (space borel) (range enum_basis)"
+ by (simp add: borel_eq_enum_basis)
+ qed
+qed
+
+lemma product_open_generates_sets_PiF_single:
+ assumes "I \<noteq> {}"
+ assumes [simp]: "finite I"
+ shows "sets (PiF {I} (\<lambda>_. borel::'b::enumerable_basis measure)) =
+ sigma_sets (space (PiF {I} (\<lambda>_. borel))) {Pi' I F |F. (\<forall>i\<in>I. F i \<in> Collect open)}"
+proof -
+ from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
+ show ?thesis
+ proof (rule sigma_fprod_algebra_sigma_eq)
+ show "finite I" by simp
+ show "I \<noteq> {}" by fact
+ show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
+ using S by (auto simp: open_enum_basis)
+ show "Collect open \<subseteq> Pow (space borel)" by simp
+ show "sets borel = sigma_sets (space borel) (Collect open)"
+ by (simp add: borel_def)
+ qed
+qed
+
+lemma product_open_generates_sets_PiM:
+ assumes "I \<noteq> {}"
+ assumes [simp]: "finite I"
+ shows "sets (PiM I (\<lambda>_. borel::'b::enumerable_basis measure)) =
+ sigma_sets (space (PiM I (\<lambda>_. borel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> Collect open}"
+proof -
+ from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
+ show ?thesis
+ proof (rule sigma_prod_algebra_sigma_eq)
+ show "finite I" by simp note[[show_types]]
+ fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
+ using S by (auto simp: open_enum_basis)
+ show "Collect open \<subseteq> Pow (space borel)" by simp
+ show "sets borel = sigma_sets (space borel) (Collect open)"
+ by (simp add: borel_def)
+ qed
+qed
+
+lemma finmap_UNIV[simp]: "(\<Union>J\<in>Collect finite. J \<leadsto> UNIV) = UNIV" by auto
+
+lemma borel_eq_PiF_borel:
+ shows "(borel :: ('i::countable \<Rightarrow>\<^isub>F 'a::polish_space) measure) =
+ PiF (Collect finite) (\<lambda>_. borel :: 'a measure)"
+proof (rule measure_eqI)
+ have C: "Collect finite \<noteq> {}" by auto
+ show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) = sets (PiF (Collect finite) (\<lambda>_. borel))"
+ proof
+ show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) \<subseteq> sets (PiF (Collect finite) (\<lambda>_. borel))"
+ apply (simp add: borel_def sets_PiF)
+ proof (rule sigma_sets_mono, safe, cases)
+ fix X::"('i \<Rightarrow>\<^isub>F 'a) set" assume "open X" "X \<noteq> {}"
+ from open_basisE[OF this] guess NA NB . note N = this
+ hence "X = (\<Union>i. Pi' (NA i) (NB i))" by simp
+ also have "\<dots> \<in>
+ sigma_sets UNIV {Pi' J S |S J. finite J \<and> S \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
+ using N by (intro Union sigma_sets.Basic) blast
+ finally show "X \<in> sigma_sets UNIV
+ {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}" .
+ qed (auto simp: Empty)
+ next
+ show "sets (PiF (Collect finite) (\<lambda>_. borel)) \<subseteq> sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure)"
+ proof
+ fix x assume x: "x \<in> sets (PiF (Collect finite::'i set set) (\<lambda>_. borel::'a measure))"
+ hence x_sp: "x \<subseteq> space (PiF (Collect finite) (\<lambda>_. borel))" by (rule sets_into_space)
+ let ?x = "\<lambda>J. x \<inter> {x. domain x = J}"
+ have "x = \<Union>{?x J |J. finite J}" by auto
+ also have "\<dots> \<in> sets borel"
+ proof (rule countable_finite_comprehension, assumption)
+ fix J::"'i set" assume "finite J"
+ { assume ef: "J = {}"
+ { assume e: "?x J = {}"
+ hence "?x J \<in> sets borel" by simp
+ } moreover {
+ assume "?x J \<noteq> {}"
+ then obtain f where "f \<in> x" "domain f = {}" using ef by auto
+ hence "?x J = {f}" using `J = {}`
+ by (auto simp: finmap_eq_iff)
+ also have "{f} \<in> sets borel" by simp
+ finally have "?x J \<in> sets borel" .
+ } ultimately have "?x J \<in> sets borel" by blast
+ } moreover {
+ assume "J \<noteq> ({}::'i set)"
+ from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'a set" . note S = this
+ have "(?x J) = x \<inter> {m. domain m \<in> {J}}" by auto
+ also have "\<dots> \<in> sets (PiF {J} (\<lambda>_. borel))"
+ using x by (rule restrict_sets_measurable) (auto simp: `finite J`)
+ also have "\<dots> = sigma_sets (space (PiF {J} (\<lambda>_. borel)))
+ {Pi' (J) F |F. (\<forall>j\<in>J. F j \<in> range enum_basis)}"
+ (is "_ = sigma_sets _ ?P")
+ by (rule enumerable_sigma_fprod_algebra_sigma_eq[OF `J \<noteq> {}` `finite J`])
+ also have "\<dots> \<subseteq> sets borel"
+ proof
+ fix x
+ assume "x \<in> sigma_sets (space (PiF {J} (\<lambda>_. borel))) ?P"
+ thus "x \<in> sets borel"
+ proof (rule sigma_sets.induct, safe)
+ fix F::"'i \<Rightarrow> 'a set"
+ assume "\<forall>j\<in>J. F j \<in> range enum_basis"
+ hence "Pi' J F \<in> range enum_basis_finmap"
+ unfolding range_enum_basis_eq
+ by (auto simp: `finite J` intro!: exI[where x=J] exI[where x=F])
+ hence "open (Pi' (J) F)" by (rule range_enum_basis_finmap_imp_open)
+ thus "Pi' (J) F \<in> sets borel" by simp
+ next
+ fix a::"('i \<Rightarrow>\<^isub>F 'a) set"
+ have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) =
+ Pi' (J) (\<lambda>_. UNIV)"
+ by (auto simp: space_PiF product_def)
+ moreover have "open (Pi' (J::'i set) (\<lambda>_. UNIV::'a set))"
+ by (intro open_Pi'I) auto
+ ultimately
+ have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) \<in> sets borel"
+ by simp
+ moreover
+ assume "a \<in> sets borel"
+ ultimately show "space (PiF {J} (\<lambda>_. borel)) - a \<in> sets borel" ..
+ qed auto
+ qed
+ finally have "(?x J) \<in> sets borel" .
+ } ultimately show "(?x J) \<in> sets borel" by blast
+ qed
+ finally show "x \<in> sets (borel)" .
+ qed
+ qed
+qed (simp add: emeasure_sigma borel_def PiF_def)
+
+subsection {* Isomorphism between Functions and Finite Maps *}
+
+lemma
+ measurable_compose:
+ fixes f::"'a \<Rightarrow> 'b"
+ assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
+ assumes "finite J"
+ shows "(\<lambda>m. compose J m f) \<in> measurable (PiM (f ` J) (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
+proof (rule measurable_PiM)
+ show "(\<lambda>m. compose J m f)
+ \<in> space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<rightarrow>
+ (J \<rightarrow> space M) \<inter> extensional J"
+ proof safe
+ fix x and i
+ assume x: "x \<in> space (PiM (f ` J) (\<lambda>_. M))" "i \<in> J"
+ with inj show "compose J x f i \<in> space M"
+ by (auto simp: space_PiM compose_def)
+ next
+ fix x assume "x \<in> space (PiM (f ` J) (\<lambda>_. M))"
+ show "(compose J x f) \<in> extensional J" by (rule compose_extensional)
+ qed
+next
+ fix S X
+ have inv: "\<And>j. j \<in> f ` J \<Longrightarrow> f (f' j) = j" using assms by auto
+ assume S: "S \<noteq> {} \<or> J = {}" "finite S" "S \<subseteq> J" and P: "\<And>i. i \<in> S \<Longrightarrow> X i \<in> sets M"
+ have "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
+ 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)))"
+ using assms inv S sets_into_space[OF P]
+ by (force simp: prod_emb_iff compose_def space_PiM extensional_def Pi_def intro: imageI)
+ also have "\<dots> \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))"
+ proof
+ from S show "f ` S \<subseteq> f ` J" by auto
+ show "(\<Pi>\<^isub>E b\<in>f ` S. X (f' b)) \<in> sets (Pi\<^isub>M (f ` S) (\<lambda>_. M))"
+ proof (rule sets_PiM_I_finite)
+ show "finite (f ` S)" using S by simp
+ fix i assume "i \<in> f ` S" hence "f' i \<in> S" using S assms by auto
+ thus "X (f' i) \<in> sets M" by (rule P)
+ qed
+ qed
+ finally show "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
+ space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))" .
+qed
+
+lemma
+ measurable_compose_inv:
+ fixes f::"'a \<Rightarrow> 'b"
+ assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
+ assumes "finite J"
+ shows "(\<lambda>m. compose (f ` J) m f') \<in> measurable (PiM J (\<lambda>_. M)) (PiM (f ` J) (\<lambda>_. M))"
+proof -
+ 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))"
+ using assms by (auto intro: measurable_compose)
+ moreover
+ from inj have "f' ` f ` J = J" by (metis (hide_lams, mono_tags) image_iff set_eqI)
+ ultimately show ?thesis by simp
+qed
+
+locale function_to_finmap =
+ fixes J::"'a set" and f :: "'a \<Rightarrow> 'b::countable" and f'
+ assumes [simp]: "finite J"
+ assumes inv: "i \<in> J \<Longrightarrow> f' (f i) = i"
+begin
+
+text {* to measure finmaps *}
+
+definition "fm = (finmap_of (f ` J)) o (\<lambda>g. compose (f ` J) g f')"
+
+lemma domain_fm[simp]: "domain (fm x) = f ` J"
+ unfolding fm_def by simp
+
+lemma fm_restrict[simp]: "fm (restrict y J) = fm y"
+ unfolding fm_def by (auto simp: compose_def inv intro: restrict_ext)
+
+lemma fm_product:
+ assumes "\<And>i. space (M i) = UNIV"
+ shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^isub>M J M) = (\<Pi>\<^isub>E j \<in> J. S (f j))"
+ using assms
+ by (auto simp: inv fm_def compose_def space_PiM Pi'_def)
+
+lemma fm_measurable:
+ assumes "f ` J \<in> N"
+ shows "fm \<in> measurable (Pi\<^isub>M J (\<lambda>_. M)) (Pi\<^isub>F N (\<lambda>_. M))"
+ unfolding fm_def
+proof (rule measurable_comp, rule measurable_compose_inv)
+ show "finmap_of (f ` J) \<in> measurable (Pi\<^isub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
+ using assms by (intro measurable_finmap_of measurable_component_singleton) auto
+qed (simp_all add: inv)
+
+lemma proj_fm:
+ assumes "x \<in> J"
+ shows "fm m (f x) = m x"
+ using assms by (auto simp: fm_def compose_def o_def inv)
+
+lemma inj_on_compose_f': "inj_on (\<lambda>g. compose (f ` J) g f') (extensional J)"
+proof (rule inj_on_inverseI)
+ fix x::"'a \<Rightarrow> 'c" assume "x \<in> extensional J"
+ thus "(\<lambda>x. compose J x f) (compose (f ` J) x f') = x"
+ by (auto simp: compose_def inv extensional_def)
+qed
+
+lemma inj_on_fm:
+ assumes "\<And>i. space (M i) = UNIV"
+ shows "inj_on fm (space (Pi\<^isub>M J M))"
+ using assms
+ apply (auto simp: fm_def space_PiM)
+ apply (rule comp_inj_on)
+ apply (rule inj_on_compose_f')
+ apply (rule finmap_of_inj_on_extensional_finite)
+ apply simp
+ apply (auto)
+ done
+
+text {* to measure functions *}
+
+definition "mf = (\<lambda>g. compose J g f) o proj"
+
+lemma
+ assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))" "finite J"
+ shows "proj (finmap_of J x) = x"
+ using assms by (auto simp: space_PiM extensional_def)
+
+lemma
+ assumes "x \<in> space (Pi\<^isub>F {J} (\<lambda>_. M))"
+ shows "finmap_of J (proj x) = x"
+ using assms by (auto simp: space_PiF Pi'_def finmap_eq_iff)
+
+lemma mf_fm:
+ assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))"
+ shows "mf (fm x) = x"
+proof -
+ have "mf (fm x) \<in> extensional J"
+ by (auto simp: mf_def extensional_def compose_def)
+ moreover
+ have "x \<in> extensional J" using assms sets_into_space
+ by (force simp: space_PiM)
+ moreover
+ { fix i assume "i \<in> J"
+ hence "mf (fm x) i = x i"
+ by (auto simp: inv mf_def compose_def fm_def)
+ }
+ ultimately
+ show ?thesis by (rule extensionalityI)
+qed
+
+lemma mf_measurable:
+ assumes "space M = UNIV"
+ shows "mf \<in> measurable (PiF {f ` J} (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
+ unfolding mf_def
+proof (rule measurable_comp, rule measurable_proj_PiM)
+ show "(\<lambda>g. compose J g f) \<in>
+ measurable (Pi\<^isub>M (f ` J) (\<lambda>x. M)) (Pi\<^isub>M J (\<lambda>_. M))"
+ by (rule measurable_compose, rule inv) auto
+qed (auto simp add: space_PiM extensional_def assms)
+
+lemma fm_image_measurable:
+ assumes "space M = UNIV"
+ assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M))"
+ shows "fm ` X \<in> sets (PiF {f ` J} (\<lambda>_. M))"
+proof -
+ have "fm ` X = (mf) -` X \<inter> space (PiF {f ` J} (\<lambda>_. M))"
+ proof safe
+ fix x assume "x \<in> X"
+ with mf_fm[of x] sets_into_space[OF assms(2)] show "fm x \<in> mf -` X" by auto
+ show "fm x \<in> space (PiF {f ` J} (\<lambda>_. M))" by (simp add: space_PiF assms)
+ next
+ fix y x
+ assume x: "mf y \<in> X"
+ assume y: "y \<in> space (PiF {f ` J} (\<lambda>_. M))"
+ thus "y \<in> fm ` X"
+ by (intro image_eqI[OF _ x], unfold finmap_eq_iff)
+ (auto simp: space_PiF fm_def mf_def compose_def inv Pi'_def)
+ qed
+ also have "\<dots> \<in> sets (PiF {f ` J} (\<lambda>_. M))"
+ using assms
+ by (intro measurable_sets[OF mf_measurable]) auto
+ finally show ?thesis .
+qed
+
+lemma fm_image_measurable_finite:
+ assumes "space M = UNIV"
+ assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M::'c measure))"
+ shows "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. M::'c measure))"
+ using fm_image_measurable[OF assms]
+ by (rule subspace_set_in_sets) (auto simp: finite_subset)
+
+text {* measure on finmaps *}
+
+definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"
+
+lemma sets_mapmeasure[simp]: "sets (mapmeasure M N) = sets (PiF (Collect finite) N)"
+ unfolding mapmeasure_def by simp
+
+lemma space_mapmeasure[simp]: "space (mapmeasure M N) = space (PiF (Collect finite) N)"
+ unfolding mapmeasure_def by simp
+
+lemma mapmeasure_PiF:
+ assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes "space N = UNIV"
+ assumes "X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
+ shows "emeasure (mapmeasure M (\<lambda>_. N)) X = emeasure M ((fm -` X \<inter> extensional J))"
+ using assms
+ by (auto simp: measurable_eqI[OF s1 refl s2 refl] mapmeasure_def emeasure_distr
+ fm_measurable space_PiM)
+
+lemma mapmeasure_PiM:
+ fixes N::"'c measure"
+ assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes N: "space N = UNIV"
+ assumes X: "X \<in> sets M"
+ shows "emeasure M X = emeasure (mapmeasure M (\<lambda>_. N)) (fm ` X)"
+ unfolding mapmeasure_def
+proof (subst emeasure_distr, subst measurable_eqI[OF s1 refl s2 refl], rule fm_measurable)
+ have "X \<subseteq> space (Pi\<^isub>M J (\<lambda>_. N))" using assms by (simp add: sets_into_space)
+ 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"
+ by (auto simp: vimage_image_eq inj_on_def)
+ thus "emeasure M X = emeasure M (fm -` fm ` X \<inter> space M)" using s1
+ by simp
+ show "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
+ by (rule fm_image_measurable_finite[OF N X[simplified s2]])
+qed simp
+
+end
+
+end
--- a/src/HOL/Probability/Probability.thy Thu Nov 15 10:49:58 2012 +0100
+++ b/src/HOL/Probability/Probability.thy Thu Nov 15 11:16:58 2012 +0100
@@ -3,7 +3,7 @@
Complete_Measure
Probability_Measure
Infinite_Product_Measure
- Regularity
+ Projective_Limit
Independent_Family
Information
begin
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Projective_Limit.thy Thu Nov 15 11:16:58 2012 +0100
@@ -0,0 +1,690 @@
+(* Title: HOL/Probability/Projective_Family.thy
+ Author: Fabian Immler, TU München
+*)
+
+header {* Projective Limit *}
+
+theory Projective_Limit
+ imports
+ Caratheodory
+ Fin_Map
+ Regularity
+ Projective_Family
+ Infinite_Product_Measure
+begin
+
+subsection {* Enumeration of Countable Union of Finite Sets *}
+
+locale finite_set_sequence =
+ fixes Js::"nat \<Rightarrow> 'a set"
+ assumes finite_seq[simp]: "finite (Js n)"
+begin
+
+text {* Enumerate finite set *}
+
+definition "enum_finite_max J = (SOME n. \<exists> f. J = f ` {i. i < n} \<and> inj_on f {i. i < n})"
+
+definition enum_finite where
+ "enum_finite J =
+ (SOME f. J = f ` {i::nat. i < enum_finite_max J} \<and> inj_on f {i. i < enum_finite_max J})"
+
+lemma enum_finite_max:
+ assumes "finite J"
+ shows "\<exists>f::nat\<Rightarrow>_. J = f ` {i. i < enum_finite_max J} \<and> inj_on f {i. i < enum_finite_max J}"
+ unfolding enum_finite_max_def
+ by (rule someI_ex) (rule finite_imp_nat_seg_image_inj_on[OF `finite J`])
+
+lemma enum_finite:
+ assumes "finite J"
+ shows "J = enum_finite J ` {i::nat. i < enum_finite_max J} \<and>
+ inj_on (enum_finite J) {i::nat. i < enum_finite_max J}"
+ unfolding enum_finite_def
+ by (rule someI_ex[of "\<lambda>f. J = f ` {i::nat. i < enum_finite_max J} \<and>
+ inj_on f {i. i < enum_finite_max J}"]) (rule enum_finite_max[OF `finite J`])
+
+lemma in_set_enum_exist:
+ assumes "finite A"
+ assumes "y \<in> A"
+ shows "\<exists>i. y = enum_finite A i"
+ using assms enum_finite by auto
+
+definition set_of_Un where "set_of_Un j = (LEAST n. j \<in> Js n)"
+
+definition index_in_set where "index_in_set J j = (SOME n. j = enum_finite J n)"
+
+definition Un_to_nat where
+ "Un_to_nat j = to_nat (set_of_Un j, index_in_set (Js (set_of_Un j)) j)"
+
+lemma inj_on_Un_to_nat:
+ shows "inj_on Un_to_nat (\<Union>n::nat. Js n)"
+proof (rule inj_onI)
+ fix x y
+ assume "x \<in> (\<Union>n. Js n)" "y \<in> (\<Union>n. Js n)"
+ then obtain ix iy where ix: "x \<in> Js ix" and iy: "y \<in> Js iy" by blast
+ assume "Un_to_nat x = Un_to_nat y"
+ hence "set_of_Un x = set_of_Un y"
+ "index_in_set (Js (set_of_Un y)) y = index_in_set (Js (set_of_Un x)) x"
+ by (auto simp: Un_to_nat_def)
+ moreover
+ {
+ fix x assume "x \<in> Js (set_of_Un x)"
+ have "x = enum_finite (Js (set_of_Un x)) (index_in_set (Js (set_of_Un x)) x)"
+ unfolding index_in_set_def
+ apply (rule someI_ex)
+ using `x \<in> Js (set_of_Un x)` finite_seq
+ apply (auto intro!: in_set_enum_exist)
+ done
+ } note H = this
+ moreover
+ have "y \<in> Js (set_of_Un y)" unfolding set_of_Un_def using iy by (rule LeastI)
+ note H[OF this]
+ moreover
+ have "x \<in> Js (set_of_Un x)" unfolding set_of_Un_def using ix by (rule LeastI)
+ note H[OF this]
+ ultimately show "x = y" by simp
+qed
+
+lemma inj_Un[simp]:
+ shows "inj_on (Un_to_nat) (Js n)"
+ by (intro subset_inj_on[OF inj_on_Un_to_nat]) (auto simp: assms)
+
+lemma Un_to_nat_injectiveD:
+ assumes "Un_to_nat x = Un_to_nat y"
+ assumes "x \<in> Js i" "y \<in> Js j"
+ shows "x = y"
+ using assms
+ by (intro inj_onD[OF inj_on_Un_to_nat]) auto
+
+end
+
+subsection {* Sequences of Finite Maps in Compact Sets *}
+
+locale finmap_seqs_into_compact =
+ fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a)" and M
+ assumes compact: "\<And>n. compact (K n)"
+ assumes f_in_K: "\<And>n. K n \<noteq> {}"
+ assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)"
+ assumes proj_in_K:
+ "\<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"
+begin
+
+lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n)"
+ using proj_in_K f_in_K
+proof cases
+ obtain k where "k \<in> K (Suc 0)" using f_in_K by auto
+ assume "\<forall>n. t \<notin> domain (f n)"
+ thus ?thesis
+ by (auto intro!: exI[where x=1] image_eqI[OF _ `k \<in> K (Suc 0)`]
+ simp: domain_K[OF `k \<in> K (Suc 0)`])
+qed blast
+
+lemma proj_in_KE:
+ obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n"
+ using proj_in_K' by blast
+
+lemma compact_projset:
+ shows "compact ((\<lambda>k. (k)\<^isub>F i) ` K n)"
+ using continuous_proj compact by (rule compact_continuous_image)
+
+end
+
+lemma compactE':
+ assumes "compact S" "\<forall>n\<ge>m. f n \<in> S"
+ obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
+proof atomize_elim
+ have "subseq (op + m)" by (simp add: subseq_def)
+ have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> S" using assms by auto
+ from compactE[OF `compact S` this] guess l r .
+ hence "l \<in> S" "subseq ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) ----> l"
+ using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def)
+ thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) ----> l" by blast
+qed
+
+sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^isub>F n) ----> l)"
+proof
+ fix n s
+ assume "subseq s"
+ from proj_in_KE[of n] guess n0 . note n0 = this
+ have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0"
+ proof safe
+ fix i assume "n0 \<le> i"
+ also have "\<dots> \<le> s i" by (rule seq_suble) fact
+ finally have "n0 \<le> s i" .
+ with n0 show "((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0 "
+ by auto
+ qed
+ from compactE'[OF compact_projset this] guess ls rs .
+ 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)
+qed
+
+lemma (in finmap_seqs_into_compact)
+ diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^isub>F n) ----> l"
+proof -
+ have "\<And>i n0. (f o seqseq i) i = f (diagseq i)" unfolding diagseq_def by simp
+ from reducer_reduces obtain l where l: "(\<lambda>i. ((f \<circ> seqseq (Suc n)) i)\<^isub>F n) ----> l"
+ unfolding seqseq_reducer
+ by auto
+ have "(\<lambda>i. (f (diagseq (i + Suc n)))\<^isub>F n) =
+ (\<lambda>i. ((f o (diagseq o (op + (Suc n)))) i)\<^isub>F n)" by (simp add: add_commute)
+ also have "\<dots> =
+ (\<lambda>i. ((f o ((seqseq (Suc n) o (\<lambda>x. fold_reduce (Suc n) x (Suc n + x))))) i)\<^isub>F n)"
+ unfolding diagseq_seqseq by simp
+ 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))"
+ by (simp add: o_def)
+ also have "\<dots> ----> l"
+ proof (rule LIMSEQ_subseq_LIMSEQ[OF _ subseq_diagonal_rest], rule tendstoI)
+ fix e::real assume "0 < e"
+ from tendstoD[OF l `0 < e`]
+ show "eventually (\<lambda>x. dist (((f \<circ> seqseq (Suc n)) x)\<^isub>F n) l < e)
+ sequentially" .
+ qed
+ finally show ?thesis by (intro exI) (rule LIMSEQ_offset)
+qed
+
+subsection {* Daniell-Kolmogorov Theorem *}
+
+text {* Existence of Projective Limit *}
+
+locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure"
+ for I::"'i set" and P
+begin
+
+abbreviation "PiB \<equiv> (\<lambda>J P. PiP J (\<lambda>_. borel) P)"
+
+lemma
+ emeasure_PiB_emb_not_empty:
+ assumes "I \<noteq> {}"
+ assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
+ shows "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (PiB J P) (Pi\<^isub>E J B)"
+proof -
+ let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space borel"
+ let ?G = generator
+ interpret G!: algebra ?\<Omega> generator by (intro algebra_generator) fact
+ note \<mu>G_mono =
+ G.additive_increasing[OF positive_\<mu>G[OF `I \<noteq> {}`] additive_\<mu>G[OF `I \<noteq> {}`], THEN increasingD]
+ have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
+ proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G,
+ OF `I \<noteq> {}`, OF `I \<noteq> {}`])
+ fix A assume "A \<in> ?G"
+ with generatorE guess J X . note JX = this
+ interpret prob_space "P J" using prob_space[OF `finite J`] .
+ show "\<mu>G A \<noteq> \<infinity>" using JX by (simp add: PiP_finite)
+ next
+ fix Z assume Z: "range Z \<subseteq> ?G" "decseq Z" "(\<Inter>i. Z i) = {}"
+ then have "decseq (\<lambda>i. \<mu>G (Z i))"
+ by (auto intro!: \<mu>G_mono simp: decseq_def)
+ moreover
+ have "(INF i. \<mu>G (Z i)) = 0"
+ proof (rule ccontr)
+ assume "(INF i. \<mu>G (Z i)) \<noteq> 0" (is "?a \<noteq> 0")
+ moreover have "0 \<le> ?a"
+ using Z positive_\<mu>G[OF `I \<noteq> {}`] by (auto intro!: INF_greatest simp: positive_def)
+ ultimately have "0 < ?a" by auto
+ hence "?a \<noteq> -\<infinity>" by auto
+ 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>
+ Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (PiB J P) B"
+ using Z by (intro allI generator_Ex) auto
+ then obtain J' B' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I"
+ "\<And>n. B' n \<in> sets (\<Pi>\<^isub>M i\<in>J' n. borel)"
+ and Z_emb: "\<And>n. Z n = emb I (J' n) (B' n)"
+ unfolding choice_iff by blast
+ moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
+ moreover def B \<equiv> "\<lambda>n. emb (J n) (J' n) (B' n)"
+ ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I"
+ "\<And>n. B n \<in> sets (\<Pi>\<^isub>M i\<in>J n. borel)"
+ by auto
+ have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
+ unfolding J_def by force
+ have "\<forall>n. \<exists>j. j \<in> J n" using J by blast
+ then obtain j where j: "\<And>n. j n \<in> J n"
+ unfolding choice_iff by blast
+ note [simp] = `\<And>n. finite (J n)`
+ from J Z_emb have Z_eq: "\<And>n. Z n = emb I (J n) (B n)" "\<And>n. Z n \<in> ?G"
+ unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto)
+ interpret prob_space "P (J i)" for i using prob_space by simp
+ have "?a \<le> \<mu>G (Z 0)" by (auto intro: INF_lower)
+ also have "\<dots> < \<infinity>" using J by (auto simp: Z_eq \<mu>G_eq PiP_finite proj_sets)
+ finally have "?a \<noteq> \<infinity>" by simp
+ have "\<And>n. \<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" unfolding Z_eq using J J_mono
+ by (subst \<mu>G_eq) (auto simp: PiP_finite proj_sets \<mu>G_eq)
+
+ interpret finite_set_sequence J by unfold_locales simp
+ def Utn \<equiv> Un_to_nat
+ interpret function_to_finmap "J n" Utn "inv_into (J n) Utn" for n
+ by unfold_locales (auto simp: Utn_def)
+ def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
+ let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
+ {
+ fix n
+ interpret finite_measure "P (J n)" by unfold_locales
+ have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
+ using J by (auto simp: P'_def mapmeasure_PiM proj_space proj_sets)
+ also
+ have "\<dots> = ?SUP n"
+ proof (rule inner_regular)
+ show "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
+ unfolding P'_def
+ by (auto simp: P'_def mapmeasure_PiF fm_measurable proj_space proj_sets)
+ show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
+ next
+ show "fm n ` B n \<in> sets borel"
+ unfolding borel_eq_PiF_borel
+ by (auto simp del: J(2) simp: P'_def fm_image_measurable_finite proj_sets J)
+ qed
+ finally
+ have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
+ } note R = this
+ have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a
+ \<and> compact K \<and> K \<subseteq> fm n ` B n"
+ proof
+ fix n
+ have "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
+ by (simp add: mapmeasure_PiF P'_def proj_space proj_sets)
+ then interpret finite_measure "P' n" ..
+ show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
+ compact K \<and> K \<subseteq> fm n ` B n"
+ unfolding R
+ proof (rule ccontr)
+ assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n) * ?a \<and>
+ compact K' \<and> K' \<subseteq> fm n ` B n)"
+ have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
+ proof (intro SUP_least)
+ fix K
+ assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
+ with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
+ by auto
+ hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
+ unfolding not_less[symmetric] by simp
+ hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
+ using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps)
+ thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
+ qed
+ hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
+ hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
+ hence "0 \<le> - (2 powr (-n) * ?a)"
+ using `?SUP _ \<noteq> \<infinity>` `?SUP _ \<noteq> - \<infinity>`
+ by (subst (asm) ereal_add_le_add_iff) (auto simp:)
+ moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a`
+ by (auto simp: ereal_zero_less_0_iff)
+ ultimately show False by simp
+ qed
+ qed
+ then obtain K' where K':
+ "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
+ "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
+ unfolding choice_iff by blast
+ def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
+ have K_sets: "\<And>n. K n \<in> sets (Pi\<^isub>M (J n) (\<lambda>_. borel))"
+ unfolding K_def
+ using compact_imp_closed[OF `compact (K' _)`]
+ by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
+ (auto simp: borel_eq_PiF_borel[symmetric])
+ have K_B: "\<And>n. K n \<subseteq> B n"
+ proof
+ fix x n
+ assume "x \<in> K n" hence fm_in: "fm n x \<in> fm n ` B n"
+ using K' by (force simp: K_def)
+ show "x \<in> B n"
+ apply (rule inj_on_image_mem_iff[OF inj_on_fm _ fm_in])
+ using `x \<in> K n` K_sets J[of n] sets_into_space
+ apply (auto simp: proj_space)
+ using J[of n] sets_into_space apply auto
+ done
+ qed
+ def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)"
+ have Z': "\<And>n. Z' n \<subseteq> Z n"
+ unfolding Z_eq unfolding Z'_def
+ proof (rule prod_emb_mono, safe)
+ fix n x assume "x \<in> K n"
+ hence "fm n x \<in> K' n" "x \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
+ by (simp_all add: K_def proj_space)
+ note this(1)
+ also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
+ finally have "fm n x \<in> fm n ` B n" .
+ thus "x \<in> B n"
+ proof safe
+ fix y assume "y \<in> B n"
+ moreover
+ hence "y \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))" using J sets_into_space[of "B n" "P (J n)"]
+ by (auto simp add: proj_space proj_sets)
+ assume "fm n x = fm n y"
+ note inj_onD[OF inj_on_fm[OF space_borel],
+ OF `fm n x = fm n y` `x \<in> space _` `y \<in> space _`]
+ ultimately show "x \<in> B n" by simp
+ qed
+ qed
+ { fix n
+ have "Z' n \<in> ?G" using K' unfolding Z'_def
+ apply (intro generatorI'[OF J(1-3)])
+ unfolding K_def proj_space
+ apply (rule measurable_sets[OF fm_measurable[of _ "Collect finite"]])
+ apply (auto simp add: P'_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
+ done
+ }
+ def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
+ hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
+ hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
+ have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
+ hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto
+ have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
+ proof -
+ fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact
+ have "Y n = (\<Inter> i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
+ by (auto simp: Y_def Z'_def)
+ also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter> i\<in>{1..n}. emb (J n) (J i) (K i))"
+ using `n \<ge> 1`
+ by (subst prod_emb_INT) auto
+ finally
+ have Y_emb:
+ "Y n = prod_emb I (\<lambda>_. borel) (J n)
+ (\<Inter> i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
+ hence "Y n \<in> ?G" using J J_mono K_sets `n \<ge> 1` by (intro generatorI[OF _ _ _ _ Y_emb]) auto
+ hence "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" unfolding Y_emb using J J_mono K_sets `n \<ge> 1`
+ by (subst \<mu>G_eq) (auto simp: PiP_finite proj_sets \<mu>G_eq)
+ interpret finite_measure "(PiP (J n) (\<lambda>_. borel) P)"
+ proof
+ have "emeasure (PiP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^isub>E space borel) \<noteq> \<infinity>"
+ using J by (subst emeasure_PiP) auto
+ thus "emeasure (PiP (J n) (\<lambda>_. borel) P) (space (PiP (J n) (\<lambda>_. borel) P)) \<noteq> \<infinity>"
+ by (simp add: space_PiM)
+ qed
+ have "\<mu>G (Z n) = PiP (J n) (\<lambda>_. borel) P (B n)"
+ unfolding Z_eq using J by (auto simp: \<mu>G_eq)
+ moreover have "\<mu>G (Y n) =
+ PiP (J n) (\<lambda>_. borel) P (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
+ unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst \<mu>G_eq) auto
+ moreover have "\<mu>G (Z n - Y n) = PiP (J n) (\<lambda>_. borel) P
+ (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
+ unfolding Z_eq Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1`
+ by (subst \<mu>G_eq) (auto intro!: Diff)
+ ultimately
+ have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
+ using J J_mono K_sets `n \<ge> 1`
+ by (simp only: emeasure_eq_measure)
+ (auto dest!: bspec[where x=n]
+ simp: extensional_restrict emeasure_eq_measure prod_emb_iff
+ intro!: measure_Diff[symmetric] set_mp[OF K_B])
+ also have subs: "Z n - Y n \<subseteq> (\<Union> i\<in>{1..n}. (Z i - Z' i))" using Z' Z `n \<ge> 1`
+ unfolding Y_def by (force simp: decseq_def)
+ have "Z n - Y n \<in> ?G" "(\<Union> i\<in>{1..n}. (Z i - Z' i)) \<in> ?G"
+ using `Z' _ \<in> ?G` `Z _ \<in> ?G` `Y _ \<in> ?G` by auto
+ hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union> i\<in>{1..n}. (Z i - Z' i))"
+ using subs G.additive_increasing[OF positive_\<mu>G[OF `I \<noteq> {}`] additive_\<mu>G[OF `I \<noteq> {}`]]
+ unfolding increasing_def by auto
+ also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> ?G` `Z' _ \<in> ?G`
+ by (intro G.subadditive[OF positive_\<mu>G additive_\<mu>G, OF `I \<noteq> {}` `I \<noteq> {}`]) auto
+ also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
+ proof (rule setsum_mono)
+ fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
+ have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
+ unfolding Z'_def Z_eq by simp
+ also have "\<dots> = P (J i) (B i - K i)"
+ apply (subst \<mu>G_eq) using J K_sets apply auto
+ apply (subst PiP_finite) apply auto
+ done
+ also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
+ apply (subst emeasure_Diff) using K_sets J `K _ \<subseteq> B _` apply (auto simp: proj_sets)
+ done
+ also have "\<dots> = P (J i) (B i) - P' i (K' i)"
+ unfolding K_def P'_def
+ by (auto simp: mapmeasure_PiF proj_space proj_sets borel_eq_PiF_borel[symmetric]
+ compact_imp_closed[OF `compact (K' _)`] space_PiM)
+ also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
+ finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
+ qed
+ also have "\<dots> = (\<Sum> i\<in>{1..n}. ereal (2 powr -real i) * ereal(real ?a))"
+ using `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` by (subst ereal_real') auto
+ also have "\<dots> = ereal (\<Sum> i\<in>{1..n}. (2 powr -real i) * (real ?a))" by simp
+ also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real i)) * real ?a)"
+ by (simp add: setsum_left_distrib)
+ also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps
+ proof (rule mult_strict_right_mono)
+ have "(\<Sum>i\<in>{1..n}. 2 powr - real i) = (\<Sum>i\<in>{1..<Suc n}. (1/2) ^ i)"
+ by (rule setsum_cong)
+ (auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide)
+ also have "{1..<Suc n} = {0..<Suc n} - {0}" by auto
+ also have "setsum (op ^ (1 / 2::real)) ({0..<Suc n} - {0}) =
+ setsum (op ^ (1 / 2)) ({0..<Suc n}) - 1" by (auto simp: setsum_diff1)
+ also have "\<dots> < 1" by (subst sumr_geometric) auto
+ finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" .
+ qed (auto simp:
+ `0 < ?a` `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` ereal_less_real_iff zero_ereal_def[symmetric])
+ also have "\<dots> = ?a" using `0 < ?a` `?a \<noteq> \<infinity>` by (auto simp: ereal_real')
+ also have "\<dots> \<le> \<mu>G (Z n)" by (auto intro: INF_lower)
+ finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
+ hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
+ using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less)
+ have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+ also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
+ apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+ finally have "\<mu>G (Y n) > 0"
+ using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric])
+ thus "Y n \<noteq> {}" using positive_\<mu>G `I \<noteq> {}` by (auto simp add: positive_def)
+ qed
+ hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
+ then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force
+ {
+ fix t and n m::nat
+ assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
+ from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto
+ also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] .
+ finally
+ have "fm n (restrict (y m) (J n)) \<in> K' n"
+ unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
+ moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
+ using J by (simp add: fm_def)
+ ultimately have "fm n (y m) \<in> K' n" by simp
+ } note fm_in_K' = this
+ interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel
+ proof
+ fix n show "compact (K' n)" by fact
+ next
+ fix n
+ from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto
+ also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto
+ finally
+ have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"
+ unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
+ thus "K' (Suc n) \<noteq> {}" by auto
+ fix k
+ assume "k \<in> K' (Suc n)"
+ with K'[of "Suc n"] sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto
+ then obtain b where "k = fm (Suc n) b" by auto
+ thus "domain k = domain (fm (Suc n) (y (Suc n)))"
+ by (simp_all add: fm_def)
+ next
+ fix t and n m::nat
+ assume "n \<le> m" hence "Suc n \<le> Suc m" by simp
+ assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
+ then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
+ hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
+ have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
+ by (intro fm_in_K') simp_all
+ show "(fm (Suc m) (y (Suc m)))\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K' (Suc n)"
+ apply (rule image_eqI[OF _ img])
+ using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
+ unfolding j by (subst proj_fm, auto)+
+ qed
+ have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z"
+ using diagonal_tendsto ..
+ then obtain z where z:
+ "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
+ unfolding choice_iff by blast
+ {
+ fix n :: nat assume "n \<ge> 1"
+ have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
+ by simp
+ moreover
+ {
+ fix t
+ assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
+ hence "t \<in> Utn ` J n" by simp
+ then obtain j where j: "t = Utn j" "j \<in> J n" by auto
+ have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
+ apply (subst (2) tendsto_iff, subst eventually_sequentially)
+ proof safe
+ fix e :: real assume "0 < e"
+ { fix i x assume "i \<ge> n" "t \<in> domain (fm n x)"
+ moreover
+ hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
+ ultimately have "(fm i x)\<^isub>F t = (fm n x)\<^isub>F t"
+ using j by (auto simp: proj_fm dest!:
+ Un_to_nat_injectiveD[simplified Utn_def[symmetric]])
+ } note index_shift = this
+ have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
+ apply (rule le_SucI)
+ apply (rule order_trans) apply simp
+ apply (rule seq_suble[OF subseq_diagseq])
+ done
+ from z
+ have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e"
+ unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
+ then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
+ dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e" by auto
+ show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e "
+ proof (rule exI[where x="max N n"], safe)
+ fix na assume "max N n \<le> na"
+ hence "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) =
+ dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^isub>F t) (z t)" using t
+ by (subst index_shift[OF I]) auto
+ also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp
+ finally show "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e" .
+ qed
+ qed
+ hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> (finmap_of (Utn ` J n) z)\<^isub>F t"
+ by (simp add: tendsto_intros)
+ } ultimately
+ have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
+ by (rule tendsto_finmap)
+ hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) ----> finmap_of (Utn ` J n) z"
+ by (intro lim_subseq) (simp add: subseq_def)
+ moreover
+ have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
+ apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI)
+ apply (rule le_trans)
+ apply (rule le_add2)
+ using seq_suble[OF subseq_diagseq]
+ apply auto
+ done
+ moreover
+ from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed)
+ ultimately
+ have "finmap_of (Utn ` J n) z \<in> K' n"
+ unfolding closed_sequential_limits by blast
+ also have "finmap_of (Utn ` J n) z = fm n (\<lambda>i. z (Utn i))"
+ by (auto simp: finmap_eq_iff fm_def compose_def f_inv_into_f)
+ finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
+ moreover
+ let ?J = "\<Union>n. J n"
+ have "(?J \<inter> J n) = J n" by auto
+ ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"
+ unfolding K_def by (auto simp: proj_space space_PiM)
+ hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def
+ using J by (auto simp: prod_emb_def extensional_def)
+ also have "\<dots> \<subseteq> Z n" using Z' by simp
+ finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .
+ } note in_Z = this
+ hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
+ hence "(\<Inter>i. Z i) \<noteq> {}" using Z INT_decseq_offset[OF `decseq Z`] by simp
+ thus False using Z by simp
+ qed
+ ultimately show "(\<lambda>i. \<mu>G (Z i)) ----> 0"
+ using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (Z i)"] by simp
+ qed
+ then guess \<mu> .. note \<mu> = this
+ def f \<equiv> "finmap_of J B"
+ show "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (PiB J P) (Pi\<^isub>E J B)"
+ proof (subst emeasure_extend_measure_Pair[OF PiP_def, of I "\<lambda>_. borel" \<mu>])
+ show "positive (sets (PiB I P)) \<mu>" "countably_additive (sets (PiB I P)) \<mu>"
+ using \<mu> unfolding sets_PiP sets_PiM_generator by (auto simp: measure_space_def)
+ next
+ show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> B \<in> J \<rightarrow> sets borel"
+ using assms by (auto simp: f_def)
+ next
+ fix J and X::"'i \<Rightarrow> 'a set"
+ show "prod_emb I (\<lambda>_. borel) J (Pi\<^isub>E J X) \<in> Pow ((I \<rightarrow> space borel) \<inter> extensional I)"
+ by (auto simp: prod_emb_def)
+ assume JX: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> J \<rightarrow> sets borel"
+ hence "emb I J (Pi\<^isub>E J X) \<in> generator" using assms
+ by (intro generatorI[where J=J and X="Pi\<^isub>E J X"]) (auto intro: sets_PiM_I_finite)
+ hence "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))" using \<mu> by simp
+ also have "\<dots> = emeasure (P J) (Pi\<^isub>E J X)"
+ using JX assms proj_sets
+ by (subst \<mu>G_eq) (auto simp: \<mu>G_eq PiP_finite intro: sets_PiM_I_finite)
+ finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = emeasure (P J) (Pi\<^isub>E J X)" .
+ next
+ show "emeasure (P J) (Pi\<^isub>E J B) = emeasure (PiP J (\<lambda>_. borel) P) (Pi\<^isub>E J B)"
+ using assms by (simp add: f_def PiP_finite Pi_def)
+ qed
+qed
+
+end
+
+sublocale polish_projective \<subseteq> P!: prob_space "(PiB I P)"
+proof
+ show "emeasure (PiB I P) (space (PiB I P)) = 1"
+ proof cases
+ assume "I = {}"
+ interpret prob_space "P {}" using prob_space by simp
+ show ?thesis
+ by (simp add: space_PiM_empty PiP_finite emeasure_space_1 `I = {}`)
+ next
+ assume "I \<noteq> {}"
+ then obtain i where "i \<in> I" by auto
+ interpret prob_space "P {i}" using prob_space by simp
+ have R: "(space (PiB I P)) = (emb I {i} (Pi\<^isub>E {i} (\<lambda>_. space borel)))"
+ by (auto simp: prod_emb_def space_PiM)
+ moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM)
+ ultimately show ?thesis using `i \<in> I`
+ apply (subst R)
+ apply (subst emeasure_PiB_emb_not_empty)
+ apply (auto simp: PiP_finite emeasure_space_1)
+ done
+ qed
+qed
+
+context polish_projective begin
+
+lemma emeasure_PiB_emb:
+ assumes X: "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
+ shows "emeasure (PiB I P) (emb I J (Pi\<^isub>E J B)) = emeasure (P J) (Pi\<^isub>E J B)"
+proof cases
+ interpret prob_space "P {}" using prob_space by simp
+ assume "J = {}"
+ moreover have "emb I {} {\<lambda>x. undefined} = space (PiB I P)"
+ by (auto simp: space_PiM prod_emb_def)
+ moreover have "{\<lambda>x. undefined} = space (PiB {} P)"
+ by (auto simp: space_PiM prod_emb_def)
+ ultimately show ?thesis
+ by (simp add: P.emeasure_space_1 PiP_finite emeasure_space_1 del: space_PiP)
+next
+ assume "J \<noteq> {}" with X show ?thesis
+ by (subst emeasure_PiB_emb_not_empty) (auto simp: PiP_finite)
+qed
+
+lemma measure_PiB_emb:
+ assumes "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
+ shows "measure (PiB I P) (emb I J (Pi\<^isub>E J B)) = measure (P J) (Pi\<^isub>E J B)"
+proof -
+ interpret prob_space "P J" using prob_space assms by simp
+ show ?thesis
+ using emeasure_PiB_emb[OF assms]
+ unfolding emeasure_eq_measure PiP_finite[OF `finite J` `J \<subseteq> I`] P.emeasure_eq_measure
+ by simp
+qed
+
+end
+
+locale polish_product_prob_space =
+ product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
+
+sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
+proof qed
+
+lemma (in polish_product_prob_space)
+ PiP_eq_PiM:
+ "I \<noteq> {} \<Longrightarrow> PiP I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
+ PiM I (\<lambda>_. borel)"
+ by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_PiB_emb)
+
+end