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