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