| author | wenzelm | 
| Mon, 25 Oct 2010 21:23:09 +0200 | |
| changeset 40133 | b61d52de66f0 | 
| parent 39302 | d7728f65b353 | 
| permissions | -rw-r--r-- | 
| 38656 | 1 | (* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *) | 
| 2 | ||
| 3 | header {*Borel spaces*}
 | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 4 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 5 | theory Borel | 
| 38656 | 6 | imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 7 | begin | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 8 | |
| 39092 | 9 | lemma LIMSEQ_max: | 
| 10 | "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0" | |
| 11 | by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) | |
| 12 | ||
| 38656 | 13 | section "Generic Borel spaces" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 14 | |
| 38656 | 15 | definition "borel_space = sigma (UNIV::'a::topological_space set) open" | 
| 16 | abbreviation "borel_measurable M \<equiv> measurable M borel_space" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 17 | |
| 38656 | 18 | interpretation borel_space: sigma_algebra borel_space | 
| 19 | using sigma_algebra_sigma by (auto simp: borel_space_def) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 20 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 21 | lemma in_borel_measurable: | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 22 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 38656 | 23 | (\<forall>S \<in> sets (sigma UNIV open). | 
| 24 | f -` S \<inter> space M \<in> sets M)" | |
| 25 | by (auto simp add: measurable_def borel_space_def) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 26 | |
| 38656 | 27 | lemma in_borel_measurable_borel_space: | 
| 28 | "f \<in> borel_measurable M \<longleftrightarrow> | |
| 29 | (\<forall>S \<in> sets borel_space. | |
| 30 | f -` S \<inter> space M \<in> sets M)" | |
| 31 | by (auto simp add: measurable_def borel_space_def) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 32 | |
| 38656 | 33 | lemma space_borel_space[simp]: "space borel_space = UNIV" | 
| 34 | unfolding borel_space_def by auto | |
| 35 | ||
| 36 | lemma borel_space_open[simp]: | |
| 37 | assumes "open A" shows "A \<in> sets borel_space" | |
| 38 | proof - | |
| 39 | have "A \<in> open" unfolding mem_def using assms . | |
| 40 | thus ?thesis unfolding borel_space_def sigma_def by (auto intro!: sigma_sets.Basic) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 41 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 42 | |
| 38656 | 43 | lemma borel_space_closed[simp]: | 
| 44 | assumes "closed A" shows "A \<in> sets borel_space" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 45 | proof - | 
| 38656 | 46 | have "space borel_space - (- A) \<in> sets borel_space" | 
| 47 | using assms unfolding closed_def by (blast intro: borel_space_open) | |
| 48 | thus ?thesis by simp | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 49 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 50 | |
| 38656 | 51 | lemma (in sigma_algebra) borel_measurable_vimage: | 
| 52 | fixes f :: "'a \<Rightarrow> 'x::t2_space" | |
| 53 | assumes borel: "f \<in> borel_measurable M" | |
| 54 |   shows "f -` {x} \<inter> space M \<in> sets M"
 | |
| 55 | proof (cases "x \<in> f ` space M") | |
| 56 | case True then obtain y where "x = f y" by auto | |
| 57 | from closed_sing[of "f y"] | |
| 58 |   have "{f y} \<in> sets borel_space" by (rule borel_space_closed)
 | |
| 59 | with assms show ?thesis | |
| 60 | unfolding in_borel_measurable_borel_space `x = f y` by auto | |
| 61 | next | |
| 62 |   case False hence "f -` {x} \<inter> space M = {}" by auto
 | |
| 63 | thus ?thesis by auto | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 64 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 65 | |
| 38656 | 66 | lemma (in sigma_algebra) borel_measurableI: | 
| 67 | fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space" | |
| 68 | assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" | |
| 69 | shows "f \<in> borel_measurable M" | |
| 70 | unfolding borel_space_def | |
| 71 | proof (rule measurable_sigma) | |
| 72 | fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M" | |
| 73 | using assms[of S] by (simp add: mem_def) | |
| 74 | qed simp_all | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 75 | |
| 38656 | 76 | lemma borel_space_singleton[simp, intro]: | 
| 77 | fixes x :: "'a::t1_space" | |
| 78 | shows "A \<in> sets borel_space \<Longrightarrow> insert x A \<in> sets borel_space" | |
| 79 | proof (rule borel_space.insert_in_sets) | |
| 80 |     show "{x} \<in> sets borel_space"
 | |
| 81 | using closed_sing[of x] by (rule borel_space_closed) | |
| 82 | qed simp | |
| 83 | ||
| 84 | lemma (in sigma_algebra) borel_measurable_const[simp, intro]: | |
| 85 | "(\<lambda>x. c) \<in> borel_measurable M" | |
| 86 | by (auto intro!: measurable_const) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 87 | |
| 39083 | 88 | lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]: | 
| 38656 | 89 | assumes A: "A \<in> sets M" | 
| 90 | shows "indicator A \<in> borel_measurable M" | |
| 91 | unfolding indicator_def_raw using A | |
| 92 | by (auto intro!: measurable_If_set borel_measurable_const) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 93 | |
| 38656 | 94 | lemma borel_measurable_translate: | 
| 95 | assumes "A \<in> sets borel_space" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel_space" | |
| 96 | shows "f -` A \<in> sets borel_space" | |
| 97 | proof - | |
| 98 | have "A \<in> sigma_sets UNIV open" using assms | |
| 99 | by (simp add: borel_space_def sigma_def) | |
| 100 | thus ?thesis | |
| 101 | proof (induct rule: sigma_sets.induct) | |
| 102 | case (Basic a) thus ?case using trans[of a] by (simp add: mem_def) | |
| 103 | next | |
| 104 | case (Compl a) | |
| 105 | moreover have "UNIV \<in> sets borel_space" | |
| 106 | by (metis borel_space.top borel_space_def_raw mem_def space_sigma) | |
| 107 | ultimately show ?case | |
| 108 | by (auto simp: vimage_Diff borel_space.Diff) | |
| 109 | qed (auto simp add: vimage_UN) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 110 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 111 | |
| 39092 | 112 | lemma (in sigma_algebra) borel_measurable_restricted: | 
| 113 |   fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
 | |
| 114 | shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow> | |
| 115 | (\<lambda>x. f x * indicator A x) \<in> borel_measurable M" | |
| 116 | (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M") | |
| 117 | proof - | |
| 118 | interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) | |
| 119 | have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R" | |
| 120 | by (auto intro!: measurable_cong) | |
| 121 | show ?thesis unfolding * | |
| 122 | unfolding in_borel_measurable_borel_space | |
| 123 | proof (simp, safe) | |
| 124 | fix S :: "'x set" assume "S \<in> sets borel_space" | |
| 125 | "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M" | |
| 126 | then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto | |
| 127 | then have f: "?f -` S \<inter> A \<in> sets M" | |
| 128 | using `A \<in> sets M` sets_into_space by fastsimp | |
| 129 | show "?f -` S \<inter> space M \<in> sets M" | |
| 130 | proof cases | |
| 131 | assume "0 \<in> S" | |
| 132 | then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)" | |
| 133 | using `A \<in> sets M` sets_into_space by auto | |
| 134 | then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff) | |
| 135 | next | |
| 136 | assume "0 \<notin> S" | |
| 137 | then have "?f -` S \<inter> space M = ?f -` S \<inter> A" | |
| 138 | using `A \<in> sets M` sets_into_space | |
| 139 | by (auto simp: indicator_def split: split_if_asm) | |
| 140 | then show ?thesis using f by auto | |
| 141 | qed | |
| 142 | next | |
| 143 | fix S :: "'x set" assume "S \<in> sets borel_space" | |
| 144 | "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M" | |
| 145 | then have f: "?f -` S \<inter> space M \<in> sets M" by auto | |
| 146 | then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M" | |
| 147 | using `A \<in> sets M` sets_into_space | |
| 148 | apply (simp add: image_iff) | |
| 149 | apply (rule bexI[OF _ f]) | |
| 150 | by auto | |
| 151 | qed | |
| 152 | qed | |
| 153 | ||
| 154 | lemma (in sigma_algebra) borel_measurable_subalgebra: | |
| 155 | assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)" | |
| 156 | shows "f \<in> borel_measurable M" | |
| 157 | using assms unfolding measurable_def by auto | |
| 158 | ||
| 38656 | 159 | section "Borel spaces on euclidean spaces" | 
| 160 | ||
| 161 | lemma lessThan_borel[simp, intro]: | |
| 162 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 163 |   shows "{..< a} \<in> sets borel_space"
 | |
| 164 | by (blast intro: borel_space_open) | |
| 165 | ||
| 166 | lemma greaterThan_borel[simp, intro]: | |
| 167 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 168 |   shows "{a <..} \<in> sets borel_space"
 | |
| 169 | by (blast intro: borel_space_open) | |
| 170 | ||
| 171 | lemma greaterThanLessThan_borel[simp, intro]: | |
| 172 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 173 |   shows "{a<..<b} \<in> sets borel_space"
 | |
| 174 | by (blast intro: borel_space_open) | |
| 175 | ||
| 176 | lemma atMost_borel[simp, intro]: | |
| 177 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 178 |   shows "{..a} \<in> sets borel_space"
 | |
| 179 | by (blast intro: borel_space_closed) | |
| 180 | ||
| 181 | lemma atLeast_borel[simp, intro]: | |
| 182 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 183 |   shows "{a..} \<in> sets borel_space"
 | |
| 184 | by (blast intro: borel_space_closed) | |
| 185 | ||
| 186 | lemma atLeastAtMost_borel[simp, intro]: | |
| 187 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 188 |   shows "{a..b} \<in> sets borel_space"
 | |
| 189 | by (blast intro: borel_space_closed) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 190 | |
| 38656 | 191 | lemma greaterThanAtMost_borel[simp, intro]: | 
| 192 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 193 |   shows "{a<..b} \<in> sets borel_space"
 | |
| 194 | unfolding greaterThanAtMost_def by blast | |
| 195 | ||
| 196 | lemma atLeastLessThan_borel[simp, intro]: | |
| 197 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 198 |   shows "{a..<b} \<in> sets borel_space"
 | |
| 199 | unfolding atLeastLessThan_def by blast | |
| 200 | ||
| 201 | lemma hafspace_less_borel[simp, intro]: | |
| 202 | fixes a :: real | |
| 203 |   shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel_space"
 | |
| 204 | by (auto intro!: borel_space_open open_halfspace_component_gt) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 205 | |
| 38656 | 206 | lemma hafspace_greater_borel[simp, intro]: | 
| 207 | fixes a :: real | |
| 208 |   shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel_space"
 | |
| 209 | by (auto intro!: borel_space_open open_halfspace_component_lt) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 210 | |
| 38656 | 211 | lemma hafspace_less_eq_borel[simp, intro]: | 
| 212 | fixes a :: real | |
| 213 |   shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel_space"
 | |
| 214 | by (auto intro!: borel_space_closed closed_halfspace_component_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 215 | |
| 38656 | 216 | lemma hafspace_greater_eq_borel[simp, intro]: | 
| 217 | fixes a :: real | |
| 218 |   shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel_space"
 | |
| 219 | by (auto intro!: borel_space_closed closed_halfspace_component_le) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 220 | |
| 38656 | 221 | lemma (in sigma_algebra) borel_measurable_less[simp, intro]: | 
| 222 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 223 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 224 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 225 |   shows "{w \<in> space M. f w < g w} \<in> sets M"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 226 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 227 |   have "{w \<in> space M. f w < g w} =
 | 
| 38656 | 228 |         (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
 | 
| 229 | using Rats_dense_in_real by (auto simp add: Rats_def) | |
| 230 | then show ?thesis using f g | |
| 231 | by simp (blast intro: measurable_sets) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 232 | qed | 
| 38656 | 233 | |
| 234 | lemma (in sigma_algebra) borel_measurable_le[simp, intro]: | |
| 235 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 236 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 237 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 238 |   shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 239 | proof - | 
| 38656 | 240 |   have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
 | 
| 241 | by auto | |
| 242 | thus ?thesis using f g | |
| 243 | by simp blast | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 244 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 245 | |
| 38656 | 246 | lemma (in sigma_algebra) borel_measurable_eq[simp, intro]: | 
| 247 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 248 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 249 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 250 |   shows "{w \<in> space M. f w = g w} \<in> sets M"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 251 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 252 |   have "{w \<in> space M. f w = g w} =
 | 
| 33536 | 253 |         {w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 254 | by auto | 
| 38656 | 255 | thus ?thesis using f g by auto | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 256 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 257 | |
| 38656 | 258 | lemma (in sigma_algebra) borel_measurable_neq[simp, intro]: | 
| 259 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 260 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 261 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 262 |   shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 263 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 264 |   have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 265 | by auto | 
| 38656 | 266 | thus ?thesis using f g by auto | 
| 267 | qed | |
| 268 | ||
| 269 | subsection "Borel space equals sigma algebras over intervals" | |
| 270 | ||
| 271 | lemma rational_boxes: | |
| 272 | fixes x :: "'a\<Colon>ordered_euclidean_space" | |
| 273 | assumes "0 < e" | |
| 274 |   shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
 | |
| 275 | proof - | |
| 276 |   def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
 | |
| 277 | then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos) | |
| 278 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i") | |
| 279 | proof | |
| 280 | fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e | |
| 281 | show "?th i" by auto | |
| 282 | qed | |
| 283 | from choice[OF this] guess a .. note a = this | |
| 284 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i") | |
| 285 | proof | |
| 286 | fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e | |
| 287 | show "?th i" by auto | |
| 288 | qed | |
| 289 | from choice[OF this] guess b .. note b = this | |
| 290 |   { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
 | |
| 291 |     have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
 | |
| 292 | unfolding setL2_def[symmetric] by (rule euclidean_dist_l2) | |
| 293 |     also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
 | |
| 294 | proof (rule real_sqrt_less_mono, rule setsum_strict_mono) | |
| 295 |       fix i assume i: "i \<in> {..<DIM('a)}"
 | |
| 296 | have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto | |
| 297 | moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto | |
| 298 | moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto | |
| 299 | ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto | |
| 300 |       then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
 | |
| 301 | unfolding e'_def by (auto simp: dist_real_def) | |
| 302 |       then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
 | |
| 303 | by (rule power_strict_mono) auto | |
| 304 |       then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
 | |
| 305 | by (simp add: power_divide) | |
| 306 | qed auto | |
| 307 | also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive) | |
| 308 | finally have "dist x y < e" . } | |
| 309 | with a b show ?thesis | |
| 310 | apply (rule_tac exI[of _ "Chi a"]) | |
| 311 | apply (rule_tac exI[of _ "Chi b"]) | |
| 312 | using eucl_less[where 'a='a] by auto | |
| 313 | qed | |
| 314 | ||
| 315 | lemma ex_rat_list: | |
| 316 | fixes x :: "'a\<Colon>ordered_euclidean_space" | |
| 317 | assumes "\<And> i. x $$ i \<in> \<rat>" | |
| 318 |   shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
 | |
| 319 | proof - | |
| 320 | have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast | |
| 321 | from choice[OF this] guess r .. | |
| 322 |   then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
 | |
| 323 | qed | |
| 324 | ||
| 325 | lemma open_UNION: | |
| 326 | fixes M :: "'a\<Colon>ordered_euclidean_space set" | |
| 327 | assumes "open M" | |
| 328 |   shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
 | |
| 329 |                    (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
 | |
| 330 | (is "M = UNION ?idx ?box") | |
| 331 | proof safe | |
| 332 | fix x assume "x \<in> M" | |
| 333 | obtain e where e: "e > 0" "ball x e \<subseteq> M" | |
| 334 | using openE[OF assms `x \<in> M`] by auto | |
| 335 |   then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
 | |
| 336 | using rational_boxes[OF e(1)] by blast | |
| 337 |   then obtain p q where pq: "length p = DIM ('a)"
 | |
| 338 |                             "length q = DIM ('a)"
 | |
| 339 |                             "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
 | |
| 340 | using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast | |
| 341 | hence p: "Chi (of_rat \<circ> op ! p) = a" | |
| 342 | using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a] | |
| 343 | unfolding o_def by auto | |
| 344 | from pq have q: "Chi (of_rat \<circ> op ! q) = b" | |
| 345 | using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b] | |
| 346 | unfolding o_def by auto | |
| 347 | have "x \<in> ?box (p, q)" | |
| 348 | using p q ab by auto | |
| 349 | thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto | |
| 350 | qed auto | |
| 351 | ||
| 352 | lemma halfspace_span_open: | |
| 353 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))
 | |
| 354 | \<subseteq> sets borel_space" | |
| 355 | by (auto intro!: borel_space.sigma_sets_subset[simplified] borel_space_open | |
| 356 | open_halfspace_component_lt simp: sets_sigma) | |
| 357 | ||
| 358 | lemma halfspace_lt_in_halfspace: | |
| 359 |   "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
 | |
| 360 | unfolding sets_sigma by (rule sigma_sets.Basic) auto | |
| 361 | ||
| 362 | lemma halfspace_gt_in_halfspace: | |
| 363 |   "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
 | |
| 364 | (is "?set \<in> sets ?SIGMA") | |
| 365 | proof - | |
| 366 | interpret sigma_algebra ?SIGMA by (rule sigma_algebra_sigma) simp | |
| 367 |   have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
 | |
| 368 | proof (safe, simp_all add: not_less) | |
| 369 | fix x assume "a < x $$ i" | |
| 370 | with reals_Archimedean[of "x $$ i - a"] | |
| 371 | obtain n where "a + 1 / real (Suc n) < x $$ i" | |
| 372 | by (auto simp: inverse_eq_divide field_simps) | |
| 373 | then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i" | |
| 374 | by (blast intro: less_imp_le) | |
| 375 | next | |
| 376 | fix x n | |
| 377 | have "a < a + 1 / real (Suc n)" by auto | |
| 378 | also assume "\<dots> \<le> x" | |
| 379 | finally show "a < x" . | |
| 380 | qed | |
| 381 | show "?set \<in> sets ?SIGMA" unfolding * | |
| 382 | by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 383 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 384 | |
| 38656 | 385 | lemma (in sigma_algebra) sets_sigma_subset: | 
| 386 | assumes "A = space M" | |
| 387 | assumes "B \<subseteq> sets M" | |
| 388 | shows "sets (sigma A B) \<subseteq> sets M" | |
| 389 | by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms) | |
| 390 | ||
| 391 | lemma open_span_halfspace: | |
| 392 |   "sets borel_space \<subseteq> sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})))"
 | |
| 393 | (is "_ \<subseteq> sets ?SIGMA") | |
| 394 | proof (unfold borel_space_def, rule sigma_algebra.sets_sigma_subset, safe) | |
| 395 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp | |
| 396 | then interpret sigma_algebra ?SIGMA . | |
| 397 | fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def . | |
| 398 | from open_UNION[OF this] | |
| 399 | obtain I where *: "S = | |
| 400 | (\<Union>(a, b)\<in>I. | |
| 401 |         (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
 | |
| 402 |         (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
 | |
| 403 | unfolding greaterThanLessThan_def | |
| 404 | unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] | |
| 405 | unfolding eucl_lessThan_eq_halfspaces[where 'a='a] | |
| 406 | by blast | |
| 407 | show "S \<in> sets ?SIGMA" | |
| 408 | unfolding * | |
| 409 | by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace) | |
| 410 | qed auto | |
| 411 | ||
| 412 | lemma halfspace_span_halfspace_le: | |
| 413 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq>
 | |
| 414 |    sets (sigma UNIV (range (\<lambda> (a, i). {x. x $$ i \<le> a})))"
 | |
| 415 | (is "_ \<subseteq> sets ?SIGMA") | |
| 416 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 417 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 418 | then interpret sigma_algebra ?SIGMA . | |
| 419 | fix a i | |
| 420 |   have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
 | |
| 421 | proof (safe, simp_all) | |
| 422 | fix x::'a assume *: "x$$i < a" | |
| 423 | with reals_Archimedean[of "a - x$$i"] | |
| 424 | obtain n where "x $$ i < a - 1 / (real (Suc n))" | |
| 425 | by (auto simp: field_simps inverse_eq_divide) | |
| 426 | then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))" | |
| 427 | by (blast intro: less_imp_le) | |
| 428 | next | |
| 429 | fix x::'a and n | |
| 430 | assume "x$$i \<le> a - 1 / real (Suc n)" | |
| 431 | also have "\<dots> < a" by auto | |
| 432 | finally show "x$$i < a" . | |
| 433 | qed | |
| 434 |   show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | |
| 435 | by (safe intro!: countable_UN) | |
| 436 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 437 | qed auto | |
| 438 | ||
| 439 | lemma halfspace_span_halfspace_ge: | |
| 440 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq> 
 | |
| 441 |    sets (sigma UNIV (range (\<lambda> (a, i). {x. a \<le> x $$ i})))"
 | |
| 442 | (is "_ \<subseteq> sets ?SIGMA") | |
| 443 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 444 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 445 | then interpret sigma_algebra ?SIGMA . | |
| 446 |   fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
 | |
| 447 |   show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | |
| 448 | by (safe intro!: Diff) | |
| 449 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 450 | qed auto | |
| 451 | ||
| 452 | lemma halfspace_le_span_halfspace_gt: | |
| 453 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq> 
 | |
| 454 |    sets (sigma UNIV (range (\<lambda> (a, i). {x. a < x $$ i})))"
 | |
| 455 | (is "_ \<subseteq> sets ?SIGMA") | |
| 456 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 457 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 458 | then interpret sigma_algebra ?SIGMA . | |
| 459 |   fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | |
| 460 |   show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | |
| 461 | by (safe intro!: Diff) | |
| 462 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 463 | qed auto | |
| 464 | ||
| 465 | lemma halfspace_le_span_atMost: | |
| 466 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
 | |
| 467 |    sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})))"
 | |
| 468 | (is "_ \<subseteq> sets ?SIGMA") | |
| 469 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 470 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 471 | then interpret sigma_algebra ?SIGMA . | |
| 472 | fix a i | |
| 473 |   show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 474 | proof cases | |
| 475 |     assume "i < DIM('a)"
 | |
| 476 |     then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
 | |
| 477 | proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) | |
| 478 | fix x | |
| 479 |       from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
 | |
| 480 |       then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
 | |
| 481 | by (subst (asm) Max_le_iff) auto | |
| 482 |       then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
 | |
| 483 | by (auto intro!: exI[of _ k]) | |
| 484 | qed | |
| 485 |     show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | |
| 486 | by (safe intro!: countable_UN) | |
| 487 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 488 | next | |
| 489 |     assume "\<not> i < DIM('a)"
 | |
| 490 |     then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 491 | using top by auto | |
| 492 | qed | |
| 493 | qed auto | |
| 494 | ||
| 495 | lemma halfspace_le_span_greaterThan: | |
| 496 |   "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
 | |
| 497 |    sets (sigma UNIV (range (\<lambda>a. {a<..})))"
 | |
| 498 | (is "_ \<subseteq> sets ?SIGMA") | |
| 499 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 500 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 501 | then interpret sigma_algebra ?SIGMA . | |
| 502 | fix a i | |
| 503 |   show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 504 | proof cases | |
| 505 |     assume "i < DIM('a)"
 | |
| 506 |     have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | |
| 507 |     also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
 | |
| 508 | proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) | |
| 509 | fix x | |
| 510 |       from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
 | |
| 511 | guess k::nat .. note k = this | |
| 512 |       { fix i assume "i < DIM('a)"
 | |
| 513 | then have "-x$$i < real k" | |
| 514 | using k by (subst (asm) Max_less_iff) auto | |
| 515 | then have "- real k < x$$i" by simp } | |
| 516 |       then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
 | |
| 517 | by (auto intro!: exI[of _ k]) | |
| 518 | qed | |
| 519 |     finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 520 | apply (simp only:) | |
| 521 | apply (safe intro!: countable_UN Diff) | |
| 522 | by (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 523 | next | |
| 524 |     assume "\<not> i < DIM('a)"
 | |
| 525 |     then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 526 | using top by auto | |
| 527 | qed | |
| 528 | qed auto | |
| 529 | ||
| 530 | lemma atMost_span_atLeastAtMost: | |
| 531 |   "sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))) \<subseteq>
 | |
| 532 |    sets (sigma UNIV (range (\<lambda>(a,b). {a..b})))"
 | |
| 533 | (is "_ \<subseteq> sets ?SIGMA") | |
| 534 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 535 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 536 | then interpret sigma_algebra ?SIGMA . | |
| 537 | fix a::'a | |
| 538 |   have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
 | |
| 539 | proof (safe, simp_all add: eucl_le[where 'a='a]) | |
| 540 | fix x | |
| 541 |     from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
 | |
| 542 | guess k::nat .. note k = this | |
| 543 |     { fix i assume "i < DIM('a)"
 | |
| 544 | with k have "- x$$i \<le> real k" | |
| 545 | by (subst (asm) Max_le_iff) (auto simp: field_simps) | |
| 546 | then have "- real k \<le> x$$i" by simp } | |
| 547 |     then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
 | |
| 548 | by (auto intro!: exI[of _ k]) | |
| 549 | qed | |
| 550 |   show "{..a} \<in> sets ?SIGMA" unfolding *
 | |
| 551 | by (safe intro!: countable_UN) | |
| 552 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 553 | qed auto | |
| 554 | ||
| 555 | lemma borel_space_eq_greaterThanLessThan: | |
| 556 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})))"
 | |
| 557 | (is "_ = sets ?SIGMA") | |
| 558 | proof (rule antisym) | |
| 559 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 560 | by (rule borel_space.sets_sigma_subset) auto | |
| 561 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 562 | unfolding borel_space_def | |
| 563 | proof (rule sigma_algebra.sets_sigma_subset, safe) | |
| 564 | show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 565 | then interpret sigma_algebra ?SIGMA . | |
| 566 | fix M :: "'a set" assume "M \<in> open" | |
| 567 | then have "open M" by (simp add: mem_def) | |
| 568 | show "M \<in> sets ?SIGMA" | |
| 569 | apply (subst open_UNION[OF `open M`]) | |
| 570 | apply (safe intro!: countable_UN) | |
| 571 | by (auto simp add: sigma_def intro!: sigma_sets.Basic) | |
| 572 | qed auto | |
| 573 | qed | |
| 574 | ||
| 575 | lemma borel_space_eq_atMost: | |
| 576 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})))"
 | |
| 577 | (is "_ = sets ?SIGMA") | |
| 578 | proof (rule antisym) | |
| 579 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 580 | using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace | |
| 581 | by auto | |
| 582 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 583 | by (rule borel_space.sets_sigma_subset) auto | |
| 584 | qed | |
| 585 | ||
| 586 | lemma borel_space_eq_atLeastAtMost: | |
| 587 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})))"
 | |
| 588 | (is "_ = sets ?SIGMA") | |
| 589 | proof (rule antisym) | |
| 590 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 591 | using atMost_span_atLeastAtMost halfspace_le_span_atMost | |
| 592 | halfspace_span_halfspace_le open_span_halfspace | |
| 593 | by auto | |
| 594 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 595 | by (rule borel_space.sets_sigma_subset) auto | |
| 596 | qed | |
| 597 | ||
| 598 | lemma borel_space_eq_greaterThan: | |
| 599 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})))"
 | |
| 600 | (is "_ = sets ?SIGMA") | |
| 601 | proof (rule antisym) | |
| 602 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 603 | using halfspace_le_span_greaterThan | |
| 604 | halfspace_span_halfspace_le open_span_halfspace | |
| 605 | by auto | |
| 606 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 607 | by (rule borel_space.sets_sigma_subset) auto | |
| 608 | qed | |
| 609 | ||
| 610 | lemma borel_space_eq_halfspace_le: | |
| 611 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})))"
 | |
| 612 | (is "_ = sets ?SIGMA") | |
| 613 | proof (rule antisym) | |
| 614 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 615 | using open_span_halfspace halfspace_span_halfspace_le by auto | |
| 616 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 617 | by (rule borel_space.sets_sigma_subset) auto | |
| 618 | qed | |
| 619 | ||
| 620 | lemma borel_space_eq_halfspace_less: | |
| 621 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})))"
 | |
| 622 | (is "_ = sets ?SIGMA") | |
| 623 | proof (rule antisym) | |
| 624 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 625 | using open_span_halfspace . | |
| 626 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 627 | by (rule borel_space.sets_sigma_subset) auto | |
| 628 | qed | |
| 629 | ||
| 630 | lemma borel_space_eq_halfspace_gt: | |
| 631 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})))"
 | |
| 632 | (is "_ = sets ?SIGMA") | |
| 633 | proof (rule antisym) | |
| 634 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 635 | using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto | |
| 636 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 637 | by (rule borel_space.sets_sigma_subset) auto | |
| 638 | qed | |
| 639 | ||
| 640 | lemma borel_space_eq_halfspace_ge: | |
| 641 |   "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})))"
 | |
| 642 | (is "_ = sets ?SIGMA") | |
| 643 | proof (rule antisym) | |
| 644 | show "sets borel_space \<subseteq> sets ?SIGMA" | |
| 645 | using halfspace_span_halfspace_ge open_span_halfspace by auto | |
| 646 | show "sets ?SIGMA \<subseteq> sets borel_space" | |
| 647 | by (rule borel_space.sets_sigma_subset) auto | |
| 648 | qed | |
| 649 | ||
| 650 | lemma (in sigma_algebra) borel_measurable_halfspacesI: | |
| 651 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 652 | assumes "sets borel_space = sets (sigma UNIV (range F))" | |
| 653 | and "\<And>a i. S a i = f -` F (a,i) \<inter> space M" | |
| 654 |   and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
 | |
| 655 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
 | |
| 656 | proof safe | |
| 657 |   fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
 | |
| 658 | then show "S a i \<in> sets M" unfolding assms | |
| 659 | by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def) | |
| 660 | next | |
| 661 |   assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
 | |
| 662 |   { fix a i have "S a i \<in> sets M"
 | |
| 663 | proof cases | |
| 664 |       assume "i < DIM('c)"
 | |
| 665 | with a show ?thesis unfolding assms(2) by simp | |
| 666 | next | |
| 667 |       assume "\<not> i < DIM('c)"
 | |
| 668 | from assms(3)[OF this] show ?thesis . | |
| 669 | qed } | |
| 670 | then have "f \<in> measurable M (sigma UNIV (range F))" | |
| 671 | by (auto intro!: measurable_sigma simp: assms(2)) | |
| 672 | then show "f \<in> borel_measurable M" unfolding measurable_def | |
| 673 | unfolding assms(1) by simp | |
| 674 | qed | |
| 675 | ||
| 676 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_le: | |
| 677 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 678 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
 | |
| 679 | by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_le]) auto | |
| 680 | ||
| 681 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_less: | |
| 682 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 683 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
 | |
| 684 | by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_less]) auto | |
| 685 | ||
| 686 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge: | |
| 687 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 688 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
 | |
| 689 | by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_ge]) auto | |
| 690 | ||
| 691 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater: | |
| 692 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 693 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
 | |
| 694 | by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_gt]) auto | |
| 695 | ||
| 696 | lemma (in sigma_algebra) borel_measurable_iff_le: | |
| 697 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
 | |
| 698 | using borel_measurable_iff_halfspace_le[where 'c=real] by simp | |
| 699 | ||
| 700 | lemma (in sigma_algebra) borel_measurable_iff_less: | |
| 701 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
 | |
| 702 | using borel_measurable_iff_halfspace_less[where 'c=real] by simp | |
| 703 | ||
| 704 | lemma (in sigma_algebra) borel_measurable_iff_ge: | |
| 705 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
 | |
| 706 | using borel_measurable_iff_halfspace_ge[where 'c=real] by simp | |
| 707 | ||
| 708 | lemma (in sigma_algebra) borel_measurable_iff_greater: | |
| 709 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
 | |
| 710 | using borel_measurable_iff_halfspace_greater[where 'c=real] by simp | |
| 711 | ||
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 712 | lemma borel_measureable_euclidean_component: | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 713 | "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel_space" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 714 | unfolding borel_space_def[where 'a=real] | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 715 | proof (rule borel_space.measurable_sigma) | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 716 | fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def . | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 717 | from open_vimage_euclidean_component[OF this] | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 718 | show "(\<lambda>x. x $$ i) -` S \<inter> space borel_space \<in> sets borel_space" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 719 | by (auto intro: borel_space_open) | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 720 | qed auto | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 721 | |
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 722 | lemma (in sigma_algebra) borel_measureable_euclidean_space: | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 723 | fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 724 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
 | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 725 | proof safe | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 726 | fix i assume "f \<in> borel_measurable M" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 727 | then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 728 | using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def] | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 729 | by (auto intro: borel_measureable_euclidean_component) | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 730 | next | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 731 |   assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
 | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 732 | then show "f \<in> borel_measurable M" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 733 | unfolding borel_measurable_iff_halfspace_le by auto | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 734 | qed | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 735 | |
| 38656 | 736 | subsection "Borel measurable operators" | 
| 737 | ||
| 738 | lemma (in sigma_algebra) affine_borel_measurable_vector: | |
| 739 | fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" | |
| 740 | assumes "f \<in> borel_measurable M" | |
| 741 | shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" | |
| 742 | proof (rule borel_measurableI) | |
| 743 | fix S :: "'x set" assume "open S" | |
| 744 | show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" | |
| 745 | proof cases | |
| 746 | assume "b \<noteq> 0" | |
| 747 | with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open") | |
| 748 | by (auto intro!: open_affinity simp: scaleR.add_right mem_def) | |
| 749 | hence "?S \<in> sets borel_space" | |
| 750 | unfolding borel_space_def by (auto simp: sigma_def intro!: sigma_sets.Basic) | |
| 751 | moreover | |
| 752 | from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" | |
| 753 | apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) | |
| 754 | ultimately show ?thesis using assms unfolding in_borel_measurable_borel_space | |
| 755 | by auto | |
| 756 | qed simp | |
| 757 | qed | |
| 758 | ||
| 759 | lemma (in sigma_algebra) affine_borel_measurable: | |
| 760 | fixes g :: "'a \<Rightarrow> real" | |
| 761 | assumes g: "g \<in> borel_measurable M" | |
| 762 | shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M" | |
| 763 | using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute) | |
| 764 | ||
| 765 | lemma (in sigma_algebra) borel_measurable_add[simp, intro]: | |
| 766 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 767 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 768 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 769 | shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 770 | proof - | 
| 38656 | 771 |   have 1: "\<And>a. {w\<in>space M. a \<le> f w + g w} = {w \<in> space M. a + g w * -1 \<le> f w}"
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 772 | by auto | 
| 38656 | 773 | have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M" | 
| 774 | by (rule affine_borel_measurable [OF g]) | |
| 775 |   then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
 | |
| 776 | by auto | |
| 777 |   then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
 | |
| 778 | by (simp add: 1) | |
| 779 | then show ?thesis | |
| 780 | by (simp add: borel_measurable_iff_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 781 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 782 | |
| 38656 | 783 | lemma (in sigma_algebra) borel_measurable_square: | 
| 784 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 785 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 786 | shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 787 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 788 |   {
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 789 | fix a | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 790 |     have "{w \<in> space M. (f w)\<twosuperior> \<le> a} \<in> sets M"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 791 | proof (cases rule: linorder_cases [of a 0]) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 792 | case less | 
| 38656 | 793 |       hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}"
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 794 | by auto (metis less order_le_less_trans power2_less_0) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 795 | also have "... \<in> sets M" | 
| 38656 | 796 | by (rule empty_sets) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 797 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 798 | next | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 799 | case equal | 
| 38656 | 800 |       hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 801 |              {w \<in> space M. f w \<le> 0} \<inter> {w \<in> space M. 0 \<le> f w}"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 802 | by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 803 | also have "... \<in> sets M" | 
| 38656 | 804 | apply (insert f) | 
| 805 | apply (rule Int) | |
| 806 | apply (simp add: borel_measurable_iff_le) | |
| 807 | apply (simp add: borel_measurable_iff_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 808 | done | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 809 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 810 | next | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 811 | case greater | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 812 | have "\<forall>x. (f x ^ 2 \<le> sqrt a ^ 2) = (- sqrt a \<le> f x & f x \<le> sqrt a)" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 813 | by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 814 | real_sqrt_le_iff real_sqrt_power) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 815 |       hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
 | 
| 38656 | 816 |              {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}"
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 817 | using greater by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 818 | also have "... \<in> sets M" | 
| 38656 | 819 | apply (insert f) | 
| 820 | apply (rule Int) | |
| 821 | apply (simp add: borel_measurable_iff_ge) | |
| 822 | apply (simp add: borel_measurable_iff_le) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 823 | done | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 824 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 825 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 826 | } | 
| 38656 | 827 | thus ?thesis by (auto simp add: borel_measurable_iff_le) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 828 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 829 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 830 | lemma times_eq_sum_squares: | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 831 | fixes x::real | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 832 | shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4" | 
| 38656 | 833 | by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric]) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 834 | |
| 38656 | 835 | lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]: | 
| 836 | fixes g :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 837 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 838 | shows "(\<lambda>x. - g x) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 839 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 840 | have "(\<lambda>x. - g x) = (\<lambda>x. 0 + (g x) * -1)" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 841 | by simp | 
| 38656 | 842 | also have "... \<in> borel_measurable M" | 
| 843 | by (fast intro: affine_borel_measurable g) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 844 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 845 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 846 | |
| 38656 | 847 | lemma (in sigma_algebra) borel_measurable_times[simp, intro]: | 
| 848 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 849 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 850 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 851 | shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 852 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 853 | have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M" | 
| 38656 | 854 | using assms by (fast intro: affine_borel_measurable borel_measurable_square) | 
| 855 | have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) = | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 856 | (\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))" | 
| 35582 | 857 | by (simp add: minus_divide_right) | 
| 38656 | 858 | also have "... \<in> borel_measurable M" | 
| 859 | using f g by (fast intro: affine_borel_measurable borel_measurable_square f g) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 860 | finally have 2: "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) \<in> borel_measurable M" . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 861 | show ?thesis | 
| 38656 | 862 | apply (simp add: times_eq_sum_squares diff_minus) | 
| 863 | using 1 2 by simp | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 864 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 865 | |
| 38656 | 866 | lemma (in sigma_algebra) borel_measurable_diff[simp, intro]: | 
| 867 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 868 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 869 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 870 | shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" | 
| 38656 | 871 | unfolding diff_minus using assms by fast | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 872 | |
| 38656 | 873 | lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]: | 
| 874 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real" | |
| 875 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | |
| 876 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" | |
| 877 | proof cases | |
| 878 | assume "finite S" | |
| 879 | thus ?thesis using assms by induct auto | |
| 880 | qed simp | |
| 35692 | 881 | |
| 38656 | 882 | lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]: | 
| 883 | fixes f :: "'a \<Rightarrow> real" | |
| 35692 | 884 | assumes "f \<in> borel_measurable M" | 
| 885 | shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M" | |
| 38656 | 886 | unfolding borel_measurable_iff_ge unfolding inverse_eq_divide | 
| 887 | proof safe | |
| 888 | fix a :: real | |
| 889 |   have *: "{w \<in> space M. a \<le> 1 / f w} =
 | |
| 890 |       ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
 | |
| 891 |       ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
 | |
| 892 |       ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
 | |
| 893 |   show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
 | |
| 894 | by (auto intro!: Int Un) | |
| 35692 | 895 | qed | 
| 896 | ||
| 38656 | 897 | lemma (in sigma_algebra) borel_measurable_divide[simp, intro]: | 
| 898 | fixes f :: "'a \<Rightarrow> real" | |
| 35692 | 899 | assumes "f \<in> borel_measurable M" | 
| 900 | and "g \<in> borel_measurable M" | |
| 901 | shows "(\<lambda>x. f x / g x) \<in> borel_measurable M" | |
| 902 | unfolding field_divide_inverse | |
| 38656 | 903 | by (rule borel_measurable_inverse borel_measurable_times assms)+ | 
| 904 | ||
| 905 | lemma (in sigma_algebra) borel_measurable_max[intro, simp]: | |
| 906 | fixes f g :: "'a \<Rightarrow> real" | |
| 907 | assumes "f \<in> borel_measurable M" | |
| 908 | assumes "g \<in> borel_measurable M" | |
| 909 | shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" | |
| 910 | unfolding borel_measurable_iff_le | |
| 911 | proof safe | |
| 912 | fix a | |
| 913 |   have "{x \<in> space M. max (g x) (f x) \<le> a} =
 | |
| 914 |     {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
 | |
| 915 |   thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
 | |
| 916 | using assms unfolding borel_measurable_iff_le | |
| 917 | by (auto intro!: Int) | |
| 918 | qed | |
| 919 | ||
| 920 | lemma (in sigma_algebra) borel_measurable_min[intro, simp]: | |
| 921 | fixes f g :: "'a \<Rightarrow> real" | |
| 922 | assumes "f \<in> borel_measurable M" | |
| 923 | assumes "g \<in> borel_measurable M" | |
| 924 | shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M" | |
| 925 | unfolding borel_measurable_iff_ge | |
| 926 | proof safe | |
| 927 | fix a | |
| 928 |   have "{x \<in> space M. a \<le> min (g x) (f x)} =
 | |
| 929 |     {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
 | |
| 930 |   thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
 | |
| 931 | using assms unfolding borel_measurable_iff_ge | |
| 932 | by (auto intro!: Int) | |
| 933 | qed | |
| 934 | ||
| 935 | lemma (in sigma_algebra) borel_measurable_abs[simp, intro]: | |
| 936 | assumes "f \<in> borel_measurable M" | |
| 937 | shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" | |
| 938 | proof - | |
| 939 | have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def) | |
| 940 | show ?thesis unfolding * using assms by auto | |
| 941 | qed | |
| 942 | ||
| 943 | section "Borel space over the real line with infinity" | |
| 35692 | 944 | |
| 38656 | 945 | lemma borel_space_Real_measurable: | 
| 946 | "A \<in> sets borel_space \<Longrightarrow> Real -` A \<in> sets borel_space" | |
| 947 | proof (rule borel_measurable_translate) | |
| 948 | fix B :: "pinfreal set" assume "open B" | |
| 949 |   then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
 | |
| 950 |     x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
 | |
| 951 | unfolding open_pinfreal_def by blast | |
| 952 | ||
| 953 |   have "Real -` B = Real -` (B - {\<omega>})" by auto
 | |
| 954 |   also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
 | |
| 955 |   also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
 | |
| 956 | apply (auto simp add: Real_eq_Real image_iff) | |
| 957 | apply (rule_tac x="max 0 x" in bexI) | |
| 958 | by (auto simp: max_def) | |
| 959 | finally show "Real -` B \<in> sets borel_space" | |
| 960 | using `open T` by auto | |
| 961 | qed simp | |
| 962 | ||
| 963 | lemma borel_space_real_measurable: | |
| 964 | "A \<in> sets borel_space \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel_space" | |
| 965 | proof (rule borel_measurable_translate) | |
| 966 | fix B :: "real set" assume "open B" | |
| 967 |   { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
 | |
| 968 | note Ex_less_real = this | |
| 969 |   have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
 | |
| 970 | by (force simp: Ex_less_real) | |
| 971 | ||
| 972 |   have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
 | |
| 973 | unfolding open_pinfreal_def using `open B` | |
| 974 |     by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
 | |
| 975 | then show "(real -` B :: pinfreal set) \<in> sets borel_space" unfolding * by auto | |
| 976 | qed simp | |
| 977 | ||
| 978 | lemma (in sigma_algebra) borel_measurable_Real[intro, simp]: | |
| 979 | assumes "f \<in> borel_measurable M" | |
| 980 | shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M" | |
| 981 | unfolding in_borel_measurable_borel_space | |
| 982 | proof safe | |
| 983 | fix S :: "pinfreal set" assume "S \<in> sets borel_space" | |
| 984 | from borel_space_Real_measurable[OF this] | |
| 985 | have "(Real \<circ> f) -` S \<inter> space M \<in> sets M" | |
| 986 | using assms | |
| 987 | unfolding vimage_compose in_borel_measurable_borel_space | |
| 988 | by auto | |
| 989 | thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def) | |
| 35748 | 990 | qed | 
| 35692 | 991 | |
| 38656 | 992 | lemma (in sigma_algebra) borel_measurable_real[intro, simp]: | 
| 993 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 994 | assumes "f \<in> borel_measurable M" | |
| 995 | shows "(\<lambda>x. real (f x)) \<in> borel_measurable M" | |
| 996 | unfolding in_borel_measurable_borel_space | |
| 997 | proof safe | |
| 998 | fix S :: "real set" assume "S \<in> sets borel_space" | |
| 999 | from borel_space_real_measurable[OF this] | |
| 1000 | have "(real \<circ> f) -` S \<inter> space M \<in> sets M" | |
| 1001 | using assms | |
| 1002 | unfolding vimage_compose in_borel_measurable_borel_space | |
| 1003 | by auto | |
| 1004 | thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def) | |
| 1005 | qed | |
| 35692 | 1006 | |
| 38656 | 1007 | lemma (in sigma_algebra) borel_measurable_Real_eq: | 
| 1008 | assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" | |
| 1009 | shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" | |
| 1010 | proof | |
| 1011 |   have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
 | |
| 1012 | by auto | |
| 1013 | assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M" | |
| 1014 | hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M" | |
| 1015 | by (rule borel_measurable_real) | |
| 1016 | moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x" | |
| 1017 | using assms by auto | |
| 1018 | ultimately show "f \<in> borel_measurable M" | |
| 1019 | by (simp cong: measurable_cong) | |
| 1020 | qed auto | |
| 35692 | 1021 | |
| 38656 | 1022 | lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real: | 
| 1023 | "f \<in> borel_measurable M \<longleftrightarrow> | |
| 1024 |     ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
 | |
| 1025 | proof safe | |
| 1026 | assume "f \<in> borel_measurable M" | |
| 1027 |   then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
 | |
| 1028 | by (auto intro: borel_measurable_vimage borel_measurable_real) | |
| 1029 | next | |
| 1030 |   assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
 | |
| 1031 |   have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
 | |
| 1032 |   with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
 | |
| 1033 | have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 1034 | by (simp add: fun_eq_iff Real_real) | 
| 38656 | 1035 | show "f \<in> borel_measurable M" | 
| 1036 | apply (subst f) | |
| 1037 | apply (rule measurable_If) | |
| 1038 | using * ** by auto | |
| 1039 | qed | |
| 1040 | ||
| 1041 | lemma (in sigma_algebra) less_eq_ge_measurable: | |
| 1042 | fixes f :: "'a \<Rightarrow> 'c::linorder" | |
| 1043 |   shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
 | |
| 1044 | proof | |
| 1045 |   assume "{x\<in>space M. f x \<le> a} \<in> sets M"
 | |
| 1046 |   moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
 | |
| 1047 |   ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
 | |
| 1048 | next | |
| 1049 |   assume "{x\<in>space M. a < f x} \<in> sets M"
 | |
| 1050 |   moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
 | |
| 1051 |   ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
 | |
| 1052 | qed | |
| 35692 | 1053 | |
| 38656 | 1054 | lemma (in sigma_algebra) greater_eq_le_measurable: | 
| 1055 | fixes f :: "'a \<Rightarrow> 'c::linorder" | |
| 1056 |   shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
 | |
| 1057 | proof | |
| 1058 |   assume "{x\<in>space M. a \<le> f x} \<in> sets M"
 | |
| 1059 |   moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
 | |
| 1060 |   ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
 | |
| 1061 | next | |
| 1062 |   assume "{x\<in>space M. f x < a} \<in> sets M"
 | |
| 1063 |   moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
 | |
| 1064 |   ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
 | |
| 1065 | qed | |
| 1066 | ||
| 1067 | lemma (in sigma_algebra) less_eq_le_pinfreal_measurable: | |
| 1068 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1069 |   shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
 | |
| 1070 | proof | |
| 1071 |   assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
 | |
| 1072 |   show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
 | |
| 1073 | proof | |
| 1074 |     fix a show "{x \<in> space M. a < f x} \<in> sets M"
 | |
| 1075 | proof (cases a) | |
| 1076 | case (preal r) | |
| 1077 |       have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
 | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1078 | proof safe | 
| 38656 | 1079 | fix x assume "a < f x" and [simp]: "x \<in> space M" | 
| 1080 | with ex_pinfreal_inverse_of_nat_Suc_less[of "f x - a"] | |
| 1081 | obtain n where "a + inverse (of_nat (Suc n)) < f x" | |
| 1082 | by (cases "f x", auto simp: pinfreal_minus_order) | |
| 1083 | then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp | |
| 1084 |         then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
 | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1085 | by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1086 | next | 
| 38656 | 1087 | fix i x assume [simp]: "x \<in> space M" | 
| 1088 | have "a < a + inverse (of_nat (Suc i))" using preal by auto | |
| 1089 | also assume "a + inverse (of_nat (Suc i)) \<le> f x" | |
| 1090 | finally show "a < f x" . | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1091 | qed | 
| 38656 | 1092 | with a show ?thesis by auto | 
| 1093 | qed simp | |
| 35582 | 1094 | qed | 
| 1095 | next | |
| 38656 | 1096 |   assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
 | 
| 1097 |   then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
 | |
| 1098 |   show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
 | |
| 1099 | proof | |
| 1100 |     fix a show "{x \<in> space M. f x < a} \<in> sets M"
 | |
| 1101 | proof (cases a) | |
| 1102 | case (preal r) | |
| 1103 | show ?thesis | |
| 1104 | proof cases | |
| 1105 | assume "a = 0" then show ?thesis by simp | |
| 1106 | next | |
| 1107 | assume "a \<noteq> 0" | |
| 1108 |         have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
 | |
| 1109 | proof safe | |
| 1110 | fix x assume "f x < a" and [simp]: "x \<in> space M" | |
| 1111 | with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"] | |
| 1112 | obtain n where "inverse (of_nat (Suc n)) < a - f x" | |
| 1113 | using preal by (cases "f x") auto | |
| 1114 | then have "f x \<le> a - inverse (of_nat (Suc n)) " | |
| 1115 | using preal by (cases "f x") (auto split: split_if_asm) | |
| 1116 |           then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
 | |
| 1117 | by auto | |
| 1118 | next | |
| 1119 | fix i x assume [simp]: "x \<in> space M" | |
| 1120 | assume "f x \<le> a - inverse (of_nat (Suc i))" | |
| 1121 | also have "\<dots> < a" using `a \<noteq> 0` preal by auto | |
| 1122 | finally show "f x < a" . | |
| 1123 | qed | |
| 1124 | with a show ?thesis by auto | |
| 1125 | qed | |
| 1126 | next | |
| 1127 | case infinite | |
| 1128 |       have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
 | |
| 1129 | proof (safe, simp_all, safe) | |
| 1130 | fix x assume *: "\<forall>n::nat. Real (real n) < f x" | |
| 1131 | show "f x = \<omega>" proof (rule ccontr) | |
| 1132 | assume "f x \<noteq> \<omega>" | |
| 1133 | with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n" | |
| 1134 | by (auto simp: pinfreal_noteq_omega_Ex) | |
| 1135 | with *[THEN spec, of n] show False by auto | |
| 1136 | qed | |
| 1137 | qed | |
| 1138 |       with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
 | |
| 1139 |       moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
 | |
| 1140 | using infinite by auto | |
| 1141 | ultimately show ?thesis by auto | |
| 1142 | qed | |
| 35582 | 1143 | qed | 
| 1144 | qed | |
| 1145 | ||
| 38656 | 1146 | lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater: | 
| 1147 |   "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
 | |
| 1148 | proof safe | |
| 1149 | fix a assume f: "f \<in> borel_measurable M" | |
| 1150 |   have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
 | |
| 1151 |   with f show "{x\<in>space M. a < f x} \<in> sets M"
 | |
| 1152 | by (auto intro!: measurable_sets) | |
| 1153 | next | |
| 1154 |   assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
 | |
| 1155 |   hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
 | |
| 1156 | unfolding less_eq_le_pinfreal_measurable | |
| 1157 | unfolding greater_eq_le_measurable . | |
| 35582 | 1158 | |
| 38656 | 1159 | show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater | 
| 1160 | proof safe | |
| 1161 |     have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
 | |
| 1162 |     then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
 | |
| 35582 | 1163 | |
| 38656 | 1164 | fix a | 
| 1165 |     have "{w \<in> space M. a < real (f w)} =
 | |
| 1166 |       (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
 | |
| 1167 | proof (split split_if, safe del: notI) | |
| 1168 | fix x assume "0 \<le> a" | |
| 1169 |       { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
 | |
| 1170 | using `0 \<le> a` by (cases "f x", auto) } | |
| 1171 |       { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
 | |
| 1172 | using `0 \<le> a` by (cases "f x", auto) } | |
| 1173 | next | |
| 1174 | fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto | |
| 1175 | qed | |
| 1176 |     then show "{w \<in> space M. a < real (f w)} \<in> sets M"
 | |
| 1177 | using \<omega> * by (auto intro!: Diff) | |
| 35582 | 1178 | qed | 
| 1179 | qed | |
| 1180 | ||
| 38656 | 1181 | lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less: | 
| 1182 |   "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
 | |
| 1183 | using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable . | |
| 1184 | ||
| 1185 | lemma (in sigma_algebra) borel_measurable_pinfreal_iff_le: | |
| 1186 |   "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
 | |
| 1187 | using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable . | |
| 1188 | ||
| 1189 | lemma (in sigma_algebra) borel_measurable_pinfreal_iff_ge: | |
| 1190 |   "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
 | |
| 1191 | using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable . | |
| 1192 | ||
| 1193 | lemma (in sigma_algebra) borel_measurable_pinfreal_eq_const: | |
| 1194 | fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" | |
| 1195 |   shows "{x\<in>space M. f x = c} \<in> sets M"
 | |
| 1196 | proof - | |
| 1197 |   have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
 | |
| 1198 | then show ?thesis using assms by (auto intro!: measurable_sets) | |
| 1199 | qed | |
| 1200 | ||
| 1201 | lemma (in sigma_algebra) borel_measurable_pinfreal_neq_const: | |
| 1202 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1203 | assumes "f \<in> borel_measurable M" | |
| 1204 |   shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 | |
| 1205 | proof - | |
| 1206 |   have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
 | |
| 1207 | then show ?thesis using assms by (auto intro!: measurable_sets) | |
| 1208 | qed | |
| 1209 | ||
| 1210 | lemma (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]: | |
| 1211 | fixes f g :: "'a \<Rightarrow> pinfreal" | |
| 1212 | assumes f: "f \<in> borel_measurable M" | |
| 1213 | assumes g: "g \<in> borel_measurable M" | |
| 1214 |   shows "{x \<in> space M. f x < g x} \<in> sets M"
 | |
| 1215 | proof - | |
| 1216 | have "(\<lambda>x. real (f x)) \<in> borel_measurable M" | |
| 1217 | "(\<lambda>x. real (g x)) \<in> borel_measurable M" | |
| 1218 | using assms by (auto intro!: borel_measurable_real) | |
| 1219 | from borel_measurable_less[OF this] | |
| 1220 |   have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
 | |
| 1221 |   moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
 | |
| 1222 |   moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
 | |
| 1223 |   moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
 | |
| 1224 |   moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
 | |
| 1225 |     ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
 | |
| 1226 | by (auto simp: real_of_pinfreal_strict_mono_iff) | |
| 1227 | ultimately show ?thesis by auto | |
| 1228 | qed | |
| 1229 | ||
| 1230 | lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]: | |
| 1231 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1232 | assumes f: "f \<in> borel_measurable M" | |
| 1233 | assumes g: "g \<in> borel_measurable M" | |
| 1234 |   shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
 | |
| 1235 | proof - | |
| 1236 |   have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
 | |
| 1237 | then show ?thesis using g f by auto | |
| 1238 | qed | |
| 1239 | ||
| 1240 | lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]: | |
| 1241 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1242 | assumes f: "f \<in> borel_measurable M" | |
| 1243 | assumes g: "g \<in> borel_measurable M" | |
| 1244 |   shows "{w \<in> space M. f w = g w} \<in> sets M"
 | |
| 1245 | proof - | |
| 1246 |   have "{x \<in> space M. f x = g x} = {x \<in> space M. g x \<le> f x} \<inter> {x \<in> space M. f x \<le> g x}" by auto
 | |
| 1247 | then show ?thesis using g f by auto | |
| 1248 | qed | |
| 1249 | ||
| 1250 | lemma (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]: | |
| 1251 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1252 | assumes f: "f \<in> borel_measurable M" | |
| 1253 | assumes g: "g \<in> borel_measurable M" | |
| 1254 |   shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | |
| 35692 | 1255 | proof - | 
| 38656 | 1256 |   have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
 | 
| 1257 | thus ?thesis using f g by auto | |
| 1258 | qed | |
| 1259 | ||
| 1260 | lemma (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]: | |
| 1261 | fixes f :: "'a \<Rightarrow> pinfreal" | |
| 1262 | assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | |
| 1263 | shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" | |
| 1264 | proof - | |
| 1265 | have *: "(\<lambda>x. f x + g x) = | |
| 1266 | (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 1267 | by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex) | 
| 38656 | 1268 | show ?thesis using assms unfolding * | 
| 1269 | by (auto intro!: measurable_If) | |
| 1270 | qed | |
| 1271 | ||
| 1272 | lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]: | |
| 1273 | fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" | |
| 1274 | shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" | |
| 1275 | proof - | |
| 1276 | have *: "(\<lambda>x. f x * g x) = | |
| 1277 | (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else | |
| 1278 | Real (real (f x) * real (g x)))" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 1279 | by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex) | 
| 38656 | 1280 | show ?thesis using assms unfolding * | 
| 1281 | by (auto intro!: measurable_If) | |
| 1282 | qed | |
| 1283 | ||
| 1284 | lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]: | |
| 1285 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal" | |
| 1286 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | |
| 1287 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" | |
| 1288 | proof cases | |
| 1289 | assume "finite S" | |
| 1290 | thus ?thesis using assms | |
| 1291 | by induct auto | |
| 1292 | qed (simp add: borel_measurable_const) | |
| 1293 | ||
| 1294 | lemma (in sigma_algebra) borel_measurable_pinfreal_min[intro, simp]: | |
| 1295 | fixes f g :: "'a \<Rightarrow> pinfreal" | |
| 1296 | assumes "f \<in> borel_measurable M" | |
| 1297 | assumes "g \<in> borel_measurable M" | |
| 1298 | shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M" | |
| 1299 | using assms unfolding min_def by (auto intro!: measurable_If) | |
| 1300 | ||
| 1301 | lemma (in sigma_algebra) borel_measurable_pinfreal_max[intro]: | |
| 1302 | fixes f g :: "'a \<Rightarrow> pinfreal" | |
| 1303 | assumes "f \<in> borel_measurable M" | |
| 1304 | and "g \<in> borel_measurable M" | |
| 1305 | shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" | |
| 1306 | using assms unfolding max_def by (auto intro!: measurable_If) | |
| 1307 | ||
| 1308 | lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]: | |
| 1309 | fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal" | |
| 1310 | assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M" | |
| 1311 | shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M") | |
| 1312 | unfolding borel_measurable_pinfreal_iff_greater | |
| 1313 | proof safe | |
| 1314 | fix a | |
| 1315 |   have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
 | |
| 38705 | 1316 | by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'c=pinfreal]) | 
| 38656 | 1317 |   then show "{x\<in>space M. a < ?sup x} \<in> sets M"
 | 
| 1318 | using assms by auto | |
| 1319 | qed | |
| 1320 | ||
| 1321 | lemma (in sigma_algebra) borel_measurable_INF[simp, intro]: | |
| 1322 | fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal" | |
| 1323 | assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M" | |
| 1324 | shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M") | |
| 1325 | unfolding borel_measurable_pinfreal_iff_less | |
| 1326 | proof safe | |
| 1327 | fix a | |
| 1328 |   have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
 | |
| 1329 | by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand) | |
| 1330 |   then show "{x\<in>space M. ?inf x < a} \<in> sets M"
 | |
| 1331 | using assms by auto | |
| 1332 | qed | |
| 1333 | ||
| 1334 | lemma (in sigma_algebra) borel_measurable_pinfreal_diff: | |
| 1335 | fixes f g :: "'a \<Rightarrow> pinfreal" | |
| 1336 | assumes "f \<in> borel_measurable M" | |
| 1337 | assumes "g \<in> borel_measurable M" | |
| 1338 | shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" | |
| 1339 | unfolding borel_measurable_pinfreal_iff_greater | |
| 1340 | proof safe | |
| 1341 | fix a | |
| 1342 |   have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
 | |
| 1343 | by (simp add: pinfreal_less_minus_iff) | |
| 1344 |   then show "{x \<in> space M. a < f x - g x} \<in> sets M"
 | |
| 1345 | using assms by auto | |
| 35692 | 1346 | qed | 
| 1347 | ||
| 39092 | 1348 | lemma (in sigma_algebra) borel_measurable_psuminf: | 
| 1349 | assumes "\<And>i. f i \<in> borel_measurable M" | |
| 1350 | shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M" | |
| 1351 | using assms unfolding psuminf_def | |
| 1352 | by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) | |
| 1353 | ||
| 1354 | section "LIMSEQ is borel measurable" | |
| 1355 | ||
| 1356 | lemma (in sigma_algebra) borel_measurable_LIMSEQ: | |
| 1357 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" | |
| 1358 | assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" | |
| 1359 | and u: "\<And>i. u i \<in> borel_measurable M" | |
| 1360 | shows "u' \<in> borel_measurable M" | |
| 1361 | proof - | |
| 1362 | let "?pu x i" = "max (u i x) 0" | |
| 1363 | let "?nu x i" = "max (- u i x) 0" | |
| 1364 | ||
| 1365 |   { fix x assume x: "x \<in> space M"
 | |
| 1366 | have "(?pu x) ----> max (u' x) 0" | |
| 1367 | "(?nu x) ----> max (- u' x) 0" | |
| 1368 | using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) | |
| 1369 | from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] | |
| 1370 | have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" | |
| 1371 | "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" | |
| 1372 | by (simp_all add: Real_max'[symmetric]) } | |
| 1373 | note eq = this | |
| 1374 | ||
| 1375 | have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x" | |
| 1376 | by auto | |
| 1377 | ||
| 1378 | have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M" | |
| 1379 | "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M" | |
| 1380 | using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real) | |
| 1381 | with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space] | |
| 1382 | have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M" | |
| 1383 | "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M" | |
| 1384 | unfolding SUPR_fun_expand INFI_fun_expand by auto | |
| 1385 | note this[THEN borel_measurable_real] | |
| 1386 | from borel_measurable_diff[OF this] | |
| 1387 | show ?thesis unfolding * . | |
| 1388 | qed | |
| 1389 | ||
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1390 | end |