| author | kuncar | 
| Mon, 26 Mar 2012 15:32:54 +0200 | |
| changeset 47116 | 529d2a949bd4 | 
| parent 46905 | 6b1c0a80a57a | 
| child 47694 | 05663f75964c | 
| permissions | -rw-r--r-- | 
| 42150 | 1 | (* Title: HOL/Probability/Borel_Space.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 38656 | 5 | |
| 6 | header {*Borel spaces*}
 | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 7 | |
| 40859 | 8 | theory Borel_Space | 
| 45288 | 9 | imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 10 | begin | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 11 | |
| 38656 | 12 | section "Generic Borel spaces" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 13 | |
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 14 | definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = {S. open S}\<rparr>"
 | 
| 40859 | 15 | abbreviation "borel_measurable M \<equiv> measurable M borel" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 16 | |
| 40859 | 17 | interpretation borel: sigma_algebra borel | 
| 18 | by (auto simp: borel_def intro!: sigma_algebra_sigma) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 19 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 20 | lemma in_borel_measurable: | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 21 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 22 |     (\<forall>S \<in> sets (sigma \<lparr> space = UNIV, sets = {S. open S}\<rparr>).
 | 
| 38656 | 23 | f -` S \<inter> space M \<in> sets M)" | 
| 40859 | 24 | by (auto simp add: measurable_def borel_def) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 25 | |
| 40859 | 26 | lemma in_borel_measurable_borel: | 
| 38656 | 27 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 40859 | 28 | (\<forall>S \<in> sets borel. | 
| 38656 | 29 | f -` S \<inter> space M \<in> sets M)" | 
| 40859 | 30 | by (auto simp add: measurable_def borel_def) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 31 | |
| 40859 | 32 | lemma space_borel[simp]: "space borel = UNIV" | 
| 33 | unfolding borel_def by auto | |
| 38656 | 34 | |
| 40859 | 35 | lemma borel_open[simp]: | 
| 36 | assumes "open A" shows "A \<in> sets borel" | |
| 38656 | 37 | proof - | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 38 |   have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
 | 
| 40859 | 39 | thus ?thesis unfolding borel_def sigma_def by (auto intro!: sigma_sets.Basic) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 40 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 41 | |
| 40859 | 42 | lemma borel_closed[simp]: | 
| 43 | assumes "closed A" shows "A \<in> sets borel" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 44 | proof - | 
| 40859 | 45 | have "space borel - (- A) \<in> sets borel" | 
| 46 | using assms unfolding closed_def by (blast intro: borel_open) | |
| 38656 | 47 | thus ?thesis by simp | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 48 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 49 | |
| 41830 | 50 | lemma borel_comp[intro,simp]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel" | 
| 51 | unfolding Compl_eq_Diff_UNIV by (intro borel.Diff) auto | |
| 52 | ||
| 38656 | 53 | lemma (in sigma_algebra) borel_measurable_vimage: | 
| 54 | fixes f :: "'a \<Rightarrow> 'x::t2_space" | |
| 55 | assumes borel: "f \<in> borel_measurable M" | |
| 56 |   shows "f -` {x} \<inter> space M \<in> sets M"
 | |
| 57 | proof (cases "x \<in> f ` space M") | |
| 58 | case True then obtain y where "x = f y" by auto | |
| 41969 | 59 | from closed_singleton[of "f y"] | 
| 40859 | 60 |   have "{f y} \<in> sets borel" by (rule borel_closed)
 | 
| 38656 | 61 | with assms show ?thesis | 
| 40859 | 62 | unfolding in_borel_measurable_borel `x = f y` by auto | 
| 38656 | 63 | next | 
| 64 |   case False hence "f -` {x} \<inter> space M = {}" by auto
 | |
| 65 | thus ?thesis by auto | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 66 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 67 | |
| 38656 | 68 | lemma (in sigma_algebra) borel_measurableI: | 
| 69 | fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space" | |
| 70 | assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" | |
| 71 | shows "f \<in> borel_measurable M" | |
| 40859 | 72 | unfolding borel_def | 
| 73 | proof (rule measurable_sigma, simp_all) | |
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 74 | fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M" | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 75 | using assms[of S] by simp | 
| 40859 | 76 | qed | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 77 | |
| 40859 | 78 | lemma borel_singleton[simp, intro]: | 
| 38656 | 79 | fixes x :: "'a::t1_space" | 
| 40859 | 80 | shows "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets borel" | 
| 81 | proof (rule borel.insert_in_sets) | |
| 82 |     show "{x} \<in> sets borel"
 | |
| 41969 | 83 | using closed_singleton[of x] by (rule borel_closed) | 
| 38656 | 84 | qed simp | 
| 85 | ||
| 86 | lemma (in sigma_algebra) borel_measurable_const[simp, intro]: | |
| 87 | "(\<lambda>x. c) \<in> borel_measurable M" | |
| 88 | by (auto intro!: measurable_const) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 89 | |
| 39083 | 90 | lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]: | 
| 38656 | 91 | assumes A: "A \<in> sets M" | 
| 92 | shows "indicator A \<in> borel_measurable M" | |
| 46905 | 93 | unfolding indicator_def [abs_def] using A | 
| 38656 | 94 | by (auto intro!: measurable_If_set borel_measurable_const) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 95 | |
| 40859 | 96 | lemma (in sigma_algebra) borel_measurable_indicator_iff: | 
| 97 |   "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
 | |
| 98 | (is "?I \<in> borel_measurable M \<longleftrightarrow> _") | |
| 99 | proof | |
| 100 | assume "?I \<in> borel_measurable M" | |
| 101 |   then have "?I -` {1} \<inter> space M \<in> sets M"
 | |
| 102 | unfolding measurable_def by auto | |
| 103 |   also have "?I -` {1} \<inter> space M = A \<inter> space M"
 | |
| 46905 | 104 | unfolding indicator_def [abs_def] by auto | 
| 40859 | 105 | finally show "A \<inter> space M \<in> sets M" . | 
| 106 | next | |
| 107 | assume "A \<inter> space M \<in> sets M" | |
| 108 | moreover have "?I \<in> borel_measurable M \<longleftrightarrow> | |
| 109 | (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M" | |
| 110 | by (intro measurable_cong) (auto simp: indicator_def) | |
| 111 | ultimately show "?I \<in> borel_measurable M" by auto | |
| 112 | qed | |
| 113 | ||
| 39092 | 114 | lemma (in sigma_algebra) borel_measurable_restricted: | 
| 43920 | 115 | fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M" | 
| 39092 | 116 | shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow> | 
| 117 | (\<lambda>x. f x * indicator A x) \<in> borel_measurable M" | |
| 118 | (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M") | |
| 119 | proof - | |
| 120 | interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) | |
| 121 | have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R" | |
| 122 | by (auto intro!: measurable_cong) | |
| 123 | show ?thesis unfolding * | |
| 40859 | 124 | unfolding in_borel_measurable_borel | 
| 39092 | 125 | proof (simp, safe) | 
| 43920 | 126 | fix S :: "ereal set" assume "S \<in> sets borel" | 
| 40859 | 127 | "\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M" | 
| 39092 | 128 | then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto | 
| 129 | then have f: "?f -` S \<inter> A \<in> sets M" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44666diff
changeset | 130 | using `A \<in> sets M` sets_into_space by fastforce | 
| 39092 | 131 | show "?f -` S \<inter> space M \<in> sets M" | 
| 132 | proof cases | |
| 133 | assume "0 \<in> S" | |
| 134 | then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)" | |
| 135 | using `A \<in> sets M` sets_into_space by auto | |
| 136 | then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff) | |
| 137 | next | |
| 138 | assume "0 \<notin> S" | |
| 139 | then have "?f -` S \<inter> space M = ?f -` S \<inter> A" | |
| 140 | using `A \<in> sets M` sets_into_space | |
| 141 | by (auto simp: indicator_def split: split_if_asm) | |
| 142 | then show ?thesis using f by auto | |
| 143 | qed | |
| 144 | next | |
| 43920 | 145 | fix S :: "ereal set" assume "S \<in> sets borel" | 
| 40859 | 146 | "\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M" | 
| 39092 | 147 | then have f: "?f -` S \<inter> space M \<in> sets M" by auto | 
| 148 | then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M" | |
| 149 | using `A \<in> sets M` sets_into_space | |
| 150 | apply (simp add: image_iff) | |
| 151 | apply (rule bexI[OF _ f]) | |
| 152 | by auto | |
| 153 | qed | |
| 154 | qed | |
| 155 | ||
| 156 | lemma (in sigma_algebra) borel_measurable_subalgebra: | |
| 41545 | 157 | assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N" | 
| 39092 | 158 | shows "f \<in> borel_measurable M" | 
| 159 | using assms unfolding measurable_def by auto | |
| 160 | ||
| 38656 | 161 | section "Borel spaces on euclidean spaces" | 
| 162 | ||
| 163 | lemma lessThan_borel[simp, intro]: | |
| 164 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 165 |   shows "{..< a} \<in> sets borel"
 | 
| 166 | by (blast intro: borel_open) | |
| 38656 | 167 | |
| 168 | lemma greaterThan_borel[simp, intro]: | |
| 169 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 170 |   shows "{a <..} \<in> sets borel"
 | 
| 171 | by (blast intro: borel_open) | |
| 38656 | 172 | |
| 173 | lemma greaterThanLessThan_borel[simp, intro]: | |
| 174 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 175 |   shows "{a<..<b} \<in> sets borel"
 | 
| 176 | by (blast intro: borel_open) | |
| 38656 | 177 | |
| 178 | lemma atMost_borel[simp, intro]: | |
| 179 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 180 |   shows "{..a} \<in> sets borel"
 | 
| 181 | by (blast intro: borel_closed) | |
| 38656 | 182 | |
| 183 | lemma atLeast_borel[simp, intro]: | |
| 184 | fixes a :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 185 |   shows "{a..} \<in> sets borel"
 | 
| 186 | by (blast intro: borel_closed) | |
| 38656 | 187 | |
| 188 | lemma atLeastAtMost_borel[simp, intro]: | |
| 189 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 190 |   shows "{a..b} \<in> sets borel"
 | 
| 191 | by (blast intro: borel_closed) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 192 | |
| 38656 | 193 | lemma greaterThanAtMost_borel[simp, intro]: | 
| 194 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 195 |   shows "{a<..b} \<in> sets borel"
 | 
| 38656 | 196 | unfolding greaterThanAtMost_def by blast | 
| 197 | ||
| 198 | lemma atLeastLessThan_borel[simp, intro]: | |
| 199 | fixes a b :: "'a\<Colon>ordered_euclidean_space" | |
| 40859 | 200 |   shows "{a..<b} \<in> sets borel"
 | 
| 38656 | 201 | unfolding atLeastLessThan_def by blast | 
| 202 | ||
| 203 | lemma hafspace_less_borel[simp, intro]: | |
| 204 | fixes a :: real | |
| 40859 | 205 |   shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel"
 | 
| 206 | by (auto intro!: borel_open open_halfspace_component_gt) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 207 | |
| 38656 | 208 | lemma hafspace_greater_borel[simp, intro]: | 
| 209 | fixes a :: real | |
| 40859 | 210 |   shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel"
 | 
| 211 | by (auto intro!: borel_open open_halfspace_component_lt) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 212 | |
| 38656 | 213 | lemma hafspace_less_eq_borel[simp, intro]: | 
| 214 | fixes a :: real | |
| 40859 | 215 |   shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel"
 | 
| 216 | by (auto intro!: borel_closed closed_halfspace_component_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 217 | |
| 38656 | 218 | lemma hafspace_greater_eq_borel[simp, intro]: | 
| 219 | fixes a :: real | |
| 40859 | 220 |   shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel"
 | 
| 221 | by (auto intro!: borel_closed closed_halfspace_component_le) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 222 | |
| 38656 | 223 | lemma (in sigma_algebra) borel_measurable_less[simp, intro]: | 
| 224 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 225 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 226 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 227 |   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 | 228 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 229 |   have "{w \<in> space M. f w < g w} =
 | 
| 38656 | 230 |         (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
 | 
| 231 | using Rats_dense_in_real by (auto simp add: Rats_def) | |
| 232 | then show ?thesis using f g | |
| 233 | by simp (blast intro: measurable_sets) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 234 | qed | 
| 38656 | 235 | |
| 236 | lemma (in sigma_algebra) borel_measurable_le[simp, intro]: | |
| 237 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 238 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 239 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 240 |   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 | 241 | proof - | 
| 38656 | 242 |   have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
 | 
| 243 | by auto | |
| 244 | thus ?thesis using f g | |
| 245 | by simp blast | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 246 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 247 | |
| 38656 | 248 | lemma (in sigma_algebra) borel_measurable_eq[simp, intro]: | 
| 249 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 250 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 251 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 252 |   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 | 253 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 254 |   have "{w \<in> space M. f w = g w} =
 | 
| 33536 | 255 |         {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 | 256 | by auto | 
| 38656 | 257 | thus ?thesis using f g by auto | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 258 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 259 | |
| 38656 | 260 | lemma (in sigma_algebra) borel_measurable_neq[simp, intro]: | 
| 261 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 262 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 263 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 264 |   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 | 265 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 266 |   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 | 267 | by auto | 
| 38656 | 268 | thus ?thesis using f g by auto | 
| 269 | qed | |
| 270 | ||
| 271 | subsection "Borel space equals sigma algebras over intervals" | |
| 272 | ||
| 273 | lemma rational_boxes: | |
| 274 | fixes x :: "'a\<Colon>ordered_euclidean_space" | |
| 275 | assumes "0 < e" | |
| 276 |   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"
 | |
| 277 | proof - | |
| 278 |   def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
 | |
| 279 | then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos) | |
| 280 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i") | |
| 281 | proof | |
| 282 | fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e | |
| 283 | show "?th i" by auto | |
| 284 | qed | |
| 285 | from choice[OF this] guess a .. note a = this | |
| 286 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i") | |
| 287 | proof | |
| 288 | fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e | |
| 289 | show "?th i" by auto | |
| 290 | qed | |
| 291 | from choice[OF this] guess b .. note b = this | |
| 292 |   { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
 | |
| 293 |     have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
 | |
| 294 | unfolding setL2_def[symmetric] by (rule euclidean_dist_l2) | |
| 295 |     also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
 | |
| 296 | proof (rule real_sqrt_less_mono, rule setsum_strict_mono) | |
| 297 |       fix i assume i: "i \<in> {..<DIM('a)}"
 | |
| 298 | have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto | |
| 299 | moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto | |
| 300 | moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto | |
| 301 | ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto | |
| 302 |       then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
 | |
| 303 | unfolding e'_def by (auto simp: dist_real_def) | |
| 304 |       then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
 | |
| 305 | by (rule power_strict_mono) auto | |
| 306 |       then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
 | |
| 307 | by (simp add: power_divide) | |
| 308 | qed auto | |
| 309 | also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive) | |
| 310 | finally have "dist x y < e" . } | |
| 311 | with a b show ?thesis | |
| 312 | apply (rule_tac exI[of _ "Chi a"]) | |
| 313 | apply (rule_tac exI[of _ "Chi b"]) | |
| 314 | using eucl_less[where 'a='a] by auto | |
| 315 | qed | |
| 316 | ||
| 317 | lemma ex_rat_list: | |
| 318 | fixes x :: "'a\<Colon>ordered_euclidean_space" | |
| 319 | assumes "\<And> i. x $$ i \<in> \<rat>" | |
| 320 |   shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
 | |
| 321 | proof - | |
| 322 | have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast | |
| 323 | from choice[OF this] guess r .. | |
| 324 |   then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
 | |
| 325 | qed | |
| 326 | ||
| 327 | lemma open_UNION: | |
| 328 | fixes M :: "'a\<Colon>ordered_euclidean_space set" | |
| 329 | assumes "open M" | |
| 330 |   shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
 | |
| 331 |                    (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
 | |
| 332 | (is "M = UNION ?idx ?box") | |
| 333 | proof safe | |
| 334 | fix x assume "x \<in> M" | |
| 335 | obtain e where e: "e > 0" "ball x e \<subseteq> M" | |
| 336 | using openE[OF assms `x \<in> M`] by auto | |
| 337 |   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"
 | |
| 338 | using rational_boxes[OF e(1)] by blast | |
| 339 |   then obtain p q where pq: "length p = DIM ('a)"
 | |
| 340 |                             "length q = DIM ('a)"
 | |
| 341 |                             "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
 | |
| 342 | using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast | |
| 343 | hence p: "Chi (of_rat \<circ> op ! p) = a" | |
| 344 | using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a] | |
| 345 | unfolding o_def by auto | |
| 346 | from pq have q: "Chi (of_rat \<circ> op ! q) = b" | |
| 347 | using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b] | |
| 348 | unfolding o_def by auto | |
| 349 | have "x \<in> ?box (p, q)" | |
| 350 | using p q ab by auto | |
| 351 | thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto | |
| 352 | qed auto | |
| 353 | ||
| 354 | lemma halfspace_span_open: | |
| 40859 | 355 |   "sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))
 | 
| 356 | \<subseteq> sets borel" | |
| 357 | by (auto intro!: borel.sigma_sets_subset[simplified] borel_open | |
| 358 | open_halfspace_component_lt) | |
| 38656 | 359 | |
| 360 | lemma halfspace_lt_in_halfspace: | |
| 40859 | 361 |   "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>)"
 | 
| 362 | by (auto intro!: sigma_sets.Basic simp: sets_sigma) | |
| 38656 | 363 | |
| 364 | lemma halfspace_gt_in_halfspace: | |
| 40859 | 365 |   "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>)"
 | 
| 366 | (is "?set \<in> sets ?SIGMA") | |
| 38656 | 367 | proof - | 
| 40859 | 368 | interpret sigma_algebra "?SIGMA" | 
| 369 | by (intro sigma_algebra_sigma_sets) (simp_all add: sets_sigma) | |
| 38656 | 370 |   have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
 | 
| 371 | proof (safe, simp_all add: not_less) | |
| 372 | fix x assume "a < x $$ i" | |
| 373 | with reals_Archimedean[of "x $$ i - a"] | |
| 374 | obtain n where "a + 1 / real (Suc n) < x $$ i" | |
| 375 | by (auto simp: inverse_eq_divide field_simps) | |
| 376 | then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i" | |
| 377 | by (blast intro: less_imp_le) | |
| 378 | next | |
| 379 | fix x n | |
| 380 | have "a < a + 1 / real (Suc n)" by auto | |
| 381 | also assume "\<dots> \<le> x" | |
| 382 | finally show "a < x" . | |
| 383 | qed | |
| 384 | show "?set \<in> sets ?SIGMA" unfolding * | |
| 385 | 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 | 386 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 387 | |
| 38656 | 388 | lemma open_span_halfspace: | 
| 40859 | 389 |   "sets borel \<subseteq> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})\<rparr>)"
 | 
| 38656 | 390 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 391 | proof - | 
| 392 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp | |
| 38656 | 393 | then interpret sigma_algebra ?SIGMA . | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 394 |   { fix S :: "'a set" assume "S \<in> {S. open S}"
 | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 395 | then have "open S" unfolding mem_Collect_eq . | 
| 40859 | 396 | from open_UNION[OF this] | 
| 397 | obtain I where *: "S = | |
| 398 | (\<Union>(a, b)\<in>I. | |
| 399 |           (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
 | |
| 400 |           (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
 | |
| 401 | unfolding greaterThanLessThan_def | |
| 402 | unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] | |
| 403 | unfolding eucl_lessThan_eq_halfspaces[where 'a='a] | |
| 404 | by blast | |
| 405 | have "S \<in> sets ?SIGMA" | |
| 406 | unfolding * | |
| 407 | by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace) } | |
| 408 | then show ?thesis unfolding borel_def | |
| 409 | by (intro sets_sigma_subset) auto | |
| 410 | qed | |
| 38656 | 411 | |
| 412 | lemma halfspace_span_halfspace_le: | |
| 40859 | 413 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
 | 
| 414 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. x $$ i \<le> a})\<rparr>)"
 | |
| 38656 | 415 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 416 | proof - | 
| 417 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 38656 | 418 | then interpret sigma_algebra ?SIGMA . | 
| 40859 | 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 |     have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | |
| 435 | by (safe intro!: countable_UN) | |
| 436 | (auto simp: sets_sigma intro!: sigma_sets.Basic) } | |
| 437 | then show ?thesis by (intro sets_sigma_subset) auto | |
| 438 | qed | |
| 38656 | 439 | |
| 440 | lemma halfspace_span_halfspace_ge: | |
| 40859 | 441 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
 | 
| 442 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a \<le> x $$ i})\<rparr>)"
 | |
| 38656 | 443 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 444 | proof - | 
| 445 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 38656 | 446 | then interpret sigma_algebra ?SIGMA . | 
| 40859 | 447 |   { fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
 | 
| 448 |     have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | |
| 449 | by (safe intro!: Diff) | |
| 450 | (auto simp: sets_sigma intro!: sigma_sets.Basic) } | |
| 451 | then show ?thesis by (intro sets_sigma_subset) auto | |
| 452 | qed | |
| 38656 | 453 | |
| 454 | lemma halfspace_le_span_halfspace_gt: | |
| 40859 | 455 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
| 456 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a < x $$ i})\<rparr>)"
 | |
| 38656 | 457 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 458 | proof - | 
| 459 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 38656 | 460 | then interpret sigma_algebra ?SIGMA . | 
| 40859 | 461 |   { fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | 
| 462 |     have "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | |
| 463 | by (safe intro!: Diff) | |
| 464 | (auto simp: sets_sigma intro!: sigma_sets.Basic) } | |
| 465 | then show ?thesis by (intro sets_sigma_subset) auto | |
| 466 | qed | |
| 38656 | 467 | |
| 468 | lemma halfspace_le_span_atMost: | |
| 40859 | 469 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
| 470 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>)"
 | |
| 38656 | 471 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 472 | proof - | 
| 473 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 38656 | 474 | then interpret sigma_algebra ?SIGMA . | 
| 40859 | 475 |   have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 38656 | 476 | proof cases | 
| 40859 | 477 |     fix a i assume "i < DIM('a)"
 | 
| 38656 | 478 |     then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
 | 
| 479 | proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) | |
| 480 | fix x | |
| 481 |       from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
 | |
| 482 |       then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
 | |
| 483 | by (subst (asm) Max_le_iff) auto | |
| 484 |       then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
 | |
| 485 | by (auto intro!: exI[of _ k]) | |
| 486 | qed | |
| 487 |     show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | |
| 488 | by (safe intro!: countable_UN) | |
| 489 | (auto simp: sets_sigma intro!: sigma_sets.Basic) | |
| 490 | next | |
| 40859 | 491 |     fix a i assume "\<not> i < DIM('a)"
 | 
| 38656 | 492 |     then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 493 | using top by auto | |
| 494 | qed | |
| 40859 | 495 | then show ?thesis by (intro sets_sigma_subset) auto | 
| 496 | qed | |
| 38656 | 497 | |
| 498 | lemma halfspace_le_span_greaterThan: | |
| 40859 | 499 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
| 500 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {a<..})\<rparr>)"
 | |
| 38656 | 501 | (is "_ \<subseteq> sets ?SIGMA") | 
| 40859 | 502 | proof - | 
| 503 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 38656 | 504 | then interpret sigma_algebra ?SIGMA . | 
| 40859 | 505 |   have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 38656 | 506 | proof cases | 
| 40859 | 507 |     fix a i assume "i < DIM('a)"
 | 
| 38656 | 508 |     have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | 
| 509 |     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)`
 | |
| 510 | proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) | |
| 511 | fix x | |
| 44666 | 512 |       from reals_Archimedean2[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
 | 
| 38656 | 513 | guess k::nat .. note k = this | 
| 514 |       { fix i assume "i < DIM('a)"
 | |
| 515 | then have "-x$$i < real k" | |
| 516 | using k by (subst (asm) Max_less_iff) auto | |
| 517 | then have "- real k < x$$i" by simp } | |
| 518 |       then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
 | |
| 519 | by (auto intro!: exI[of _ k]) | |
| 520 | qed | |
| 521 |     finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | |
| 522 | apply (simp only:) | |
| 523 | apply (safe intro!: countable_UN Diff) | |
| 46731 | 524 | apply (auto simp: sets_sigma intro!: sigma_sets.Basic) | 
| 525 | done | |
| 38656 | 526 | next | 
| 40859 | 527 |     fix a i assume "\<not> i < DIM('a)"
 | 
| 38656 | 528 |     then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 529 | using top by auto | |
| 530 | qed | |
| 40859 | 531 | then show ?thesis by (intro sets_sigma_subset) auto | 
| 532 | qed | |
| 533 | ||
| 534 | lemma halfspace_le_span_lessThan: | |
| 535 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i})\<rparr>) \<subseteq>
 | |
| 536 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..<a})\<rparr>)"
 | |
| 537 | (is "_ \<subseteq> sets ?SIGMA") | |
| 538 | proof - | |
| 539 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 540 | then interpret sigma_algebra ?SIGMA . | |
| 541 |   have "\<And>a i. {x. a \<le> x$$i} \<in> sets ?SIGMA"
 | |
| 542 | proof cases | |
| 543 |     fix a i assume "i < DIM('a)"
 | |
| 544 |     have "{x::'a. a \<le> x$$i} = space ?SIGMA - {x::'a. x$$i < a}" by auto
 | |
| 545 |     also have *: "{x::'a. x$$i < a} = (\<Union>k::nat. {..< (\<chi>\<chi> n. if n = i then a else real k)})" using `i <DIM('a)`
 | |
| 546 | proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) | |
| 547 | fix x | |
| 44666 | 548 |       from reals_Archimedean2[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
 | 
| 40859 | 549 | guess k::nat .. note k = this | 
| 550 |       { fix i assume "i < DIM('a)"
 | |
| 551 | then have "x$$i < real k" | |
| 552 | using k by (subst (asm) Max_less_iff) auto | |
| 553 | then have "x$$i < real k" by simp } | |
| 554 |       then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
 | |
| 555 | by (auto intro!: exI[of _ k]) | |
| 556 | qed | |
| 557 |     finally show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
 | |
| 558 | apply (simp only:) | |
| 559 | apply (safe intro!: countable_UN Diff) | |
| 46731 | 560 | apply (auto simp: sets_sigma intro!: sigma_sets.Basic) | 
| 561 | done | |
| 40859 | 562 | next | 
| 563 |     fix a i assume "\<not> i < DIM('a)"
 | |
| 564 |     then show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
 | |
| 565 | using top by auto | |
| 566 | qed | |
| 567 | then show ?thesis by (intro sets_sigma_subset) auto | |
| 568 | qed | |
| 569 | ||
| 570 | lemma atMost_span_atLeastAtMost: | |
| 571 |   "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>) \<subseteq>
 | |
| 572 |    sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a,b). {a..b})\<rparr>)"
 | |
| 573 | (is "_ \<subseteq> sets ?SIGMA") | |
| 574 | proof - | |
| 575 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 576 | then interpret sigma_algebra ?SIGMA . | |
| 577 |   { fix a::'a
 | |
| 578 |     have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
 | |
| 579 | proof (safe, simp_all add: eucl_le[where 'a='a]) | |
| 580 | fix x | |
| 581 |       from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
 | |
| 582 | guess k::nat .. note k = this | |
| 583 |       { fix i assume "i < DIM('a)"
 | |
| 584 | with k have "- x$$i \<le> real k" | |
| 585 | by (subst (asm) Max_le_iff) (auto simp: field_simps) | |
| 586 | then have "- real k \<le> x$$i" by simp } | |
| 587 |       then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
 | |
| 588 | by (auto intro!: exI[of _ k]) | |
| 589 | qed | |
| 590 |     have "{..a} \<in> sets ?SIGMA" unfolding *
 | |
| 591 | by (safe intro!: countable_UN) | |
| 592 | (auto simp: sets_sigma intro!: sigma_sets.Basic) } | |
| 593 | then show ?thesis by (intro sets_sigma_subset) auto | |
| 594 | qed | |
| 595 | ||
| 596 | lemma borel_eq_atMost: | |
| 597 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})\<rparr>)"
 | |
| 598 | (is "_ = ?SIGMA") | |
| 40869 | 599 | proof (intro algebra.equality antisym) | 
| 40859 | 600 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 601 | using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace | |
| 602 | by auto | |
| 603 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 604 | by (rule borel.sets_sigma_subset) auto | |
| 605 | qed auto | |
| 606 | ||
| 607 | lemma borel_eq_atLeastAtMost: | |
| 608 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})\<rparr>)"
 | |
| 609 | (is "_ = ?SIGMA") | |
| 40869 | 610 | proof (intro algebra.equality antisym) | 
| 40859 | 611 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 612 | using atMost_span_atLeastAtMost halfspace_le_span_atMost | |
| 613 | halfspace_span_halfspace_le open_span_halfspace | |
| 614 | by auto | |
| 615 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 616 | by (rule borel.sets_sigma_subset) auto | |
| 617 | qed auto | |
| 618 | ||
| 619 | lemma borel_eq_greaterThan: | |
| 620 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})\<rparr>)"
 | |
| 621 | (is "_ = ?SIGMA") | |
| 40869 | 622 | proof (intro algebra.equality antisym) | 
| 40859 | 623 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 624 | using halfspace_le_span_greaterThan | |
| 625 | halfspace_span_halfspace_le open_span_halfspace | |
| 626 | by auto | |
| 627 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 628 | by (rule borel.sets_sigma_subset) auto | |
| 629 | qed auto | |
| 630 | ||
| 631 | lemma borel_eq_lessThan: | |
| 632 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {..< a})\<rparr>)"
 | |
| 633 | (is "_ = ?SIGMA") | |
| 40869 | 634 | proof (intro algebra.equality antisym) | 
| 40859 | 635 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 636 | using halfspace_le_span_lessThan | |
| 637 | halfspace_span_halfspace_ge open_span_halfspace | |
| 638 | by auto | |
| 639 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 640 | by (rule borel.sets_sigma_subset) auto | |
| 641 | qed auto | |
| 642 | ||
| 643 | lemma borel_eq_greaterThanLessThan: | |
| 644 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})\<rparr>)"
 | |
| 645 | (is "_ = ?SIGMA") | |
| 40869 | 646 | proof (intro algebra.equality antisym) | 
| 40859 | 647 | show "sets ?SIGMA \<subseteq> sets borel" | 
| 648 | by (rule borel.sets_sigma_subset) auto | |
| 649 | show "sets borel \<subseteq> sets ?SIGMA" | |
| 650 | proof - | |
| 651 | have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto | |
| 652 | then interpret sigma_algebra ?SIGMA . | |
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 653 |     { fix M :: "'a set" assume "M \<in> {S. open S}"
 | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 654 | then have "open M" by simp | 
| 40859 | 655 | have "M \<in> sets ?SIGMA" | 
| 656 | apply (subst open_UNION[OF `open M`]) | |
| 657 | apply (safe intro!: countable_UN) | |
| 46731 | 658 | apply (auto simp add: sigma_def intro!: sigma_sets.Basic) | 
| 659 | done } | |
| 40859 | 660 | then show ?thesis | 
| 661 | unfolding borel_def by (intro sets_sigma_subset) auto | |
| 662 | qed | |
| 38656 | 663 | qed auto | 
| 664 | ||
| 42862 | 665 | lemma borel_eq_atLeastLessThan: | 
| 666 |   "borel = sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real})\<rparr>" (is "_ = ?S")
 | |
| 667 | proof (intro algebra.equality antisym) | |
| 668 | interpret sigma_algebra ?S | |
| 669 | by (rule sigma_algebra_sigma) auto | |
| 670 | show "sets borel \<subseteq> sets ?S" | |
| 671 | unfolding borel_eq_lessThan | |
| 672 | proof (intro sets_sigma_subset subsetI) | |
| 673 | have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto | |
| 674 | fix A :: "real set" assume "A \<in> sets \<lparr>space = UNIV, sets = range lessThan\<rparr>" | |
| 675 |     then obtain x where "A = {..< x}" by auto
 | |
| 676 |     then have "A = (\<Union>i::nat. {-real i ..< x})"
 | |
| 677 | by (auto simp: move_uminus real_arch_simple) | |
| 678 | then show "A \<in> sets ?S" | |
| 679 | by (auto simp: sets_sigma intro!: sigma_sets.intros) | |
| 680 | qed simp | |
| 681 | show "sets ?S \<subseteq> sets borel" | |
| 682 | by (intro borel.sets_sigma_subset) auto | |
| 683 | qed simp_all | |
| 684 | ||
| 40859 | 685 | lemma borel_eq_halfspace_le: | 
| 686 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})\<rparr>)"
 | |
| 687 | (is "_ = ?SIGMA") | |
| 40869 | 688 | proof (intro algebra.equality antisym) | 
| 40859 | 689 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 690 | using open_span_halfspace halfspace_span_halfspace_le by auto | |
| 691 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 692 | by (rule borel.sets_sigma_subset) auto | |
| 693 | qed auto | |
| 694 | ||
| 695 | lemma borel_eq_halfspace_less: | |
| 696 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})\<rparr>)"
 | |
| 697 | (is "_ = ?SIGMA") | |
| 40869 | 698 | proof (intro algebra.equality antisym) | 
| 40859 | 699 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 700 | using open_span_halfspace . | |
| 701 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 702 | by (rule borel.sets_sigma_subset) auto | |
| 38656 | 703 | qed auto | 
| 704 | ||
| 40859 | 705 | lemma borel_eq_halfspace_gt: | 
| 706 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})\<rparr>)"
 | |
| 707 | (is "_ = ?SIGMA") | |
| 40869 | 708 | proof (intro algebra.equality antisym) | 
| 40859 | 709 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 710 | using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto | |
| 711 | show "sets ?SIGMA \<subseteq> sets borel" | |
| 712 | by (rule borel.sets_sigma_subset) auto | |
| 713 | qed auto | |
| 38656 | 714 | |
| 40859 | 715 | lemma borel_eq_halfspace_ge: | 
| 716 |   "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})\<rparr>)"
 | |
| 717 | (is "_ = ?SIGMA") | |
| 40869 | 718 | proof (intro algebra.equality antisym) | 
| 40859 | 719 | show "sets borel \<subseteq> sets ?SIGMA" | 
| 38656 | 720 | using halfspace_span_halfspace_ge open_span_halfspace by auto | 
| 40859 | 721 | show "sets ?SIGMA \<subseteq> sets borel" | 
| 722 | by (rule borel.sets_sigma_subset) auto | |
| 723 | qed auto | |
| 38656 | 724 | |
| 725 | lemma (in sigma_algebra) borel_measurable_halfspacesI: | |
| 726 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 40859 | 727 | assumes "borel = (sigma \<lparr>space=UNIV, sets=range F\<rparr>)" | 
| 38656 | 728 | and "\<And>a i. S a i = f -` F (a,i) \<inter> space M" | 
| 729 |   and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
 | |
| 730 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
 | |
| 731 | proof safe | |
| 732 |   fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
 | |
| 733 | then show "S a i \<in> sets M" unfolding assms | |
| 734 | by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def) | |
| 735 | next | |
| 736 |   assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
 | |
| 737 |   { fix a i have "S a i \<in> sets M"
 | |
| 738 | proof cases | |
| 739 |       assume "i < DIM('c)"
 | |
| 740 | with a show ?thesis unfolding assms(2) by simp | |
| 741 | next | |
| 742 |       assume "\<not> i < DIM('c)"
 | |
| 743 | from assms(3)[OF this] show ?thesis . | |
| 744 | qed } | |
| 40859 | 745 | then have "f \<in> measurable M (sigma \<lparr>space=UNIV, sets=range F\<rparr>)" | 
| 38656 | 746 | by (auto intro!: measurable_sigma simp: assms(2)) | 
| 747 | then show "f \<in> borel_measurable M" unfolding measurable_def | |
| 748 | unfolding assms(1) by simp | |
| 749 | qed | |
| 750 | ||
| 751 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_le: | |
| 752 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 753 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
 | |
| 40859 | 754 | by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto | 
| 38656 | 755 | |
| 756 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_less: | |
| 757 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 758 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
 | |
| 40859 | 759 | by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto | 
| 38656 | 760 | |
| 761 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge: | |
| 762 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 763 |   shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
 | |
| 40859 | 764 | by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto | 
| 38656 | 765 | |
| 766 | lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater: | |
| 767 | fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" | |
| 768 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
 | |
| 40859 | 769 | by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_gt]) auto | 
| 38656 | 770 | |
| 771 | lemma (in sigma_algebra) borel_measurable_iff_le: | |
| 772 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
 | |
| 773 | using borel_measurable_iff_halfspace_le[where 'c=real] by simp | |
| 774 | ||
| 775 | lemma (in sigma_algebra) borel_measurable_iff_less: | |
| 776 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
 | |
| 777 | using borel_measurable_iff_halfspace_less[where 'c=real] by simp | |
| 778 | ||
| 779 | lemma (in sigma_algebra) borel_measurable_iff_ge: | |
| 780 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
 | |
| 781 | using borel_measurable_iff_halfspace_ge[where 'c=real] by simp | |
| 782 | ||
| 783 | lemma (in sigma_algebra) borel_measurable_iff_greater: | |
| 784 |   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
 | |
| 785 | using borel_measurable_iff_halfspace_greater[where 'c=real] by simp | |
| 786 | ||
| 41025 | 787 | lemma borel_measurable_euclidean_component: | 
| 40859 | 788 | "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel" | 
| 789 | unfolding borel_def[where 'a=real] | |
| 790 | proof (rule borel.measurable_sigma, simp_all) | |
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 791 | fix S::"real set" assume "open S" | 
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 792 | from open_vimage_euclidean_component[OF this] | 
| 40859 | 793 | show "(\<lambda>x. x $$ i) -` S \<in> sets borel" | 
| 794 | by (auto intro: borel_open) | |
| 795 | qed | |
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 796 | |
| 41025 | 797 | lemma (in sigma_algebra) borel_measurable_euclidean_space: | 
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 798 | fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 799 |   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 | 800 | proof safe | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 801 | fix i assume "f \<in> borel_measurable M" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 802 | 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 | 803 | using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def] | 
| 41025 | 804 | by (auto intro: borel_measurable_euclidean_component) | 
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 805 | next | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 806 |   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 | 807 | then show "f \<in> borel_measurable M" | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 808 | unfolding borel_measurable_iff_halfspace_le by auto | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 809 | qed | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 810 | |
| 38656 | 811 | subsection "Borel measurable operators" | 
| 812 | ||
| 813 | lemma (in sigma_algebra) affine_borel_measurable_vector: | |
| 814 | fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" | |
| 815 | assumes "f \<in> borel_measurable M" | |
| 816 | shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" | |
| 817 | proof (rule borel_measurableI) | |
| 818 | fix S :: "'x set" assume "open S" | |
| 819 | show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" | |
| 820 | proof cases | |
| 821 | assume "b \<noteq> 0" | |
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 822 | with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 823 | by (auto intro!: open_affinity simp: scaleR_add_right) | 
| 40859 | 824 | hence "?S \<in> sets borel" | 
| 825 | unfolding borel_def by (auto simp: sigma_def intro!: sigma_sets.Basic) | |
| 38656 | 826 | moreover | 
| 827 | from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" | |
| 828 | apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) | |
| 40859 | 829 | ultimately show ?thesis using assms unfolding in_borel_measurable_borel | 
| 38656 | 830 | by auto | 
| 831 | qed simp | |
| 832 | qed | |
| 833 | ||
| 834 | lemma (in sigma_algebra) affine_borel_measurable: | |
| 835 | fixes g :: "'a \<Rightarrow> real" | |
| 836 | assumes g: "g \<in> borel_measurable M" | |
| 837 | shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M" | |
| 838 | using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute) | |
| 839 | ||
| 840 | lemma (in sigma_algebra) borel_measurable_add[simp, intro]: | |
| 841 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 842 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 843 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 844 | 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 | 845 | proof - | 
| 38656 | 846 |   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 | 847 | by auto | 
| 38656 | 848 | have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M" | 
| 849 | by (rule affine_borel_measurable [OF g]) | |
| 850 |   then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
 | |
| 851 | by auto | |
| 852 |   then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
 | |
| 853 | by (simp add: 1) | |
| 854 | then show ?thesis | |
| 855 | by (simp add: borel_measurable_iff_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 856 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 857 | |
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 858 | lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]: | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 859 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 860 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 861 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 862 | proof cases | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 863 | assume "finite S" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 864 | thus ?thesis using assms by induct auto | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 865 | qed simp | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 866 | |
| 38656 | 867 | lemma (in sigma_algebra) borel_measurable_square: | 
| 868 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 869 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 870 | shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 871 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 872 |   {
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 873 | fix a | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 874 |     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 | 875 | proof (cases rule: linorder_cases [of a 0]) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 876 | case less | 
| 38656 | 877 |       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 | 878 | 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 | 879 | also have "... \<in> sets M" | 
| 38656 | 880 | by (rule empty_sets) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 881 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 882 | next | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 883 | case equal | 
| 38656 | 884 |       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 | 885 |              {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 | 886 | by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 887 | also have "... \<in> sets M" | 
| 38656 | 888 | apply (insert f) | 
| 889 | apply (rule Int) | |
| 890 | apply (simp add: borel_measurable_iff_le) | |
| 891 | apply (simp add: borel_measurable_iff_ge) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 892 | done | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 893 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 894 | next | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 895 | case greater | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 896 | 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 | 897 | 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 | 898 | real_sqrt_le_iff real_sqrt_power) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 899 |       hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
 | 
| 38656 | 900 |              {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 | 901 | using greater by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 902 | also have "... \<in> sets M" | 
| 38656 | 903 | apply (insert f) | 
| 904 | apply (rule Int) | |
| 905 | apply (simp add: borel_measurable_iff_ge) | |
| 906 | apply (simp add: borel_measurable_iff_le) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 907 | done | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 908 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 909 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 910 | } | 
| 38656 | 911 | 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 | 912 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 913 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 914 | lemma times_eq_sum_squares: | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 915 | fixes x::real | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 916 | shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4" | 
| 38656 | 917 | 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 | 918 | |
| 38656 | 919 | lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]: | 
| 920 | fixes g :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 921 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 922 | shows "(\<lambda>x. - g x) \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 923 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 924 | 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 | 925 | by simp | 
| 38656 | 926 | also have "... \<in> borel_measurable M" | 
| 927 | by (fast intro: affine_borel_measurable g) | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 928 | finally show ?thesis . | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 929 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 930 | |
| 38656 | 931 | lemma (in sigma_algebra) borel_measurable_times[simp, intro]: | 
| 932 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 933 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 934 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 935 | 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 | 936 | proof - | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 937 | have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M" | 
| 38656 | 938 | using assms by (fast intro: affine_borel_measurable borel_measurable_square) | 
| 939 | 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 | 940 | (\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))" | 
| 35582 | 941 | by (simp add: minus_divide_right) | 
| 38656 | 942 | also have "... \<in> borel_measurable M" | 
| 943 | 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 | 944 | 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 | 945 | show ?thesis | 
| 38656 | 946 | apply (simp add: times_eq_sum_squares diff_minus) | 
| 947 | using 1 2 by simp | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 948 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 949 | |
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 950 | lemma (in sigma_algebra) borel_measurable_setprod[simp, intro]: | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 951 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 952 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 953 | shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 954 | proof cases | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 955 | assume "finite S" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 956 | thus ?thesis using assms by induct auto | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 957 | qed simp | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 958 | |
| 38656 | 959 | lemma (in sigma_algebra) borel_measurable_diff[simp, intro]: | 
| 960 | fixes f :: "'a \<Rightarrow> real" | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 961 | assumes f: "f \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 962 | assumes g: "g \<in> borel_measurable M" | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 963 | shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" | 
| 38656 | 964 | unfolding diff_minus using assms by fast | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 965 | |
| 38656 | 966 | lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]: | 
| 967 | fixes f :: "'a \<Rightarrow> real" | |
| 35692 | 968 | assumes "f \<in> borel_measurable M" | 
| 969 | shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M" | |
| 38656 | 970 | unfolding borel_measurable_iff_ge unfolding inverse_eq_divide | 
| 971 | proof safe | |
| 972 | fix a :: real | |
| 973 |   have *: "{w \<in> space M. a \<le> 1 / f w} =
 | |
| 974 |       ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
 | |
| 975 |       ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
 | |
| 976 |       ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
 | |
| 977 |   show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
 | |
| 978 | by (auto intro!: Int Un) | |
| 35692 | 979 | qed | 
| 980 | ||
| 38656 | 981 | lemma (in sigma_algebra) borel_measurable_divide[simp, intro]: | 
| 982 | fixes f :: "'a \<Rightarrow> real" | |
| 35692 | 983 | assumes "f \<in> borel_measurable M" | 
| 984 | and "g \<in> borel_measurable M" | |
| 985 | shows "(\<lambda>x. f x / g x) \<in> borel_measurable M" | |
| 986 | unfolding field_divide_inverse | |
| 38656 | 987 | by (rule borel_measurable_inverse borel_measurable_times assms)+ | 
| 988 | ||
| 989 | lemma (in sigma_algebra) borel_measurable_max[intro, simp]: | |
| 990 | fixes f g :: "'a \<Rightarrow> real" | |
| 991 | assumes "f \<in> borel_measurable M" | |
| 992 | assumes "g \<in> borel_measurable M" | |
| 993 | shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" | |
| 994 | unfolding borel_measurable_iff_le | |
| 995 | proof safe | |
| 996 | fix a | |
| 997 |   have "{x \<in> space M. max (g x) (f x) \<le> a} =
 | |
| 998 |     {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
 | |
| 999 |   thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
 | |
| 1000 | using assms unfolding borel_measurable_iff_le | |
| 1001 | by (auto intro!: Int) | |
| 1002 | qed | |
| 1003 | ||
| 1004 | lemma (in sigma_algebra) borel_measurable_min[intro, simp]: | |
| 1005 | fixes f g :: "'a \<Rightarrow> real" | |
| 1006 | assumes "f \<in> borel_measurable M" | |
| 1007 | assumes "g \<in> borel_measurable M" | |
| 1008 | shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M" | |
| 1009 | unfolding borel_measurable_iff_ge | |
| 1010 | proof safe | |
| 1011 | fix a | |
| 1012 |   have "{x \<in> space M. a \<le> min (g x) (f x)} =
 | |
| 1013 |     {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
 | |
| 1014 |   thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
 | |
| 1015 | using assms unfolding borel_measurable_iff_ge | |
| 1016 | by (auto intro!: Int) | |
| 1017 | qed | |
| 1018 | ||
| 1019 | lemma (in sigma_algebra) borel_measurable_abs[simp, intro]: | |
| 1020 | assumes "f \<in> borel_measurable M" | |
| 1021 | shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" | |
| 1022 | proof - | |
| 1023 | have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def) | |
| 1024 | show ?thesis unfolding * using assms by auto | |
| 1025 | qed | |
| 1026 | ||
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1027 | lemma borel_measurable_nth[simp, intro]: | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1028 | "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel" | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1029 | using borel_measurable_euclidean_component | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1030 | unfolding nth_conv_component by auto | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1031 | |
| 41830 | 1032 | lemma borel_measurable_continuous_on1: | 
| 1033 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space" | |
| 1034 | assumes "continuous_on UNIV f" | |
| 1035 | shows "f \<in> borel_measurable borel" | |
| 1036 | apply(rule borel.borel_measurableI) | |
| 1037 | using continuous_open_preimage[OF assms] unfolding vimage_def by auto | |
| 1038 | ||
| 1039 | lemma borel_measurable_continuous_on: | |
| 1040 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space" | |
| 42990 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1041 | assumes cont: "continuous_on A f" "open A" | 
| 41830 | 1042 | shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _") | 
| 1043 | proof (rule borel.borel_measurableI) | |
| 1044 | fix S :: "'b set" assume "open S" | |
| 42990 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1045 |   then have "open {x\<in>A. f x \<in> S}"
 | 
| 41830 | 1046 | by (intro continuous_open_preimage[OF cont]) auto | 
| 42990 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1047 |   then have *: "{x\<in>A. f x \<in> S} \<in> sets borel" by auto
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1048 | have "?f -` S \<inter> space borel = | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1049 |     {x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1050 | by (auto split: split_if_asm) | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1051 | also have "\<dots> \<in> sets borel" | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1052 | using * `open A` by (auto simp del: space_borel intro!: borel.Un) | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1053 | finally show "?f -` S \<inter> space borel \<in> sets borel" . | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1054 | qed | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1055 | |
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1056 | lemma (in sigma_algebra) convex_measurable: | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1057 | fixes a b :: real | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1058 |   assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> { a <..< b}"
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1059 |   assumes q: "convex_on { a <..< b} q"
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1060 | shows "q \<circ> X \<in> borel_measurable M" | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1061 | proof - | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1062 |   have "(\<lambda>x. if x \<in> {a <..< b} then q x else 0) \<in> borel_measurable borel"
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1063 | proof (rule borel_measurable_continuous_on) | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1064 |     show "open {a<..<b}" by auto
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1065 |     from this q show "continuous_on {a<..<b} q"
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1066 | by (rule convex_on_continuous) | 
| 41830 | 1067 | qed | 
| 42990 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1068 |   then have "(\<lambda>x. if x \<in> {a <..< b} then q x else 0) \<circ> X \<in> borel_measurable M" (is ?qX)
 | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1069 | using X by (intro measurable_comp) auto | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1070 | moreover have "?qX \<longleftrightarrow> q \<circ> X \<in> borel_measurable M" | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1071 | using X by (intro measurable_cong) auto | 
| 
3706951a6421
composition of convex and measurable function is measurable
 hoelzl parents: 
42950diff
changeset | 1072 | ultimately show ?thesis by simp | 
| 41830 | 1073 | qed | 
| 1074 | ||
| 1075 | lemma borel_measurable_borel_log: assumes "1 < b" shows "log b \<in> borel_measurable borel" | |
| 1076 | proof - | |
| 1077 |   { fix x :: real assume x: "x \<le> 0"
 | |
| 1078 |     { fix x::real assume "x \<le> 0" then have "\<And>u. exp u = x \<longleftrightarrow> False" by auto }
 | |
| 1079 | from this[of x] x this[of 0] have "log b 0 = log b x" | |
| 1080 | by (auto simp: ln_def log_def) } | |
| 1081 | note log_imp = this | |
| 1082 |   have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) \<in> borel_measurable borel"
 | |
| 1083 | proof (rule borel_measurable_continuous_on) | |
| 1084 |     show "continuous_on {0<..} (log b)"
 | |
| 1085 | by (auto intro!: continuous_at_imp_continuous_on DERIV_log DERIV_isCont | |
| 1086 | simp: continuous_isCont[symmetric]) | |
| 1087 |     show "open ({0<..}::real set)" by auto
 | |
| 1088 | qed | |
| 1089 |   also have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) = log b"
 | |
| 1090 | by (simp add: fun_eq_iff not_less log_imp) | |
| 1091 | finally show ?thesis . | |
| 1092 | qed | |
| 1093 | ||
| 1094 | lemma (in sigma_algebra) borel_measurable_log[simp,intro]: | |
| 1095 | assumes f: "f \<in> borel_measurable M" and "1 < b" | |
| 1096 | shows "(\<lambda>x. log b (f x)) \<in> borel_measurable M" | |
| 1097 | using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]] | |
| 1098 | by (simp add: comp_def) | |
| 1099 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1100 | subsection "Borel space on the extended reals" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1101 | |
| 43920 | 1102 | lemma borel_measurable_ereal_borel: | 
| 1103 | "ereal \<in> borel_measurable borel" | |
| 1104 | unfolding borel_def[where 'a=ereal] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1105 | proof (rule borel.measurable_sigma) | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1106 |   fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = {S. open S} \<rparr>"
 | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1107 | then have "open X" by simp | 
| 43920 | 1108 | then have "open (ereal -` X \<inter> space borel)" | 
| 1109 | by (simp add: open_ereal_vimage) | |
| 1110 | then show "ereal -` X \<inter> space borel \<in> sets borel" by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1111 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1112 | |
| 43920 | 1113 | lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]: | 
| 1114 | assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" | |
| 1115 | using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def . | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1116 | |
| 43920 | 1117 | lemma borel_measurable_real_of_ereal_borel: | 
| 1118 | "(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1119 | unfolding borel_def[where 'a=real] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1120 | proof (rule borel.measurable_sigma) | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1121 |   fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = {S. open S} \<rparr>"
 | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1122 | then have "open B" by simp | 
| 43920 | 1123 |   have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
 | 
| 1124 |   have open_real: "open (real -` (B - {0}) :: ereal set)"
 | |
| 1125 | unfolding open_ereal_def * using `open B` by auto | |
| 1126 | show "(real -` B \<inter> space borel :: ereal set) \<in> sets borel" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1127 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1128 | assume "0 \<in> B" | 
| 43923 | 1129 |     then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0::ereal}"
 | 
| 43920 | 1130 | by (auto simp add: real_of_ereal_eq_0) | 
| 1131 | then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1132 | using open_real by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1133 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1134 | assume "0 \<notin> B" | 
| 43920 | 1135 |     then have *: "(real -` B :: ereal set) = real -` (B - {0})"
 | 
| 1136 | by (auto simp add: real_of_ereal_eq_0) | |
| 1137 | then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1138 | using open_real by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1139 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1140 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1141 | |
| 43920 | 1142 | lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]: | 
| 1143 | assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M" | |
| 1144 | using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def . | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1145 | |
| 43920 | 1146 | lemma (in sigma_algebra) borel_measurable_ereal_iff: | 
| 1147 | shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1148 | proof | 
| 43920 | 1149 | assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" | 
| 1150 | from borel_measurable_real_of_ereal[OF this] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1151 | show "f \<in> borel_measurable M" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1152 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1153 | |
| 43920 | 1154 | lemma (in sigma_algebra) borel_measurable_ereal_iff_real: | 
| 43923 | 1155 | fixes f :: "'a \<Rightarrow> ereal" | 
| 1156 | shows "f \<in> borel_measurable M \<longleftrightarrow> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1157 |     ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1158 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1159 |   assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1160 |   have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1161 |   with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
 | 
| 46731 | 1162 | let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real (f x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1163 | have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto | 
| 43920 | 1164 | also have "?f = f" by (auto simp: fun_eq_iff ereal_real) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1165 | finally show "f \<in> borel_measurable M" . | 
| 43920 | 1166 | qed (auto intro: measurable_sets borel_measurable_real_of_ereal) | 
| 41830 | 1167 | |
| 38656 | 1168 | lemma (in sigma_algebra) less_eq_ge_measurable: | 
| 1169 | fixes f :: "'a \<Rightarrow> 'c::linorder" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1170 |   shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
 | 
| 38656 | 1171 | proof | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1172 |   assume "f -` {a <..} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1173 |   moreover have "f -` {..a} \<inter> space M = space M - f -` {a <..} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1174 |   ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1175 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1176 |   assume "f -` {..a} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1177 |   moreover have "f -` {a <..} \<inter> space M = space M - f -` {..a} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1178 |   ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1179 | qed | 
| 35692 | 1180 | |
| 38656 | 1181 | lemma (in sigma_algebra) greater_eq_le_measurable: | 
| 1182 | fixes f :: "'a \<Rightarrow> 'c::linorder" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1183 |   shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
 | 
| 38656 | 1184 | proof | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1185 |   assume "f -` {a ..} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1186 |   moreover have "f -` {..< a} \<inter> space M = space M - f -` {a ..} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1187 |   ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1188 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1189 |   assume "f -` {..< a} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1190 |   moreover have "f -` {a ..} \<inter> space M = space M - f -` {..< a} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1191 |   ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1192 | qed | 
| 1193 | ||
| 43920 | 1194 | lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal: | 
| 1195 | "(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1196 | proof (subst borel_def, rule borel.measurable_sigma) | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1197 |   fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = {S. open S}\<rparr>"
 | 
| 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 1198 | then have "open X" by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1199 | have "uminus -` X = uminus ` X" by (force simp: image_iff) | 
| 43920 | 1200 | then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1201 | then show "uminus -` X \<inter> space borel \<in> sets borel" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1202 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1203 | |
| 43920 | 1204 | lemma (in sigma_algebra) borel_measurable_uminus_ereal[intro]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1205 | assumes "f \<in> borel_measurable M" | 
| 43920 | 1206 | shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M" | 
| 1207 | using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1208 | |
| 43920 | 1209 | lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]: | 
| 1210 | "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") | |
| 38656 | 1211 | proof | 
| 43920 | 1212 | assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1213 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1214 | |
| 43920 | 1215 | lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal: | 
| 43923 | 1216 | fixes f :: "'a \<Rightarrow> ereal" | 
| 1217 |   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1218 | proof (intro iffI allI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1219 |   assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1220 | show "f \<in> borel_measurable M" | 
| 43920 | 1221 | unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1222 | proof (intro conjI allI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1223 | fix a :: real | 
| 43920 | 1224 |     { fix x :: ereal assume *: "\<forall>i::nat. real i < x"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1225 | have "x = \<infinity>" | 
| 43920 | 1226 | proof (rule ereal_top) | 
| 44666 | 1227 | fix B from reals_Archimedean2[of B] guess n .. | 
| 43920 | 1228 | then have "ereal B < real n" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1229 | with * show "B \<le> x" by (metis less_trans less_imp_le) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1230 | qed } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1231 |     then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1232 | by (auto simp: not_le) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1233 |     then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1234 | moreover | 
| 43923 | 1235 |     have "{-\<infinity>::ereal} = {..-\<infinity>}" by auto
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1236 |     then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
 | 
| 43920 | 1237 |     moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
 | 
| 1238 | using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1239 |     moreover have "{w \<in> space M. real (f w) \<le> a} =
 | 
| 43920 | 1240 |       (if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
 | 
| 1241 |       else {w \<in> space M. f w \<le> ereal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1242 | proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1243 |     ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
 | 
| 35582 | 1244 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1245 | qed (simp add: measurable_sets) | 
| 35582 | 1246 | |
| 43920 | 1247 | lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal: | 
| 1248 |   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1249 | proof | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1250 |   assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1251 |   moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
 | 
| 43920 | 1252 | by (auto simp: ereal_uminus_le_reorder) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1253 | ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M" | 
| 43920 | 1254 | unfolding borel_measurable_eq_atMost_ereal by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1255 | then show "f \<in> borel_measurable M" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1256 | qed (simp add: measurable_sets) | 
| 35582 | 1257 | |
| 43920 | 1258 | lemma (in sigma_algebra) borel_measurable_ereal_iff_less: | 
| 1259 |   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
 | |
| 1260 | unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable .. | |
| 38656 | 1261 | |
| 43920 | 1262 | lemma (in sigma_algebra) borel_measurable_ereal_iff_ge: | 
| 1263 |   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
 | |
| 1264 | unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable .. | |
| 38656 | 1265 | |
| 43920 | 1266 | lemma (in sigma_algebra) borel_measurable_ereal_eq_const: | 
| 1267 | fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" | |
| 38656 | 1268 |   shows "{x\<in>space M. f x = c} \<in> sets M"
 | 
| 1269 | proof - | |
| 1270 |   have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
 | |
| 1271 | then show ?thesis using assms by (auto intro!: measurable_sets) | |
| 1272 | qed | |
| 1273 | ||
| 43920 | 1274 | lemma (in sigma_algebra) borel_measurable_ereal_neq_const: | 
| 1275 | fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" | |
| 38656 | 1276 |   shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 | 
| 1277 | proof - | |
| 1278 |   have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
 | |
| 1279 | then show ?thesis using assms by (auto intro!: measurable_sets) | |
| 1280 | qed | |
| 1281 | ||
| 43920 | 1282 | lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]: | 
| 1283 | fixes f g :: "'a \<Rightarrow> ereal" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1284 | assumes f: "f \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1285 | assumes g: "g \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1286 |   shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1287 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1288 |   have "{x \<in> space M. f x \<le> g x} =
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1289 |     {x \<in> space M. real (f x) \<le> real (g x)} - (f -` {\<infinity>, -\<infinity>} \<inter> space M \<union> g -` {\<infinity>, -\<infinity>} \<inter> space M) \<union>
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1290 |     f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1291 | proof (intro set_eqI) | 
| 43920 | 1292 | fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: ereal2_cases[of "f x" "g x"]) auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1293 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1294 | with f g show ?thesis by (auto intro!: Un simp: measurable_sets) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1295 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1296 | |
| 43920 | 1297 | lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]: | 
| 1298 | fixes f :: "'a \<Rightarrow> ereal" | |
| 38656 | 1299 | assumes f: "f \<in> borel_measurable M" | 
| 1300 | assumes g: "g \<in> borel_measurable M" | |
| 1301 |   shows "{x \<in> space M. f x < g x} \<in> sets M"
 | |
| 1302 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1303 |   have "{x \<in> space M. f x < g x} = space M - {x \<in> space M. g x \<le> f x}" by auto
 | 
| 38656 | 1304 | then show ?thesis using g f by auto | 
| 1305 | qed | |
| 1306 | ||
| 43920 | 1307 | lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]: | 
| 1308 | fixes f :: "'a \<Rightarrow> ereal" | |
| 38656 | 1309 | assumes f: "f \<in> borel_measurable M" | 
| 1310 | assumes g: "g \<in> borel_measurable M" | |
| 1311 |   shows "{w \<in> space M. f w = g w} \<in> sets M"
 | |
| 1312 | proof - | |
| 1313 |   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
 | |
| 1314 | then show ?thesis using g f by auto | |
| 1315 | qed | |
| 1316 | ||
| 43920 | 1317 | lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]: | 
| 1318 | fixes f :: "'a \<Rightarrow> ereal" | |
| 38656 | 1319 | assumes f: "f \<in> borel_measurable M" | 
| 1320 | assumes g: "g \<in> borel_measurable M" | |
| 1321 |   shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | |
| 35692 | 1322 | proof - | 
| 38656 | 1323 |   have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
 | 
| 1324 | thus ?thesis using f g by auto | |
| 1325 | qed | |
| 1326 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1327 | lemma (in sigma_algebra) split_sets: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1328 |   "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1329 |   "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1330 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1331 | |
| 43920 | 1332 | lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]: | 
| 1333 | fixes f :: "'a \<Rightarrow> ereal" | |
| 41025 | 1334 | assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 38656 | 1335 | shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" | 
| 1336 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1337 |   { fix x assume "x \<in> space M" then have "f x + g x =
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1338 | (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1339 | else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity> | 
| 43920 | 1340 | else ereal (real (f x) + real (g x)))" | 
| 1341 | by (cases rule: ereal2_cases[of "f x" "g x"]) auto } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1342 | with assms show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1343 | by (auto cong: measurable_cong simp: split_sets | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1344 | intro!: Un measurable_If measurable_sets) | 
| 38656 | 1345 | qed | 
| 1346 | ||
| 43920 | 1347 | lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]: | 
| 1348 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" | |
| 41096 | 1349 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 1350 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" | |
| 1351 | proof cases | |
| 1352 | assume "finite S" | |
| 1353 | thus ?thesis using assms | |
| 1354 | by induct auto | |
| 1355 | qed (simp add: borel_measurable_const) | |
| 1356 | ||
| 43920 | 1357 | lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]: | 
| 1358 | fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1359 | shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1360 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1361 |   { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1362 | then show ?thesis using assms by (auto intro!: measurable_If) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1363 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1364 | |
| 43920 | 1365 | lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]: | 
| 1366 | fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" | |
| 38656 | 1367 | shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" | 
| 1368 | proof - | |
| 43920 | 1369 |   { fix f g :: "'a \<Rightarrow> ereal"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1370 | assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1371 | and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1372 |     { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1373 | else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity> | 
| 43920 | 1374 | else ereal (real (f x) * real (g x)))" | 
| 1375 | apply (cases rule: ereal2_cases[of "f x" "g x"]) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1376 | using pos[of x] by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1377 | with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1378 | by (auto cong: measurable_cong simp: split_sets | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1379 | intro!: Un measurable_If measurable_sets) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1380 | note pos_times = this | 
| 38656 | 1381 | have *: "(\<lambda>x. f x * g x) = | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1382 | (\<lambda>x. if 0 \<le> f x \<and> 0 \<le> g x \<or> f x < 0 \<and> g x < 0 then \<bar>f x\<bar> * \<bar>g x\<bar> else - (\<bar>f x\<bar> * \<bar>g x\<bar>))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1383 | by (auto simp: fun_eq_iff) | 
| 38656 | 1384 | show ?thesis using assms unfolding * | 
| 43920 | 1385 | by (intro measurable_If pos_times borel_measurable_uminus_ereal) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1386 | (auto simp: split_sets intro!: Int) | 
| 38656 | 1387 | qed | 
| 1388 | ||
| 43920 | 1389 | lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]: | 
| 1390 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" | |
| 38656 | 1391 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 41096 | 1392 | shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" | 
| 38656 | 1393 | proof cases | 
| 1394 | assume "finite S" | |
| 41096 | 1395 | thus ?thesis using assms by induct auto | 
| 1396 | qed simp | |
| 38656 | 1397 | |
| 43920 | 1398 | lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]: | 
| 1399 | fixes f g :: "'a \<Rightarrow> ereal" | |
| 38656 | 1400 | assumes "f \<in> borel_measurable M" | 
| 1401 | assumes "g \<in> borel_measurable M" | |
| 1402 | shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M" | |
| 1403 | using assms unfolding min_def by (auto intro!: measurable_If) | |
| 1404 | ||
| 43920 | 1405 | lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]: | 
| 1406 | fixes f g :: "'a \<Rightarrow> ereal" | |
| 38656 | 1407 | assumes "f \<in> borel_measurable M" | 
| 1408 | and "g \<in> borel_measurable M" | |
| 1409 | shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" | |
| 1410 | using assms unfolding max_def by (auto intro!: measurable_If) | |
| 1411 | ||
| 1412 | lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]: | |
| 43920 | 1413 | fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 38656 | 1414 | assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M" | 
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41096diff
changeset | 1415 | shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M") | 
| 43920 | 1416 | unfolding borel_measurable_ereal_iff_ge | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1417 | proof | 
| 38656 | 1418 | fix a | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1419 |   have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
 | 
| 46884 | 1420 | by (auto simp: less_SUP_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1421 |   then show "?sup -` {a<..} \<inter> space M \<in> sets M"
 | 
| 38656 | 1422 | using assms by auto | 
| 1423 | qed | |
| 1424 | ||
| 1425 | lemma (in sigma_algebra) borel_measurable_INF[simp, intro]: | |
| 43920 | 1426 | fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 38656 | 1427 | assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M" | 
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41096diff
changeset | 1428 | shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M") | 
| 43920 | 1429 | unfolding borel_measurable_ereal_iff_less | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1430 | proof | 
| 38656 | 1431 | fix a | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1432 |   have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
 | 
| 46884 | 1433 | by (auto simp: INF_less_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1434 |   then show "?inf -` {..<a} \<inter> space M \<in> sets M"
 | 
| 38656 | 1435 | using assms by auto | 
| 1436 | qed | |
| 1437 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1438 | lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]: | 
| 43920 | 1439 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1440 | assumes "\<And>i. f i \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1441 | shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1442 | unfolding liminf_SUPR_INFI using assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1443 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1444 | lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]: | 
| 43920 | 1445 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1446 | assumes "\<And>i. f i \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1447 | shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1448 | unfolding limsup_INFI_SUPR using assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1449 | |
| 43920 | 1450 | lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]: | 
| 1451 | fixes f g :: "'a \<Rightarrow> ereal" | |
| 38656 | 1452 | assumes "f \<in> borel_measurable M" | 
| 1453 | assumes "g \<in> borel_measurable M" | |
| 1454 | shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" | |
| 43920 | 1455 | unfolding minus_ereal_def using assms by auto | 
| 35692 | 1456 | |
| 40870 
94427db32392
Tuned setup for borel_measurable with min, max and psuminf.
 hoelzl parents: 
40869diff
changeset | 1457 | lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]: | 
| 43920 | 1458 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1459 | assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1460 | shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1461 | apply (subst measurable_cong) | 
| 43920 | 1462 | apply (subst suminf_ereal_eq_SUPR) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1463 | apply (rule pos) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1464 | using assms by auto | 
| 39092 | 1465 | |
| 1466 | section "LIMSEQ is borel measurable" | |
| 1467 | ||
| 1468 | lemma (in sigma_algebra) borel_measurable_LIMSEQ: | |
| 1469 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" | |
| 1470 | assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" | |
| 1471 | and u: "\<And>i. u i \<in> borel_measurable M" | |
| 1472 | shows "u' \<in> borel_measurable M" | |
| 1473 | proof - | |
| 43920 | 1474 | have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)" | 
| 46731 | 1475 | using u' by (simp add: lim_imp_Liminf) | 
| 43920 | 1476 | moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M" | 
| 39092 | 1477 | by auto | 
| 43920 | 1478 | ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) | 
| 39092 | 1479 | qed | 
| 1480 | ||
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 1481 | end |