1(*  Title:      HOL/Probability/Fin_Map.thy
2    Author:     Fabian Immler, TU M��nchen
3*)
4
5section \<open>Finite Maps\<close>
6
7theory Fin_Map
8  imports "HOL-Analysis.Finite_Product_Measure" "HOL-Library.Finite_Map"
9begin
10
11text \<open>The \<^type>\<open>fmap\<close> type can be instantiated to \<^class>\<open>polish_space\<close>, needed for the proof of
12  projective limit. \<^const>\<open>extensional\<close> functions are used for the representation in order to
13  stay close to the developments of (finite) products \<^const>\<open>Pi\<^sub>E\<close> and their sigma-algebra
14  \<^const>\<open>Pi\<^sub>M\<close>.\<close>
15
16type_notation fmap ("(_ \<Rightarrow>\<^sub>F /_)" [22, 21] 21)
17
18unbundle fmap.lifting
19
20
21subsection \<open>Domain and Application\<close>
22
23lift_definition domain::"('i \<Rightarrow>\<^sub>F 'a) \<Rightarrow> 'i set" is dom .
24
25lemma finite_domain[simp, intro]: "finite (domain P)"
26  by transfer simp
27
28lift_definition proj :: "('i \<Rightarrow>\<^sub>F 'a) \<Rightarrow> 'i \<Rightarrow> 'a" ("'((_)')\<^sub>F" [0] 1000) is
29  "\<lambda>f x. if x \<in> dom f then the (f x) else undefined" .
30
31declare [[coercion proj]]
32
33lemma extensional_proj[simp, intro]: "(P)\<^sub>F \<in> extensional (domain P)"
34  by transfer (auto simp: extensional_def)
35
36lemma proj_undefined[simp, intro]: "i \<notin> domain P \<Longrightarrow> P i = undefined"
37  using extensional_proj[of P] unfolding extensional_def by auto
38
39lemma finmap_eq_iff: "P = Q \<longleftrightarrow> (domain P = domain Q \<and> (\<forall>i\<in>domain P. P i = Q i))"
40  apply transfer
41  apply (safe intro!: ext)
42  subgoal for P Q x
43    by (cases "x \<in> dom P"; cases "P x") (auto dest!: bspec[where x=x])
44  done
45
46
47subsection \<open>Constructor of Finite Maps\<close>
48
49lift_definition finmap_of::"'i set \<Rightarrow> ('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow>\<^sub>F 'a)" is
50  "\<lambda>I f x. if x \<in> I \<and> finite I then Some (f x) else None"
51  by (simp add: dom_def)
52
53lemma proj_finmap_of[simp]:
54  assumes "finite inds"
55  shows "(finmap_of inds f)\<^sub>F = restrict f inds"
56  using assms
57  by transfer force
58
59lemma domain_finmap_of[simp]:
60  assumes "finite inds"
61  shows "domain (finmap_of inds f) = inds"
62  using assms
63  by transfer (auto split: if_splits)
64
65lemma finmap_of_eq_iff[simp]:
66  assumes "finite i" "finite j"
67  shows "finmap_of i m = finmap_of j n \<longleftrightarrow> i = j \<and> (\<forall>k\<in>i. m k= n k)"
68  using assms by (auto simp: finmap_eq_iff)
69
70lemma finmap_of_inj_on_extensional_finite:
71  assumes "finite K"
72  assumes "S \<subseteq> extensional K"
73  shows "inj_on (finmap_of K) S"
74proof (rule inj_onI)
75  fix x y::"'a \<Rightarrow> 'b"
76  assume "finmap_of K x = finmap_of K y"
77  hence "(finmap_of K x)\<^sub>F = (finmap_of K y)\<^sub>F" by simp
78  moreover
79  assume "x \<in> S" "y \<in> S" hence "x \<in> extensional K" "y \<in> extensional K" using assms by auto
80  ultimately
81  show "x = y" using assms by (simp add: extensional_restrict)
82qed
83
84subsection \<open>Product set of Finite Maps\<close>
85
86text \<open>This is \<^term>\<open>Pi\<close> for Finite Maps, most of this is copied\<close>
87
88definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^sub>F 'a) set" where
89  "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^sub>F i \<in> A i) } "
90
91syntax
92  "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>'' _\<in>_./ _)"   10)
93translations
94  "\<Pi>' x\<in>A. B" == "CONST Pi' A (\<lambda>x. B)"
95
96subsubsection\<open>Basic Properties of \<^term>\<open>Pi'\<close>\<close>
97
98lemma Pi'_I[intro!]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
99  by (simp add: Pi'_def)
100
101lemma Pi'_I'[simp]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
102  by (simp add:Pi'_def)
103
104lemma Pi'_mem: "f\<in> Pi' A B \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<in> B x"
105  by (simp add: Pi'_def)
106
107lemma Pi'_iff: "f \<in> Pi' I X \<longleftrightarrow> domain f = I \<and> (\<forall>i\<in>I. f i \<in> X i)"
108  unfolding Pi'_def by auto
109
110lemma Pi'E [elim]:
111  "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"
112  by(auto simp: Pi'_def)
113
114lemma in_Pi'_cong:
115  "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"
116  by (auto simp: Pi'_def)
117
118lemma Pi'_eq_empty[simp]:
119  assumes "finite A" shows "(Pi' A B) = {} \<longleftrightarrow> (\<exists>x\<in>A. B x = {})"
120  using assms
121  apply (simp add: Pi'_def, auto)
122  apply (drule_tac x = "finmap_of A (\<lambda>u. SOME y. y \<in> B u)" in spec, auto)
123  apply (cut_tac P= "%y. y \<in> B i" in some_eq_ex, auto)
124  done
125
126lemma Pi'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> Pi' A B \<subseteq> Pi' A C"
127  by (auto simp: Pi'_def)
128
129lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^sub>E A B) = proj ` Pi' A B"
130  apply (auto simp: Pi'_def Pi_def extensional_def)
131  apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI)
132  apply auto
133  done
134
135subsection \<open>Topological Space of Finite Maps\<close>
136
137instantiation fmap :: (type, topological_space) topological_space
138begin
139
140definition open_fmap :: "('a \<Rightarrow>\<^sub>F 'b) set \<Rightarrow> bool" where
141   [code del]: "open_fmap = generate_topology {Pi' a b|a b. \<forall>i\<in>a. open (b i)}"
142
143lemma open_Pi'I: "(\<And>i. i \<in> I \<Longrightarrow> open (A i)) \<Longrightarrow> open (Pi' I A)"
144  by (auto intro: generate_topology.Basis simp: open_fmap_def)
145
146instance using topological_space_generate_topology
147  by intro_classes (auto simp: open_fmap_def class.topological_space_def)
148
149end
150
151lemma open_restricted_space:
152  shows "open {m. P (domain m)}"
153proof -
154  have "{m. P (domain m)} = (\<Union>i \<in> Collect P. {m. domain m = i})" by auto
155  also have "open \<dots>"
156  proof (rule, safe, cases)
157    fix i::"'a set"
158    assume "finite i"
159    hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
160    also have "open \<dots>" by (auto intro: open_Pi'I simp: \<open>finite i\<close>)
161    finally show "open {m. domain m = i}" .
162  next
163    fix i::"'a set"
164    assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
165    also have "open \<dots>" by simp
166    finally show "open {m. domain m = i}" .
167  qed
168  finally show ?thesis .
169qed
170
171lemma closed_restricted_space:
172  shows "closed {m. P (domain m)}"
173  using open_restricted_space[of "\<lambda>x. \<not> P x"]
174  unfolding closed_def by (rule back_subst) auto
175
176lemma tendsto_proj: "((\<lambda>x. x) \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. (x)\<^sub>F i) \<longlongrightarrow> (a)\<^sub>F i) F"
177  unfolding tendsto_def
178proof safe
179  fix S::"'b set"
180  let ?S = "Pi' (domain a) (\<lambda>x. if x = i then S else UNIV)"
181  assume "open S" hence "open ?S" by (auto intro!: open_Pi'I)
182  moreover assume "\<forall>S. open S \<longrightarrow> a \<in> S \<longrightarrow> eventually (\<lambda>x. x \<in> S) F" "a i \<in> S"
183  ultimately have "eventually (\<lambda>x. x \<in> ?S) F" by auto
184  thus "eventually (\<lambda>x. (x)\<^sub>F i \<in> S) F"
185    by eventually_elim (insert \<open>a i \<in> S\<close>, force simp: Pi'_iff split: if_split_asm)
186qed
187
188lemma continuous_proj:
189  shows "continuous_on s (\<lambda>x. (x)\<^sub>F i)"
190  unfolding continuous_on_def by (safe intro!: tendsto_proj tendsto_ident_at)
191
192instance fmap :: (type, first_countable_topology) first_countable_topology
193proof
194  fix x::"'a\<Rightarrow>\<^sub>F'b"
195  have "\<forall>i. \<exists>A. countable A \<and> (\<forall>a\<in>A. x i \<in> a) \<and> (\<forall>a\<in>A. open a) \<and>
196    (\<forall>S. open S \<and> x i \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)) \<and> (\<forall>a b. a \<in> A \<longrightarrow> b \<in> A \<longrightarrow> a \<inter> b \<in> A)" (is "\<forall>i. ?th i")
197  proof
198    fix i from first_countable_basis_Int_stableE[of "x i"] guess A .
199    thus "?th i" by (intro exI[where x=A]) simp
200  qed
201  then guess A unfolding choice_iff .. note A = this
202  hence open_sub: "\<And>i S. i\<in>domain x \<Longrightarrow> open (S i) \<Longrightarrow> x i\<in>(S i) \<Longrightarrow> (\<exists>a\<in>A i. a\<subseteq>(S i))" by auto
203  have A_notempty: "\<And>i. i \<in> domain x \<Longrightarrow> A i \<noteq> {}" using open_sub[of _ "\<lambda>_. UNIV"] by auto
204  let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^sub>E (domain x) A)"
205  show "\<exists>A::nat \<Rightarrow> ('a\<Rightarrow>\<^sub>F'b) set. (\<forall>i. x \<in> (A i) \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
206  proof (rule first_countableI[of "?A"], safe)
207    show "countable ?A" using A by (simp add: countable_PiE)
208  next
209    fix S::"('a \<Rightarrow>\<^sub>F 'b) set" assume "open S" "x \<in> S"
210    thus "\<exists>a\<in>?A. a \<subseteq> S" unfolding open_fmap_def
211    proof (induct rule: generate_topology.induct)
212      case UNIV thus ?case by (auto simp add: ex_in_conv PiE_eq_empty_iff A_notempty)
213    next
214      case (Int a b)
215      then obtain f g where
216        "f \<in> Pi\<^sub>E (domain x) A" "Pi' (domain x) f \<subseteq> a" "g \<in> Pi\<^sub>E (domain x) A" "Pi' (domain x) g \<subseteq> b"
217        by auto
218      thus ?case using A
219        by (auto simp: Pi'_iff PiE_iff extensional_def Int_stable_def
220            intro!: bexI[where x="\<lambda>i. f i \<inter> g i"])
221    next
222      case (UN B)
223      then obtain b where "x \<in> b" "b \<in> B" by auto
224      hence "\<exists>a\<in>?A. a \<subseteq> b" using UN by simp
225      thus ?case using \<open>b \<in> B\<close> by blast
226    next
227      case (Basis s)
228      then obtain a b where xs: "x\<in> Pi' a b" "s = Pi' a b" "\<And>i. i\<in>a \<Longrightarrow> open (b i)" by auto
229      have "\<forall>i. \<exists>a. (i \<in> domain x \<and> open (b i) \<and> (x)\<^sub>F i \<in> b i) \<longrightarrow> (a\<in>A i \<and> a \<subseteq> b i)"
230        using open_sub[of _ b] by auto
231      then obtain b'
232        where "\<And>i. i \<in> domain x \<Longrightarrow> open (b i) \<Longrightarrow> (x)\<^sub>F i \<in> b i \<Longrightarrow> (b' i \<in>A i \<and> b' i \<subseteq> b i)"
233          unfolding choice_iff by auto
234      with xs have "\<And>i. i \<in> a \<Longrightarrow> (b' i \<in>A i \<and> b' i \<subseteq> b i)" "Pi' a b' \<subseteq> Pi' a b"
235        by (auto simp: Pi'_iff intro!: Pi'_mono)
236      thus ?case using xs
237        by (intro bexI[where x="Pi' a b'"])
238          (auto simp: Pi'_iff intro!: image_eqI[where x="restrict b' (domain x)"])
239    qed
240  qed (insert A,auto simp: PiE_iff intro!: open_Pi'I)
241qed
242
243subsection \<open>Metric Space of Finite Maps\<close>
244
245(* TODO: Product of uniform spaces and compatibility with metric_spaces! *)
246
247instantiation fmap :: (type, metric_space) dist
248begin
249
250definition dist_fmap where
251  "dist P Q = Max (range (\<lambda>i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i))) + (if domain P = domain Q then 0 else 1)"
252
253instance ..
254end
255
256instantiation fmap :: (type, metric_space) uniformity_dist
257begin
258
259definition [code del]:
260  "(uniformity :: (('a, 'b) fmap \<times> ('a \<Rightarrow>\<^sub>F 'b)) filter) =
261    (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
262
263instance
264  by standard (rule uniformity_fmap_def)
265end
266
267declare uniformity_Abort[where 'a="('a \<Rightarrow>\<^sub>F 'b::metric_space)", code]
268
269instantiation fmap :: (type, metric_space) metric_space
270begin
271
272lemma finite_proj_image': "x \<notin> domain P \<Longrightarrow> finite ((P)\<^sub>F ` S)"
273  by (rule finite_subset[of _ "proj P ` (domain P \<inter> S \<union> {x})"]) auto
274
275lemma finite_proj_image: "finite ((P)\<^sub>F ` S)"
276 by (cases "\<exists>x. x \<notin> domain P") (auto intro: finite_proj_image' finite_subset[where B="domain P"])
277
278lemma finite_proj_diag: "finite ((\<lambda>i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S)"
279proof -
280  have "(\<lambda>i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S = (\<lambda>(i, j). d i j) ` ((\<lambda>i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S)" by auto
281  moreover have "((\<lambda>i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S) \<subseteq> (\<lambda>i. (P)\<^sub>F i) ` S \<times> (\<lambda>i. (Q)\<^sub>F i) ` S" by auto
282  moreover have "finite \<dots>" using finite_proj_image[of P S] finite_proj_image[of Q S]
283    by (intro finite_cartesian_product) simp_all
284  ultimately show ?thesis by (simp add: finite_subset)
285qed
286
287lemma dist_le_1_imp_domain_eq:
288  shows "dist P Q < 1 \<Longrightarrow> domain P = domain Q"
289  by (simp add: dist_fmap_def finite_proj_diag split: if_split_asm)
290
291lemma dist_proj:
292  shows "dist ((x)\<^sub>F i) ((y)\<^sub>F i) \<le> dist x y"
293proof -
294  have "dist (x i) (y i) \<le> Max (range (\<lambda>i. dist (x i) (y i)))"
295    by (simp add: Max_ge_iff finite_proj_diag)
296  also have "\<dots> \<le> dist x y" by (simp add: dist_fmap_def)
297  finally show ?thesis .
298qed
299
300lemma dist_finmap_lessI:
301  assumes "domain P = domain Q"
302  assumes "0 < e"
303  assumes "\<And>i. i \<in> domain P \<Longrightarrow> dist (P i) (Q i) < e"
304  shows "dist P Q < e"
305proof -
306  have "dist P Q = Max (range (\<lambda>i. dist (P i) (Q i)))"
307    using assms by (simp add: dist_fmap_def finite_proj_diag)
308  also have "\<dots> < e"
309  proof (subst Max_less_iff, safe)
310    fix i
311    show "dist ((P)\<^sub>F i) ((Q)\<^sub>F i) < e" using assms
312      by (cases "i \<in> domain P") simp_all
313  qed (simp add: finite_proj_diag)
314  finally show ?thesis .
315qed
316
317instance
318proof
319  fix S::"('a \<Rightarrow>\<^sub>F 'b) set"
320  have *: "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" (is "_ = ?od")
321  proof
322    assume "open S"
323    thus ?od
324      unfolding open_fmap_def
325    proof (induct rule: generate_topology.induct)
326      case UNIV thus ?case by (auto intro: zero_less_one)
327    next
328      case (Int a b)
329      show ?case
330      proof safe
331        fix x assume x: "x \<in> a" "x \<in> b"
332        with Int x obtain e1 e2 where
333          "e1>0" "\<forall>y. dist y x < e1 \<longrightarrow> y \<in> a" "e2>0" "\<forall>y. dist y x < e2 \<longrightarrow> y \<in> b" by force
334        thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> a \<inter> b"
335          by (auto intro!: exI[where x="min e1 e2"])
336      qed
337    next
338      case (UN K)
339      show ?case
340      proof safe
341        fix x X assume "x \<in> X" and X: "X \<in> K"
342        with UN obtain e where "e>0" "\<And>y. dist y x < e \<longrightarrow> y \<in> X" by force
343        with X show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> \<Union>K" by auto
344      qed
345    next
346      case (Basis s) then obtain a b where s: "s = Pi' a b" and b: "\<And>i. i\<in>a \<Longrightarrow> open (b i)" by auto
347      show ?case
348      proof safe
349        fix x assume "x \<in> s"
350        hence [simp]: "finite a" and a_dom: "a = domain x" using s by (auto simp: Pi'_iff)
351        obtain es where es: "\<forall>i \<in> a. es i > 0 \<and> (\<forall>y. dist y (proj x i) < es i \<longrightarrow> y \<in> b i)"
352          using b \<open>x \<in> s\<close> by atomize_elim (intro bchoice, auto simp: open_dist s)
353        hence in_b: "\<And>i y. i \<in> a \<Longrightarrow> dist y (proj x i) < es i \<Longrightarrow> y \<in> b i" by auto
354        show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> s"
355        proof (cases, rule, safe)
356          assume "a \<noteq> {}"
357          show "0 < min 1 (Min (es ` a))" using es by (auto simp: \<open>a \<noteq> {}\<close>)
358          fix y assume d: "dist y x < min 1 (Min (es ` a))"
359          show "y \<in> s" unfolding s
360          proof
361            show "domain y = a" using d s \<open>a \<noteq> {}\<close> by (auto simp: dist_le_1_imp_domain_eq a_dom)
362            fix i assume i: "i \<in> a"
363            hence "dist ((y)\<^sub>F i) ((x)\<^sub>F i) < es i" using d
364              by (auto simp: dist_fmap_def \<open>a \<noteq> {}\<close> intro!: le_less_trans[OF dist_proj])
365            with i show "y i \<in> b i" by (rule in_b)
366          qed
367        next
368          assume "\<not>a \<noteq> {}"
369          thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> s"
370            using s \<open>x \<in> s\<close> by (auto simp: Pi'_def dist_le_1_imp_domain_eq intro!: exI[where x=1])
371        qed
372      qed
373    qed
374  next
375    assume "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S"
376    then obtain e where e_pos: "\<And>x. x \<in> S \<Longrightarrow> e x > 0" and
377      e_in:  "\<And>x y . x \<in> S \<Longrightarrow> dist y x < e x \<Longrightarrow> y \<in> S"
378      unfolding bchoice_iff
379      by auto
380    have S_eq: "S = \<Union>{Pi' a b| a b. \<exists>x\<in>S. domain x = a \<and> b = (\<lambda>i. ball (x i) (e x))}"
381    proof safe
382      fix x assume "x \<in> S"
383      thus "x \<in> \<Union>{Pi' a b| a b. \<exists>x\<in>S. domain x = a \<and> b = (\<lambda>i. ball (x i) (e x))}"
384        using e_pos by (auto intro!: exI[where x="Pi' (domain x) (\<lambda>i. ball (x i) (e x))"])
385    next
386      fix x y
387      assume "y \<in> S"
388      moreover
389      assume "x \<in> (\<Pi>' i\<in>domain y. ball (y i) (e y))"
390      hence "dist x y < e y" using e_pos \<open>y \<in> S\<close>
391        by (auto simp: dist_fmap_def Pi'_iff finite_proj_diag dist_commute)
392      ultimately show "x \<in> S" by (rule e_in)
393    qed
394    also have "open \<dots>"
395      unfolding open_fmap_def
396      by (intro generate_topology.UN) (auto intro: generate_topology.Basis)
397    finally show "open S" .
398  qed
399  show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)"
400    unfolding * eventually_uniformity_metric
401    by (simp del: split_paired_All add: dist_fmap_def dist_commute eq_commute)
402next
403  fix P Q::"'a \<Rightarrow>\<^sub>F 'b"
404  have Max_eq_iff: "\<And>A m. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> (Max A = m) = (m \<in> A \<and> (\<forall>a\<in>A. a \<le> m))"
405    by (auto intro: Max_in Max_eqI)
406  show "dist P Q = 0 \<longleftrightarrow> P = Q"
407    by (auto simp: finmap_eq_iff dist_fmap_def Max_ge_iff finite_proj_diag Max_eq_iff
408        add_nonneg_eq_0_iff
409      intro!: Max_eqI image_eqI[where x=undefined])
410next
411  fix P Q R::"'a \<Rightarrow>\<^sub>F 'b"
412  let ?dists = "\<lambda>P Q i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i)"
413  let ?dpq = "?dists P Q" and ?dpr = "?dists P R" and ?dqr = "?dists Q R"
414  let ?dom = "\<lambda>P Q. (if domain P = domain Q then 0 else 1::real)"
415  have "dist P Q = Max (range ?dpq) + ?dom P Q"
416    by (simp add: dist_fmap_def)
417  also obtain t where "t \<in> range ?dpq" "t = Max (range ?dpq)" by (simp add: finite_proj_diag)
418  then obtain i where "Max (range ?dpq) = ?dpq i" by auto
419  also have "?dpq i \<le> ?dpr i + ?dqr i" by (rule dist_triangle2)
420  also have "?dpr i \<le> Max (range ?dpr)" by (simp add: finite_proj_diag)
421  also have "?dqr i \<le> Max (range ?dqr)" by (simp add: finite_proj_diag)
422  also have "?dom P Q \<le> ?dom P R + ?dom Q R" by simp
423  finally show "dist P Q \<le> dist P R + dist Q R" by (simp add: dist_fmap_def ac_simps)
424qed
425
426end
427
428subsection \<open>Complete Space of Finite Maps\<close>
429
430lemma tendsto_finmap:
431  fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^sub>F ('a::metric_space))"
432  assumes ind_f:  "\<And>n. domain (f n) = domain g"
433  assumes proj_g:  "\<And>i. i \<in> domain g \<Longrightarrow> (\<lambda>n. (f n) i) \<longlonglongrightarrow> g i"
434  shows "f \<longlonglongrightarrow> g"
435  unfolding tendsto_iff
436proof safe
437  fix e::real assume "0 < e"
438  let ?dists = "\<lambda>x i. dist ((f x)\<^sub>F i) ((g)\<^sub>F i)"
439  have "eventually (\<lambda>x. \<forall>i\<in>domain g. ?dists x i < e) sequentially"
440    using finite_domain[of g] proj_g
441  proof induct
442    case (insert i G)
443    with \<open>0 < e\<close> have "eventually (\<lambda>x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff)
444    moreover
445    from insert have "eventually (\<lambda>x. \<forall>i\<in>G. dist ((f x)\<^sub>F i) ((g)\<^sub>F i) < e) sequentially" by simp
446    ultimately show ?case by eventually_elim auto
447  qed simp
448  thus "eventually (\<lambda>x. dist (f x) g < e) sequentially"
449    by eventually_elim (auto simp add: dist_fmap_def finite_proj_diag ind_f \<open>0 < e\<close>)
450qed
451
452instance fmap :: (type, complete_space) complete_space
453proof
454  fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^sub>F 'b"
455  assume "Cauchy P"
456  then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
457    by (force simp: Cauchy_altdef2)
458  define d where "d = domain (P Nd)"
459  with Nd have dim: "\<And>n. n \<ge> Nd \<Longrightarrow> domain (P n) = d" using dist_le_1_imp_domain_eq by auto
460  have [simp]: "finite d" unfolding d_def by simp
461  define p where "p i n = P n i" for i n
462  define q where "q i = lim (p i)" for i
463  define Q where "Q = finmap_of d q"
464  have q: "\<And>i. i \<in> d \<Longrightarrow> q i = Q i" by (auto simp add: Q_def Abs_fmap_inverse)
465  {
466    fix i assume "i \<in> d"
467    have "Cauchy (p i)" unfolding Cauchy_altdef2 p_def
468    proof safe
469      fix e::real assume "0 < e"
470      with \<open>Cauchy P\<close> obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
471        by (force simp: Cauchy_altdef2 min_def)
472      hence "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto
473      with dim have dim: "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = d" by (metis nat_le_linear)
474      show "\<exists>N. \<forall>n\<ge>N. dist ((P n) i) ((P N) i) < e"
475      proof (safe intro!: exI[where x="N"])
476        fix n assume "N \<le> n" have "N \<le> N" by simp
477        have "dist ((P n) i) ((P N) i) \<le> dist (P n) (P N)"
478          using dim[OF \<open>N \<le> n\<close>]  dim[OF \<open>N \<le> N\<close>] \<open>i \<in> d\<close>
479          by (auto intro!: dist_proj)
480        also have "\<dots> < e" using N[OF \<open>N \<le> n\<close>] by simp
481        finally show "dist ((P n) i) ((P N) i) < e" .
482      qed
483    qed
484    hence "convergent (p i)" by (metis Cauchy_convergent_iff)
485    hence "p i \<longlonglongrightarrow> q i" unfolding q_def convergent_def by (metis limI)
486  } note p = this
487  have "P \<longlonglongrightarrow> Q"
488  proof (rule metric_LIMSEQ_I)
489    fix e::real assume "0 < e"
490    have "\<exists>ni. \<forall>i\<in>d. \<forall>n\<ge>ni i. dist (p i n) (q i) < e"
491    proof (safe intro!: bchoice)
492      fix i assume "i \<in> d"
493      from p[OF \<open>i \<in> d\<close>, THEN metric_LIMSEQ_D, OF \<open>0 < e\<close>]
494      show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e" .
495    qed then guess ni .. note ni = this
496    define N where "N = max Nd (Max (ni ` d))"
497    show "\<exists>N. \<forall>n\<ge>N. dist (P n) Q < e"
498    proof (safe intro!: exI[where x="N"])
499      fix n assume "N \<le> n"
500      hence dom: "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q"
501        using dim by (simp_all add: N_def Q_def dim_def Abs_fmap_inverse)
502      show "dist (P n) Q < e"
503      proof (rule dist_finmap_lessI[OF dom(3) \<open>0 < e\<close>])
504        fix i
505        assume "i \<in> domain (P n)"
506        hence "ni i \<le> Max (ni ` d)" using dom by simp
507        also have "\<dots> \<le> N" by (simp add: N_def)
508        finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni \<open>i \<in> domain (P n)\<close> \<open>N \<le> n\<close> dom
509          by (auto simp: p_def q N_def less_imp_le)
510      qed
511    qed
512  qed
513  thus "convergent P" by (auto simp: convergent_def)
514qed
515
516subsection \<open>Second Countable Space of Finite Maps\<close>
517
518instantiation fmap :: (countable, second_countable_topology) second_countable_topology
519begin
520
521definition basis_proj::"'b set set"
522  where "basis_proj = (SOME B. countable B \<and> topological_basis B)"
523
524lemma countable_basis_proj: "countable basis_proj" and basis_proj: "topological_basis basis_proj"
525  unfolding basis_proj_def by (intro is_basis countable_basis)+
526
527definition basis_finmap::"('a \<Rightarrow>\<^sub>F 'b) set set"
528  where "basis_finmap = {Pi' I S|I S. finite I \<and> (\<forall>i \<in> I. S i \<in> basis_proj)}"
529
530lemma in_basis_finmapI:
531  assumes "finite I" assumes "\<And>i. i \<in> I \<Longrightarrow> S i \<in> basis_proj"
532  shows "Pi' I S \<in> basis_finmap"
533  using assms unfolding basis_finmap_def by auto
534
535lemma basis_finmap_eq:
536  assumes "basis_proj \<noteq> {}"
537  shows "basis_finmap = (\<lambda>f. Pi' (domain f) (\<lambda>i. from_nat_into basis_proj ((f)\<^sub>F i))) `
538    (UNIV::('a \<Rightarrow>\<^sub>F nat) set)" (is "_ = ?f ` _")
539  unfolding basis_finmap_def
540proof safe
541  fix I::"'a set" and S::"'a \<Rightarrow> 'b set"
542  assume "finite I" "\<forall>i\<in>I. S i \<in> basis_proj"
543  hence "Pi' I S = ?f (finmap_of I (\<lambda>x. to_nat_on basis_proj (S x)))"
544    by (force simp: Pi'_def countable_basis_proj)
545  thus "Pi' I S \<in> range ?f" by simp
546next
547  fix x and f::"'a \<Rightarrow>\<^sub>F nat"
548  show "\<exists>I S. (\<Pi>' i\<in>domain f. from_nat_into basis_proj ((f)\<^sub>F i)) = Pi' I S \<and>
549    finite I \<and> (\<forall>i\<in>I. S i \<in> basis_proj)"
550    using assms by (auto intro: from_nat_into)
551qed
552
553lemma basis_finmap_eq_empty: "basis_proj = {} \<Longrightarrow> basis_finmap = {Pi' {} undefined}"
554  by (auto simp: Pi'_iff basis_finmap_def)
555
556lemma countable_basis_finmap: "countable basis_finmap"
557  by (cases "basis_proj = {}") (auto simp: basis_finmap_eq basis_finmap_eq_empty)
558
559lemma finmap_topological_basis:
560  "topological_basis basis_finmap"
561proof (subst topological_basis_iff, safe)
562  fix B' assume "B' \<in> basis_finmap"
563  thus "open B'"
564    by (auto intro!: open_Pi'I topological_basis_open[OF basis_proj]
565      simp: topological_basis_def basis_finmap_def Let_def)
566next
567  fix O'::"('a \<Rightarrow>\<^sub>F 'b) set" and x
568  assume O': "open O'" "x \<in> O'"
569  then obtain a where a:
570    "x \<in> Pi' (domain x) a" "Pi' (domain x) a \<subseteq> O'" "\<And>i. i\<in>domain x \<Longrightarrow> open (a i)"
571    unfolding open_fmap_def
572  proof (atomize_elim, induct rule: generate_topology.induct)
573    case (Int a b)
574    let ?p="\<lambda>a f. x \<in> Pi' (domain x) f \<and> Pi' (domain x) f \<subseteq> a \<and> (\<forall>i. i \<in> domain x \<longrightarrow> open (f i))"
575    from Int obtain f g where "?p a f" "?p b g" by auto
576    thus ?case by (force intro!: exI[where x="\<lambda>i. f i \<inter> g i"] simp: Pi'_def)
577  next
578    case (UN k)
579    then obtain kk a where "x \<in> kk" "kk \<in> k" "x \<in> Pi' (domain x) a" "Pi' (domain x) a \<subseteq> kk"
580      "\<And>i. i\<in>domain x \<Longrightarrow> open (a i)"
581      by force
582    thus ?case by blast
583  qed (auto simp: Pi'_def)
584  have "\<exists>B.
585    (\<forall>i\<in>domain x. x i \<in> B i \<and> B i \<subseteq> a i \<and> B i \<in> basis_proj)"
586  proof (rule bchoice, safe)
587    fix i assume "i \<in> domain x"
588    hence "open (a i)" "x i \<in> a i" using a by auto
589    from topological_basisE[OF basis_proj this] guess b' .
590    thus "\<exists>y. x i \<in> y \<and> y \<subseteq> a i \<and> y \<in> basis_proj" by auto
591  qed
592  then guess B .. note B = this
593  define B' where "B' = Pi' (domain x) (\<lambda>i. (B i)::'b set)"
594  have "B' \<subseteq> Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def)
595  also note \<open>\<dots> \<subseteq> O'\<close>
596  finally show "\<exists>B'\<in>basis_finmap. x \<in> B' \<and> B' \<subseteq> O'" using B
597    by (auto intro!: bexI[where x=B'] Pi'_mono in_basis_finmapI simp: B'_def)
598qed
599
600lemma range_enum_basis_finmap_imp_open:
601  assumes "x \<in> basis_finmap"
602  shows "open x"
603  using finmap_topological_basis assms by (auto simp: topological_basis_def)
604
605instance proof qed (blast intro: finmap_topological_basis countable_basis_finmap topological_basis_imp_subbasis)
606
607end
608
609subsection \<open>Polish Space of Finite Maps\<close>
610
611instance fmap :: (countable, polish_space) polish_space proof qed
612
613
614subsection \<open>Product Measurable Space of Finite Maps\<close>
615
616definition "PiF I M \<equiv>
617  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))}"
618
619abbreviation
620  "Pi\<^sub>F I M \<equiv> PiF I M"
621
622syntax
623  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^sub>F _\<in>_./ _)"  10)
624translations
625  "\<Pi>\<^sub>F x\<in>I. M" == "CONST PiF I (%x. M)"
626
627lemma 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>
628    Pow (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
629  by (auto simp: Pi'_def) (blast dest: sets.sets_into_space)
630
631lemma space_PiF: "space (PiF I M) = (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
632  unfolding PiF_def using PiF_gen_subset by (rule space_measure_of)
633
634lemma sets_PiF:
635  "sets (PiF I M) = sigma_sets (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
636    {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
637  unfolding PiF_def using PiF_gen_subset by (rule sets_measure_of)
638
639lemma sets_PiF_singleton:
640  "sets (PiF {I} M) = sigma_sets (\<Pi>' j\<in>I. space (M j))
641    {(\<Pi>' j\<in>I. X j) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
642  unfolding sets_PiF by simp
643
644lemma in_sets_PiFI:
645  assumes "X = (Pi' J S)" "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
646  shows "X \<in> sets (PiF I M)"
647  unfolding sets_PiF
648  using assms by blast
649
650lemma product_in_sets_PiFI:
651  assumes "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
652  shows "(Pi' J S) \<in> sets (PiF I M)"
653  unfolding sets_PiF
654  using assms by blast
655
656lemma singleton_space_subset_in_sets:
657  fixes J
658  assumes "J \<in> I"
659  assumes "finite J"
660  shows "space (PiF {J} M) \<in> sets (PiF I M)"
661  using assms
662  by (intro in_sets_PiFI[where J=J and S="\<lambda>i. space (M i)"])
663      (auto simp: product_def space_PiF)
664
665lemma singleton_subspace_set_in_sets:
666  assumes A: "A \<in> sets (PiF {J} M)"
667  assumes "finite J"
668  assumes "J \<in> I"
669  shows "A \<in> sets (PiF I M)"
670  using A[unfolded sets_PiF]
671  apply (induct A)
672  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
673  using assms
674  by (auto intro: in_sets_PiFI intro!: singleton_space_subset_in_sets)
675
676lemma finite_measurable_singletonI:
677  assumes "finite I"
678  assumes "\<And>J. J \<in> I \<Longrightarrow> finite J"
679  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
680  shows "A \<in> measurable (PiF I M) N"
681  unfolding measurable_def
682proof safe
683  fix y assume "y \<in> sets N"
684  have "A -` y \<inter> space (PiF I M) = (\<Union>J\<in>I. A -` y \<inter> space (PiF {J} M))"
685    by (auto simp: space_PiF)
686  also have "\<dots> \<in> sets (PiF I M)"
687  proof (rule sets.finite_UN)
688    show "finite I" by fact
689    fix J assume "J \<in> I"
690    with assms have "finite J" by simp
691    show "A -` y \<inter> space (PiF {J} M) \<in> sets (PiF I M)"
692      by (rule singleton_subspace_set_in_sets[OF measurable_sets[OF assms(3)]]) fact+
693  qed
694  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
695next
696  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
697    using MN[of "domain x"]
698    by (auto simp: space_PiF measurable_space Pi'_def)
699qed
700
701lemma countable_finite_comprehension:
702  fixes f :: "'a::countable set \<Rightarrow> _"
703  assumes "\<And>s. P s \<Longrightarrow> finite s"
704  assumes "\<And>s. P s \<Longrightarrow> f s \<in> sets M"
705  shows "\<Union>{f s|s. P s} \<in> sets M"
706proof -
707  have "\<Union>{f s|s. P s} = (\<Union>n::nat. let s = set (from_nat n) in if P s then f s else {})"
708  proof safe
709    fix x X s assume *: "x \<in> f s" "P s"
710    with assms obtain l where "s = set l" using finite_list by blast
711    with * show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using \<open>P s\<close>
712      by (auto intro!: exI[where x="to_nat l"])
713  next
714    fix x n assume "x \<in> (let s = set (from_nat n) in if P s then f s else {})"
715    thus "x \<in> \<Union>{f s|s. P s}" using assms by (auto simp: Let_def split: if_split_asm)
716  qed
717  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
718  also have "\<dots> \<in> sets M" using assms by (auto simp: Let_def)
719  finally show ?thesis .
720qed
721
722lemma space_subset_in_sets:
723  fixes J::"'a::countable set set"
724  assumes "J \<subseteq> I"
725  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
726  shows "space (PiF J M) \<in> sets (PiF I M)"
727proof -
728  have "space (PiF J M) = \<Union>{space (PiF {j} M)|j. j \<in> J}"
729    unfolding space_PiF by blast
730  also have "\<dots> \<in> sets (PiF I M)" using assms
731    by (intro countable_finite_comprehension) (auto simp: singleton_space_subset_in_sets)
732  finally show ?thesis .
733qed
734
735lemma subspace_set_in_sets:
736  fixes J::"'a::countable set set"
737  assumes A: "A \<in> sets (PiF J M)"
738  assumes "J \<subseteq> I"
739  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
740  shows "A \<in> sets (PiF I M)"
741  using A[unfolded sets_PiF]
742  apply (induct A)
743  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
744  using assms
745  by (auto intro: in_sets_PiFI intro!: space_subset_in_sets)
746
747lemma countable_measurable_PiFI:
748  fixes I::"'a::countable set set"
749  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> finite J \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
750  shows "A \<in> measurable (PiF I M) N"
751  unfolding measurable_def
752proof safe
753  fix y assume "y \<in> sets N"
754  have "A -` y = (\<Union>{A -` y \<inter> {x. domain x = J}|J. finite J})" by auto
755  { fix x::"'a \<Rightarrow>\<^sub>F 'b"
756    from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto
757    hence "\<exists>n. domain x = set (from_nat n)"
758      by (intro exI[where x="to_nat xs"]) auto }
759  hence "A -` y \<inter> space (PiF I M) = (\<Union>n. A -` y \<inter> space (PiF ({set (from_nat n)}\<inter>I) M))"
760    by (auto simp: space_PiF Pi'_def)
761  also have "\<dots> \<in> sets (PiF I M)"
762    apply (intro sets.Int sets.countable_nat_UN subsetI, safe)
763    apply (case_tac "set (from_nat i) \<in> I")
764    apply simp_all
765    apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]])
766    using assms \<open>y \<in> sets N\<close>
767    apply (auto simp: space_PiF)
768    done
769  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
770next
771  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
772    using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def)
773qed
774
775lemma measurable_PiF:
776  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))"
777  assumes S: "\<And>J S. J \<in> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> S i \<in> sets (M i)) \<Longrightarrow>
778    f -` (Pi' J S) \<inter> space N \<in> sets N"
779  shows "f \<in> measurable N (PiF I M)"
780  unfolding PiF_def
781  using PiF_gen_subset
782  apply (rule measurable_measure_of)
783  using f apply force
784  apply (insert S, auto)
785  done
786
787lemma restrict_sets_measurable:
788  assumes A: "A \<in> sets (PiF I M)" and "J \<subseteq> I"
789  shows "A \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
790  using A[unfolded sets_PiF]
791proof (induct A)
792  case (Basic a)
793  then obtain K S where S: "a = Pi' K S" "K \<in> I" "(\<forall>i\<in>K. S i \<in> sets (M i))"
794    by auto
795  show ?case
796  proof cases
797    assume "K \<in> J"
798    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
799      by (auto intro!: exI[where x=K] exI[where x=S] simp: Pi'_def)
800    also have "\<dots> \<subseteq> sets (PiF J M)" unfolding sets_PiF by auto
801    finally show ?thesis .
802  next
803    assume "K \<notin> J"
804    hence "a \<inter> {m. domain m \<in> J} = {}" using S by (auto simp: Pi'_def)
805    also have "\<dots> \<in> sets (PiF J M)" by simp
806    finally show ?thesis .
807  qed
808next
809  case (Union a)
810  have "\<Union>(a ` UNIV) \<inter> {m. domain m \<in> J} = (\<Union>i. (a i \<inter> {m. domain m \<in> J}))"
811    by simp
812  also have "\<dots> \<in> sets (PiF J M)" using Union by (intro sets.countable_nat_UN) auto
813  finally show ?case .
814next
815  case (Compl a)
816  have "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} = (space (PiF J M) - (a \<inter> {m. domain m \<in> J}))"
817    using \<open>J \<subseteq> I\<close> by (auto simp: space_PiF Pi'_def)
818  also have "\<dots> \<in> sets (PiF J M)" using Compl by auto
819  finally show ?case by (simp add: space_PiF)
820qed simp
821
822lemma measurable_finmap_of:
823  assumes f: "\<And>i. (\<exists>x \<in> space N. i \<in> J x) \<Longrightarrow> (\<lambda>x. f x i) \<in> measurable N (M i)"
824  assumes J: "\<And>x. x \<in> space N \<Longrightarrow> J x \<in> I" "\<And>x. x \<in> space N \<Longrightarrow> finite (J x)"
825  assumes JN: "\<And>S. {x. J x = S} \<inter> space N \<in> sets N"
826  shows "(\<lambda>x. finmap_of (J x) (f x)) \<in> measurable N (PiF I M)"
827proof (rule measurable_PiF)
828  fix x assume "x \<in> space N"
829  with J[of x] measurable_space[OF f]
830  show "domain (finmap_of (J x) (f x)) \<in> I \<and>
831        (\<forall>i\<in>domain (finmap_of (J x) (f x)). (finmap_of (J x) (f x)) i \<in> space (M i))"
832    by auto
833next
834  fix K S assume "K \<in> I" and *: "\<And>i. i \<in> K \<Longrightarrow> S i \<in> sets (M i)"
835  with J have eq: "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N =
836    (if \<exists>x \<in> space N. K = J x \<and> finite K then if K = {} then {x \<in> space N. J x = K}
837      else (\<Inter>i\<in>K. (\<lambda>x. f x i) -` S i \<inter> {x \<in> space N. J x = K}) else {})"
838    by (auto simp: Pi'_def)
839  have r: "{x \<in> space N. J x = K} = space N \<inter> ({x. J x = K} \<inter> space N)" by auto
840  show "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N \<in> sets N"
841    unfolding eq r
842    apply (simp del: INT_simps add: )
843    apply (intro conjI impI sets.finite_INT JN sets.Int[OF sets.top])
844    apply simp apply assumption
845    apply (subst Int_assoc[symmetric])
846    apply (rule sets.Int)
847    apply (intro measurable_sets[OF f] *) apply force apply assumption
848    apply (intro JN)
849    done
850qed
851
852lemma measurable_PiM_finmap_of:
853  assumes "finite J"
854  shows "finmap_of J \<in> measurable (Pi\<^sub>M J M) (PiF {J} M)"
855  apply (rule measurable_finmap_of)
856  apply (rule measurable_component_singleton)
857  apply simp
858  apply rule
859  apply (rule \<open>finite J\<close>)
860  apply simp
861  done
862
863lemma proj_measurable_singleton:
864  assumes "A \<in> sets (M i)"
865  shows "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
866proof cases
867  assume "i \<in> I"
868  hence "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) =
869    Pi' I (\<lambda>x. if x = i then A else space (M x))"
870    using sets.sets_into_space[OF ] \<open>A \<in> sets (M i)\<close> assms
871    by (auto simp: space_PiF Pi'_def)
872  thus ?thesis  using assms \<open>A \<in> sets (M i)\<close>
873    by (intro in_sets_PiFI) auto
874next
875  assume "i \<notin> I"
876  hence "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) =
877    (if undefined \<in> A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def)
878  thus ?thesis by simp
879qed
880
881lemma measurable_proj_singleton:
882  assumes "i \<in> I"
883  shows "(\<lambda>x. (x)\<^sub>F i) \<in> measurable (PiF {I} M) (M i)"
884  by (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
885     (insert \<open>i \<in> I\<close>, auto simp: space_PiF)
886
887lemma measurable_proj_countable:
888  fixes I::"'a::countable set set"
889  assumes "y \<in> space (M i)"
890  shows "(\<lambda>x. if i \<in> domain x then (x)\<^sub>F i else y) \<in> measurable (PiF I M) (M i)"
891proof (rule countable_measurable_PiFI)
892  fix J assume "J \<in> I" "finite J"
893  show "(\<lambda>x. if i \<in> domain x then x i else y) \<in> measurable (PiF {J} M) (M i)"
894    unfolding measurable_def
895  proof safe
896    fix z assume "z \<in> sets (M i)"
897    have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
898      (\<lambda>x. if i \<in> J then (x)\<^sub>F i else y) -` z \<inter> space (PiF {J} M)"
899      by (auto simp: space_PiF Pi'_def)
900    also have "\<dots> \<in> sets (PiF {J} M)" using \<open>z \<in> sets (M i)\<close> \<open>finite J\<close>
901      by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
902    finally show "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) \<in>
903      sets (PiF {J} M)" .
904  qed (insert \<open>y \<in> space (M i)\<close>, auto simp: space_PiF Pi'_def)
905qed
906
907lemma measurable_restrict_proj:
908  assumes "J \<in> II" "finite J"
909  shows "finmap_of J \<in> measurable (PiM J M) (PiF II M)"
910  using assms
911  by (intro measurable_finmap_of measurable_component_singleton) auto
912
913lemma measurable_proj_PiM:
914  fixes J K ::"'a::countable set" and I::"'a set set"
915  assumes "finite J" "J \<in> I"
916  assumes "x \<in> space (PiM J M)"
917  shows "proj \<in> measurable (PiF {J} M) (PiM J M)"
918proof (rule measurable_PiM_single)
919  show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^sub>E i \<in> J. space (M i))"
920    using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def)
921next
922  fix A i assume A: "i \<in> J" "A \<in> sets (M i)"
923  show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^sub>F i \<in> A} \<in> sets (PiF {J} M)"
924  proof
925    have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^sub>F i \<in> A} =
926      (\<lambda>\<omega>. (\<omega>)\<^sub>F i) -` A \<inter> space (PiF {J} M)" by auto
927    also have "\<dots> \<in> sets (PiF {J} M)"
928      using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM)
929    finally show ?thesis .
930  qed simp
931qed
932
933lemma space_PiF_singleton_eq_product:
934  assumes "finite I"
935  shows "space (PiF {I} M) = (\<Pi>' i\<in>I. space (M i))"
936  by (auto simp: product_def space_PiF assms)
937
938text \<open>adapted from @{thm sets_PiM_single}\<close>
939
940lemma sets_PiF_single:
941  assumes "finite I" "I \<noteq> {}"
942  shows "sets (PiF {I} M) =
943    sigma_sets (\<Pi>' i\<in>I. space (M i))
944      {{f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
945    (is "_ = sigma_sets ?\<Omega> ?R")
946  unfolding sets_PiF_singleton
947proof (rule sigma_sets_eqI)
948  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
949  fix A assume "A \<in> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
950  then obtain X where X: "A = Pi' I X" "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
951  show "A \<in> sigma_sets ?\<Omega> ?R"
952  proof -
953    from \<open>I \<noteq> {}\<close> X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
954      using sets.sets_into_space
955      by (auto simp: space_PiF product_def) blast
956    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
957      using X \<open>I \<noteq> {}\<close> assms by (intro R.finite_INT) (auto simp: space_PiF)
958    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
959  qed
960next
961  fix A assume "A \<in> ?R"
962  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)"
963    by auto
964  then have "A = (\<Pi>' j \<in> I. if j = i then B else space (M j))"
965    using sets.sets_into_space[OF A(3)]
966    apply (auto simp: Pi'_iff split: if_split_asm)
967    apply blast
968    done
969  also have "\<dots> \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
970    using A
971    by (intro sigma_sets.Basic )
972       (auto intro: exI[where x="\<lambda>j. if j = i then B else space (M j)"])
973  finally show "A \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" .
974qed
975
976text \<open>adapted from @{thm PiE_cong}\<close>
977
978lemma Pi'_cong:
979  assumes "finite I"
980  assumes "\<And>i. i \<in> I \<Longrightarrow> f i = g i"
981  shows "Pi' I f = Pi' I g"
982using assms by (auto simp: Pi'_def)
983
984text \<open>adapted from @{thm Pi_UN}\<close>
985
986lemma Pi'_UN:
987  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
988  assumes "finite I"
989  assumes mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
990  shows "(\<Union>n. Pi' I (A n)) = Pi' I (\<lambda>i. \<Union>n. A n i)"
991proof (intro set_eqI iffI)
992  fix f assume "f \<in> Pi' I (\<lambda>i. \<Union>n. A n i)"
993  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: \<open>finite I\<close> Pi'_def)
994  from bchoice[OF this(1)] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
995  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
996    using \<open>finite I\<close> finite_nat_set_iff_bounded_le[of "n`I"] by auto
997  have "f \<in> Pi' I (\<lambda>i. A k i)"
998  proof
999    fix i assume "i \<in> I"
1000    from mono[OF this, of "n i" k] k[OF this] n[OF this] \<open>domain f = I\<close> \<open>i \<in> I\<close>
1001    show "f i \<in> A k i " by (auto simp: \<open>finite I\<close>)
1002  qed (simp add: \<open>domain f = I\<close> \<open>finite I\<close>)
1003  then show "f \<in> (\<Union>n. Pi' I (A n))" by auto
1004qed (auto simp: Pi'_def \<open>finite I\<close>)
1005
1006text \<open>adapted from @{thm sets_PiM_sigma}\<close>
1007
1008lemma sigma_fprod_algebra_sigma_eq:
1009  fixes E :: "'i \<Rightarrow> 'a set set" and S :: "'i \<Rightarrow> nat \<Rightarrow> 'a set"
1010  assumes [simp]: "finite I" "I \<noteq> {}"
1011    and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
1012    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
1013  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
1014    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
1015  defines "P == { Pi' I F | F. \<forall>i\<in>I. F i \<in> E i }"
1016  shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P"
1017proof
1018  let ?P = "sigma (space (Pi\<^sub>F {I} M)) P"
1019  from \<open>finite I\<close>[THEN ex_bij_betw_finite_nat] guess T ..
1020  then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i"
1021    by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f simp del: \<open>finite I\<close>)
1022  have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>F {I} M))"
1023    using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq)
1024  then have space_P: "space ?P = (\<Pi>' i\<in>I. space (M i))"
1025    by (simp add: space_PiF)
1026  have "sets (PiF {I} M) =
1027      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)}"
1028    using sets_PiF_single[of I M] by (simp add: space_P)
1029  also have "\<dots> \<subseteq> sets (sigma (space (PiF {I} M)) P)"
1030  proof (safe intro!: sets.sigma_sets_subset)
1031    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
1032    have "(\<lambda>x. (x)\<^sub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
1033    proof (subst measurable_iff_measure_of)
1034      show "E i \<subseteq> Pow (space (M i))" using \<open>i \<in> I\<close> by fact
1035      from space_P \<open>i \<in> I\<close> show "(\<lambda>x. (x)\<^sub>F i) \<in> space ?P \<rightarrow> space (M i)"
1036        by auto
1037      show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
1038      proof
1039        fix A assume A: "A \<in> E i"
1040        from T have *: "\<exists>xs. length xs = card I
1041          \<and> (\<forall>j\<in>I. (g)\<^sub>F j \<in> (if i = j then A else S j (xs ! T j)))"
1042          if "domain g = I" and "\<forall>j\<in>I. (g)\<^sub>F j \<in> (if i = j then A else S j (f j))" for g f
1043          using that by (auto intro!: exI [of _ "map (\<lambda>n. f (the_inv_into I T n)) [0..<card I]"])
1044        from A have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
1045          using E_closed \<open>i \<in> I\<close> by (auto simp: space_P Pi_iff subset_eq split: if_split_asm)
1046        also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
1047          by (intro Pi'_cong) (simp_all add: S_union)
1048        also have "\<dots> = {g. domain g = I \<and> (\<exists>f. \<forall>j\<in>I. (g)\<^sub>F j \<in> (if i = j then A else S j (f j)))}"
1049          by (rule set_eqI) (simp del: if_image_distrib add: Pi'_iff bchoice_iff)
1050        also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>' j\<in>I. if i = j then A else S j (xs ! T j))"
1051          using * by (auto simp add: Pi'_iff split del: if_split)
1052        also have "\<dots> \<in> sets ?P"
1053        proof (safe intro!: sets.countable_UN)
1054          fix xs show "(\<Pi>' j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P"
1055            using A S_in_E
1056            by (simp add: P_closed)
1057               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"])
1058        qed
1059        finally show "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
1060          using P_closed by simp
1061      qed
1062    qed
1063    from measurable_sets[OF this, of A] A \<open>i \<in> I\<close> E_closed
1064    have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
1065      by (simp add: E_generates)
1066    also have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
1067      using P_closed by (auto simp: space_PiF)
1068    finally show "\<dots> \<in> sets ?P" .
1069  qed
1070  finally show "sets (PiF {I} M) \<subseteq> sigma_sets (space (PiF {I} M)) P"
1071    by (simp add: P_closed)
1072  show "sigma_sets (space (PiF {I} M)) P \<subseteq> sets (PiF {I} M)"
1073    using \<open>finite I\<close> \<open>I \<noteq> {}\<close>
1074    by (auto intro!: sets.sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def)
1075qed
1076
1077lemma product_open_generates_sets_PiF_single:
1078  assumes "I \<noteq> {}"
1079  assumes [simp]: "finite I"
1080  shows "sets (PiF {I} (\<lambda>_. borel::'b::second_countable_topology measure)) =
1081    sigma_sets (space (PiF {I} (\<lambda>_. borel))) {Pi' I F |F. (\<forall>i\<in>I. F i \<in> Collect open)}"
1082proof -
1083  from open_countable_basisE[OF open_UNIV] guess S::"'b set set" . note S = this
1084  show ?thesis
1085  proof (rule sigma_fprod_algebra_sigma_eq)
1086    show "finite I" by simp
1087    show "I \<noteq> {}" by fact
1088    define S' where "S' = from_nat_into S"
1089    show "(\<Union>j. S' j) = space borel"
1090      using S
1091      apply (auto simp add: from_nat_into countable_basis_proj S'_def basis_proj_def)
1092      apply (metis (lifting, mono_tags) UNIV_I UnionE basis_proj_def countable_basis_proj countable_subset from_nat_into_surj)
1093      done
1094    show "range S' \<subseteq> Collect open"
1095      using S
1096      apply (auto simp add: from_nat_into countable_basis_proj S'_def)
1097      apply (metis UNIV_not_empty Union_empty from_nat_into subsetD topological_basis_open[OF basis_proj] basis_proj_def)
1098      done
1099    show "Collect open \<subseteq> Pow (space borel)" by simp
1100    show "sets borel = sigma_sets (space borel) (Collect open)"
1101      by (simp add: borel_def)
1102  qed
1103qed
1104
1105lemma finmap_UNIV[simp]: "(\<Union>J\<in>Collect finite. \<Pi>' j\<in>J. UNIV) = UNIV" by auto
1106
1107lemma borel_eq_PiF_borel:
1108  shows "(borel :: ('i::countable \<Rightarrow>\<^sub>F 'a::polish_space) measure) =
1109    PiF (Collect finite) (\<lambda>_. borel :: 'a measure)"
1110  unfolding borel_def PiF_def
1111proof (rule measure_eqI, clarsimp, rule sigma_sets_eqI)
1112  fix a::"('i \<Rightarrow>\<^sub>F 'a) set" assume "a \<in> Collect open" hence "open a" by simp
1113  then obtain B' where B': "B'\<subseteq>basis_finmap" "a = \<Union>B'"
1114    using finmap_topological_basis by (force simp add: topological_basis_def)
1115  have "a \<in> sigma UNIV {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
1116    unfolding \<open>a = \<Union>B'\<close>
1117  proof (rule sets.countable_Union)
1118    from B' countable_basis_finmap show "countable B'" by (metis countable_subset)
1119  next
1120    show "B' \<subseteq> sets (sigma UNIV
1121      {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)})" (is "_ \<subseteq> sets ?s")
1122    proof
1123      fix x assume "x \<in> B'" with B' have "x \<in> basis_finmap" by auto
1124      then obtain J X where "x = Pi' J X" "finite J" "X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)"
1125        by (auto simp: basis_finmap_def topological_basis_open[OF basis_proj])
1126      thus "x \<in> sets ?s" by auto
1127    qed
1128  qed
1129  thus "a \<in> sigma_sets UNIV {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
1130    by simp
1131next
1132  fix b::"('i \<Rightarrow>\<^sub>F 'a) set"
1133  assume "b \<in> {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
1134  hence b': "b \<in> sets (Pi\<^sub>F (Collect finite) (\<lambda>_. borel))" by (auto simp: sets_PiF borel_def)
1135  let ?b = "\<lambda>J. b \<inter> {x. domain x = J}"
1136  have "b = \<Union>((\<lambda>J. ?b J) ` Collect finite)" by auto
1137  also have "\<dots> \<in> sets borel"
1138  proof (rule sets.countable_Union, safe)
1139    fix J::"'i set" assume "finite J"
1140    { assume ef: "J = {}"
1141      have "?b J \<in> sets borel"
1142      proof cases
1143        assume "?b J \<noteq> {}"
1144        then obtain f where "f \<in> b" "domain f = {}" using ef by auto
1145        hence "?b J = {f}" using \<open>J = {}\<close>
1146          by (auto simp: finmap_eq_iff)
1147        also have "{f} \<in> sets borel" by simp
1148        finally show ?thesis .
1149      qed simp
1150    } moreover {
1151      assume "J \<noteq> ({}::'i set)"
1152      have "(?b J) = b \<inter> {m. domain m \<in> {J}}" by auto
1153      also have "\<dots> \<in> sets (PiF {J} (\<lambda>_. borel))"
1154        using b' by (rule restrict_sets_measurable) (auto simp: \<open>finite J\<close>)
1155      also have "\<dots> = sigma_sets (space (PiF {J} (\<lambda>_. borel)))
1156        {Pi' (J) F |F. (\<forall>j\<in>J. F j \<in> Collect open)}"
1157        (is "_ = sigma_sets _ ?P")
1158       by (rule product_open_generates_sets_PiF_single[OF \<open>J \<noteq> {}\<close> \<open>finite J\<close>])
1159      also have "\<dots> \<subseteq> sigma_sets UNIV (Collect open)"
1160        by (intro sigma_sets_mono'') (auto intro!: open_Pi'I simp: space_PiF)
1161      finally have "(?b J) \<in> sets borel" by (simp add: borel_def)
1162    } ultimately show "(?b J) \<in> sets borel" by blast
1163  qed (simp add: countable_Collect_finite)
1164  finally show "b \<in> sigma_sets UNIV (Collect open)" by (simp add: borel_def)
1165qed (simp add: emeasure_sigma borel_def PiF_def)
1166
1167subsection \<open>Isomorphism between Functions and Finite Maps\<close>
1168
1169lemma measurable_finmap_compose:
1170  shows "(\<lambda>m. compose J m f) \<in> measurable (PiM (f ` J) (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
1171  unfolding compose_def by measurable
1172
1173lemma measurable_compose_inv:
1174  assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
1175  shows "(\<lambda>m. compose (f ` J) m f') \<in> measurable (PiM J (\<lambda>_. M)) (PiM (f ` J) (\<lambda>_. M))"
1176  unfolding compose_def by (rule measurable_restrict) (auto simp: inj)
1177
1178locale function_to_finmap =
1179  fixes J::"'a set" and f :: "'a \<Rightarrow> 'b::countable" and f'
1180  assumes [simp]: "finite J"
1181  assumes inv: "i \<in> J \<Longrightarrow> f' (f i) = i"
1182begin
1183
1184text \<open>to measure finmaps\<close>
1185
1186definition "fm = (finmap_of (f ` J)) o (\<lambda>g. compose (f ` J) g f')"
1187
1188lemma domain_fm[simp]: "domain (fm x) = f ` J"
1189  unfolding fm_def by simp
1190
1191lemma fm_restrict[simp]: "fm (restrict y J) = fm y"
1192  unfolding fm_def by (auto simp: compose_def inv intro: restrict_ext)
1193
1194lemma fm_product:
1195  assumes "\<And>i. space (M i) = UNIV"
1196  shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^sub>M J M) = (\<Pi>\<^sub>E j \<in> J. S (f j))"
1197  using assms
1198  by (auto simp: inv fm_def compose_def space_PiM Pi'_def)
1199
1200lemma fm_measurable:
1201  assumes "f ` J \<in> N"
1202  shows "fm \<in> measurable (Pi\<^sub>M J (\<lambda>_. M)) (Pi\<^sub>F N (\<lambda>_. M))"
1203  unfolding fm_def
1204proof (rule measurable_comp, rule measurable_compose_inv)
1205  show "finmap_of (f ` J) \<in> measurable (Pi\<^sub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
1206    using assms by (intro measurable_finmap_of measurable_component_singleton) auto
1207qed (simp_all add: inv)
1208
1209lemma proj_fm:
1210  assumes "x \<in> J"
1211  shows "fm m (f x) = m x"
1212  using assms by (auto simp: fm_def compose_def o_def inv)
1213
1214lemma inj_on_compose_f': "inj_on (\<lambda>g. compose (f ` J) g f') (extensional J)"
1215proof (rule inj_on_inverseI)
1216  fix x::"'a \<Rightarrow> 'c" assume "x \<in> extensional J"
1217  thus "(\<lambda>x. compose J x f) (compose (f ` J) x f') = x"
1218    by (auto simp: compose_def inv extensional_def)
1219qed
1220
1221lemma inj_on_fm:
1222  assumes "\<And>i. space (M i) = UNIV"
1223  shows "inj_on fm (space (Pi\<^sub>M J M))"
1224  using assms
1225  apply (auto simp: fm_def space_PiM PiE_def)
1226  apply (rule comp_inj_on)
1227  apply (rule inj_on_compose_f')
1228  apply (rule finmap_of_inj_on_extensional_finite)
1229  apply simp
1230  apply (auto)
1231  done
1232
1233text \<open>to measure functions\<close>
1234
1235definition "mf = (\<lambda>g. compose J g f) o proj"
1236
1237lemma mf_fm:
1238  assumes "x \<in> space (Pi\<^sub>M J (\<lambda>_. M))"
1239  shows "mf (fm x) = x"
1240proof -
1241  have "mf (fm x) \<in> extensional J"
1242    by (auto simp: mf_def extensional_def compose_def)
1243  moreover
1244  have "x \<in> extensional J" using assms sets.sets_into_space
1245    by (force simp: space_PiM PiE_def)
1246  moreover
1247  { fix i assume "i \<in> J"
1248    hence "mf (fm x) i = x i"
1249      by (auto simp: inv mf_def compose_def fm_def)
1250  }
1251  ultimately
1252  show ?thesis by (rule extensionalityI)
1253qed
1254
1255lemma mf_measurable:
1256  assumes "space M = UNIV"
1257  shows "mf \<in> measurable (PiF {f ` J} (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
1258  unfolding mf_def
1259proof (rule measurable_comp, rule measurable_proj_PiM)
1260  show "(\<lambda>g. compose J g f) \<in> measurable (Pi\<^sub>M (f ` J) (\<lambda>x. M)) (Pi\<^sub>M J (\<lambda>_. M))"
1261    by (rule measurable_finmap_compose)
1262qed (auto simp add: space_PiM extensional_def assms)
1263
1264lemma fm_image_measurable:
1265  assumes "space M = UNIV"
1266  assumes "X \<in> sets (Pi\<^sub>M J (\<lambda>_. M))"
1267  shows "fm ` X \<in> sets (PiF {f ` J} (\<lambda>_. M))"
1268proof -
1269  have "fm ` X = (mf) -` X \<inter> space (PiF {f ` J} (\<lambda>_. M))"
1270  proof safe
1271    fix x assume "x \<in> X"
1272    with mf_fm[of x] sets.sets_into_space[OF assms(2)] show "fm x \<in> mf -` X" by auto
1273    show "fm x \<in> space (PiF {f ` J} (\<lambda>_. M))" by (simp add: space_PiF assms)
1274  next
1275    fix y x
1276    assume x: "mf y \<in> X"
1277    assume y: "y \<in> space (PiF {f ` J} (\<lambda>_. M))"
1278    thus "y \<in> fm ` X"
1279      by (intro image_eqI[OF _ x], unfold finmap_eq_iff)
1280         (auto simp: space_PiF fm_def mf_def compose_def inv Pi'_def)
1281  qed
1282  also have "\<dots> \<in> sets (PiF {f ` J} (\<lambda>_. M))"
1283    using assms
1284    by (intro measurable_sets[OF mf_measurable]) auto
1285  finally show ?thesis .
1286qed
1287
1288lemma fm_image_measurable_finite:
1289  assumes "space M = UNIV"
1290  assumes "X \<in> sets (Pi\<^sub>M J (\<lambda>_. M::'c measure))"
1291  shows "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. M::'c measure))"
1292  using fm_image_measurable[OF assms]
1293  by (rule subspace_set_in_sets) (auto simp: finite_subset)
1294
1295text \<open>measure on finmaps\<close>
1296
1297definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"
1298
1299lemma sets_mapmeasure[simp]: "sets (mapmeasure M N) = sets (PiF (Collect finite) N)"
1300  unfolding mapmeasure_def by simp
1301
1302lemma space_mapmeasure[simp]: "space (mapmeasure M N) = space (PiF (Collect finite) N)"
1303  unfolding mapmeasure_def by simp
1304
1305lemma mapmeasure_PiF:
1306  assumes s1: "space M = space (Pi\<^sub>M J (\<lambda>_. N))"
1307  assumes s2: "sets M = sets (Pi\<^sub>M J (\<lambda>_. N))"
1308  assumes "space N = UNIV"
1309  assumes "X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
1310  shows "emeasure (mapmeasure M (\<lambda>_. N)) X = emeasure M ((fm -` X \<inter> extensional J))"
1311  using assms
1312  by (auto simp: measurable_cong_sets[OF s2 refl] mapmeasure_def emeasure_distr
1313    fm_measurable space_PiM PiE_def)
1314
1315lemma mapmeasure_PiM:
1316  fixes N::"'c measure"
1317  assumes s1: "space M = space (Pi\<^sub>M J (\<lambda>_. N))"
1318  assumes s2: "sets M = (Pi\<^sub>M J (\<lambda>_. N))"
1319  assumes N: "space N = UNIV"
1320  assumes X: "X \<in> sets M"
1321  shows "emeasure M X = emeasure (mapmeasure M (\<lambda>_. N)) (fm ` X)"
1322  unfolding mapmeasure_def
1323proof (subst emeasure_distr, subst measurable_cong_sets[OF s2 refl], rule fm_measurable)
1324  have "X \<subseteq> space (Pi\<^sub>M J (\<lambda>_. N))" using assms by (simp add: sets.sets_into_space)
1325  from assms inj_on_fm[of "\<lambda>_. N"] subsetD[OF this] have "fm -` fm ` X \<inter> space (Pi\<^sub>M J (\<lambda>_. N)) = X"
1326    by (auto simp: vimage_image_eq inj_on_def)
1327  thus "emeasure M X = emeasure M (fm -` fm ` X \<inter> space M)" using s1
1328    by simp
1329  show "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
1330    by (rule fm_image_measurable_finite[OF N X[simplified s2]])
1331qed simp
1332
1333end
1334
1335end
1336