| author | hoelzl | 
| Tue, 19 Jul 2011 14:38:29 +0200 | |
| changeset 43923 | ab93d0190a5d | 
| parent 43920 | cedb5cb948fd | 
| child 44282 | f0de18b62d63 | 
| 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  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
9  | 
imports Sigma_Algebra 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  | 
|
| 40859 | 14  | 
definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = open\<rparr>"  | 
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>  | 
| 40859 | 22  | 
(\<forall>S \<in> sets (sigma \<lparr> space = UNIV, sets = open\<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 -  | 
38  | 
have "A \<in> open" unfolding mem_def 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)  | 
|
| 38656 | 74  | 
fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M"  | 
75  | 
using assms[of S] by (simp add: mem_def)  | 
|
| 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"  | 
|
93  | 
unfolding indicator_def_raw using A  | 
|
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"
 | 
|
104  | 
unfolding indicator_def_raw by auto  | 
|
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"  | 
|
130  | 
using `A \<in> sets M` sets_into_space by fastsimp  | 
|
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 .  | 
| 40859 | 394  | 
  { fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def .
 | 
395  | 
from open_UNION[OF this]  | 
|
396  | 
obtain I where *: "S =  | 
|
397  | 
(\<Union>(a, b)\<in>I.  | 
|
398  | 
          (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
 | 
|
399  | 
          (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
 | 
|
400  | 
unfolding greaterThanLessThan_def  | 
|
401  | 
unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]  | 
|
402  | 
unfolding eucl_lessThan_eq_halfspaces[where 'a='a]  | 
|
403  | 
by blast  | 
|
404  | 
have "S \<in> sets ?SIGMA"  | 
|
405  | 
unfolding *  | 
|
406  | 
by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace) }  | 
|
407  | 
then show ?thesis unfolding borel_def  | 
|
408  | 
by (intro sets_sigma_subset) auto  | 
|
409  | 
qed  | 
|
| 38656 | 410  | 
|
411  | 
lemma halfspace_span_halfspace_le:  | 
|
| 40859 | 412  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
 | 
413  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. x $$ i \<le> a})\<rparr>)"
 | 
|
| 38656 | 414  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
| 40859 | 415  | 
proof -  | 
416  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
| 38656 | 417  | 
then interpret sigma_algebra ?SIGMA .  | 
| 40859 | 418  | 
  { fix a i
 | 
419  | 
    have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
 | 
|
420  | 
proof (safe, simp_all)  | 
|
421  | 
fix x::'a assume *: "x$$i < a"  | 
|
422  | 
with reals_Archimedean[of "a - x$$i"]  | 
|
423  | 
obtain n where "x $$ i < a - 1 / (real (Suc n))"  | 
|
424  | 
by (auto simp: field_simps inverse_eq_divide)  | 
|
425  | 
then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"  | 
|
426  | 
by (blast intro: less_imp_le)  | 
|
427  | 
next  | 
|
428  | 
fix x::'a and n  | 
|
429  | 
assume "x$$i \<le> a - 1 / real (Suc n)"  | 
|
430  | 
also have "\<dots> < a" by auto  | 
|
431  | 
finally show "x$$i < a" .  | 
|
432  | 
qed  | 
|
433  | 
    have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | 
|
434  | 
by (safe intro!: countable_UN)  | 
|
435  | 
(auto simp: sets_sigma intro!: sigma_sets.Basic) }  | 
|
436  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
|
437  | 
qed  | 
|
| 38656 | 438  | 
|
439  | 
lemma halfspace_span_halfspace_ge:  | 
|
| 40859 | 440  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
 | 
441  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a \<le> x $$ i})\<rparr>)"
 | 
|
| 38656 | 442  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
| 40859 | 443  | 
proof -  | 
444  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
| 38656 | 445  | 
then interpret sigma_algebra ?SIGMA .  | 
| 40859 | 446  | 
  { fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
 | 
447  | 
    have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
 | 
|
448  | 
by (safe intro!: Diff)  | 
|
449  | 
(auto simp: sets_sigma intro!: sigma_sets.Basic) }  | 
|
450  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
|
451  | 
qed  | 
|
| 38656 | 452  | 
|
453  | 
lemma halfspace_le_span_halfspace_gt:  | 
|
| 40859 | 454  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
455  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a < x $$ i})\<rparr>)"
 | 
|
| 38656 | 456  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
| 40859 | 457  | 
proof -  | 
458  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
| 38656 | 459  | 
then interpret sigma_algebra ?SIGMA .  | 
| 40859 | 460  | 
  { fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | 
461  | 
    have "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | 
|
462  | 
by (safe intro!: Diff)  | 
|
463  | 
(auto simp: sets_sigma intro!: sigma_sets.Basic) }  | 
|
464  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
|
465  | 
qed  | 
|
| 38656 | 466  | 
|
467  | 
lemma halfspace_le_span_atMost:  | 
|
| 40859 | 468  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
469  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>)"
 | 
|
| 38656 | 470  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
| 40859 | 471  | 
proof -  | 
472  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
| 38656 | 473  | 
then interpret sigma_algebra ?SIGMA .  | 
| 40859 | 474  | 
  have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 38656 | 475  | 
proof cases  | 
| 40859 | 476  | 
    fix a i assume "i < DIM('a)"
 | 
| 38656 | 477  | 
    then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
 | 
478  | 
proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)  | 
|
479  | 
fix x  | 
|
480  | 
      from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
 | 
|
481  | 
      then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
 | 
|
482  | 
by (subst (asm) Max_le_iff) auto  | 
|
483  | 
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
 | 
|
484  | 
by (auto intro!: exI[of _ k])  | 
|
485  | 
qed  | 
|
486  | 
    show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
 | 
|
487  | 
by (safe intro!: countable_UN)  | 
|
488  | 
(auto simp: sets_sigma intro!: sigma_sets.Basic)  | 
|
489  | 
next  | 
|
| 40859 | 490  | 
    fix a i assume "\<not> i < DIM('a)"
 | 
| 38656 | 491  | 
    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
492  | 
using top by auto  | 
|
493  | 
qed  | 
|
| 40859 | 494  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
495  | 
qed  | 
|
| 38656 | 496  | 
|
497  | 
lemma halfspace_le_span_greaterThan:  | 
|
| 40859 | 498  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
 | 
499  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {a<..})\<rparr>)"
 | 
|
| 38656 | 500  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
| 40859 | 501  | 
proof -  | 
502  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
| 38656 | 503  | 
then interpret sigma_algebra ?SIGMA .  | 
| 40859 | 504  | 
  have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
| 38656 | 505  | 
proof cases  | 
| 40859 | 506  | 
    fix a i assume "i < DIM('a)"
 | 
| 38656 | 507  | 
    have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
 | 
508  | 
    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)`
 | 
|
509  | 
proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)  | 
|
510  | 
fix x  | 
|
511  | 
      from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
 | 
|
512  | 
guess k::nat .. note k = this  | 
|
513  | 
      { fix i assume "i < DIM('a)"
 | 
|
514  | 
then have "-x$$i < real k"  | 
|
515  | 
using k by (subst (asm) Max_less_iff) auto  | 
|
516  | 
then have "- real k < x$$i" by simp }  | 
|
517  | 
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
 | 
|
518  | 
by (auto intro!: exI[of _ k])  | 
|
519  | 
qed  | 
|
520  | 
    finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
|
521  | 
apply (simp only:)  | 
|
522  | 
apply (safe intro!: countable_UN Diff)  | 
|
523  | 
by (auto simp: sets_sigma intro!: sigma_sets.Basic)  | 
|
524  | 
next  | 
|
| 40859 | 525  | 
    fix a i assume "\<not> i < DIM('a)"
 | 
| 38656 | 526  | 
    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
 | 
527  | 
using top by auto  | 
|
528  | 
qed  | 
|
| 40859 | 529  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
530  | 
qed  | 
|
531  | 
||
532  | 
lemma halfspace_le_span_lessThan:  | 
|
533  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i})\<rparr>) \<subseteq>
 | 
|
534  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..<a})\<rparr>)"
 | 
|
535  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
|
536  | 
proof -  | 
|
537  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
538  | 
then interpret sigma_algebra ?SIGMA .  | 
|
539  | 
  have "\<And>a i. {x. a \<le> x$$i} \<in> sets ?SIGMA"
 | 
|
540  | 
proof cases  | 
|
541  | 
    fix a i assume "i < DIM('a)"
 | 
|
542  | 
    have "{x::'a. a \<le> x$$i} = space ?SIGMA - {x::'a. x$$i < a}" by auto
 | 
|
543  | 
    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)`
 | 
|
544  | 
proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)  | 
|
545  | 
fix x  | 
|
546  | 
      from real_arch_lt[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
 | 
|
547  | 
guess k::nat .. note k = this  | 
|
548  | 
      { fix i assume "i < DIM('a)"
 | 
|
549  | 
then have "x$$i < real k"  | 
|
550  | 
using k by (subst (asm) Max_less_iff) auto  | 
|
551  | 
then have "x$$i < real k" by simp }  | 
|
552  | 
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
 | 
|
553  | 
by (auto intro!: exI[of _ k])  | 
|
554  | 
qed  | 
|
555  | 
    finally show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
 | 
|
556  | 
apply (simp only:)  | 
|
557  | 
apply (safe intro!: countable_UN Diff)  | 
|
558  | 
by (auto simp: sets_sigma intro!: sigma_sets.Basic)  | 
|
559  | 
next  | 
|
560  | 
    fix a i assume "\<not> i < DIM('a)"
 | 
|
561  | 
    then show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
 | 
|
562  | 
using top by auto  | 
|
563  | 
qed  | 
|
564  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
|
565  | 
qed  | 
|
566  | 
||
567  | 
lemma atMost_span_atLeastAtMost:  | 
|
568  | 
  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>) \<subseteq>
 | 
|
569  | 
   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a,b). {a..b})\<rparr>)"
 | 
|
570  | 
(is "_ \<subseteq> sets ?SIGMA")  | 
|
571  | 
proof -  | 
|
572  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
573  | 
then interpret sigma_algebra ?SIGMA .  | 
|
574  | 
  { fix a::'a
 | 
|
575  | 
    have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
 | 
|
576  | 
proof (safe, simp_all add: eucl_le[where 'a='a])  | 
|
577  | 
fix x  | 
|
578  | 
      from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
 | 
|
579  | 
guess k::nat .. note k = this  | 
|
580  | 
      { fix i assume "i < DIM('a)"
 | 
|
581  | 
with k have "- x$$i \<le> real k"  | 
|
582  | 
by (subst (asm) Max_le_iff) (auto simp: field_simps)  | 
|
583  | 
then have "- real k \<le> x$$i" by simp }  | 
|
584  | 
      then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
 | 
|
585  | 
by (auto intro!: exI[of _ k])  | 
|
586  | 
qed  | 
|
587  | 
    have "{..a} \<in> sets ?SIGMA" unfolding *
 | 
|
588  | 
by (safe intro!: countable_UN)  | 
|
589  | 
(auto simp: sets_sigma intro!: sigma_sets.Basic) }  | 
|
590  | 
then show ?thesis by (intro sets_sigma_subset) auto  | 
|
591  | 
qed  | 
|
592  | 
||
593  | 
lemma borel_eq_atMost:  | 
|
594  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})\<rparr>)"
 | 
|
595  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 596  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 597  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
598  | 
using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace  | 
|
599  | 
by auto  | 
|
600  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
601  | 
by (rule borel.sets_sigma_subset) auto  | 
|
602  | 
qed auto  | 
|
603  | 
||
604  | 
lemma borel_eq_atLeastAtMost:  | 
|
605  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})\<rparr>)"
 | 
|
606  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 607  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 608  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
609  | 
using atMost_span_atLeastAtMost halfspace_le_span_atMost  | 
|
610  | 
halfspace_span_halfspace_le open_span_halfspace  | 
|
611  | 
by auto  | 
|
612  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
613  | 
by (rule borel.sets_sigma_subset) auto  | 
|
614  | 
qed auto  | 
|
615  | 
||
616  | 
lemma borel_eq_greaterThan:  | 
|
617  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})\<rparr>)"
 | 
|
618  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 619  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 620  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
621  | 
using halfspace_le_span_greaterThan  | 
|
622  | 
halfspace_span_halfspace_le open_span_halfspace  | 
|
623  | 
by auto  | 
|
624  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
625  | 
by (rule borel.sets_sigma_subset) auto  | 
|
626  | 
qed auto  | 
|
627  | 
||
628  | 
lemma borel_eq_lessThan:  | 
|
629  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {..< a})\<rparr>)"
 | 
|
630  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 631  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 632  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
633  | 
using halfspace_le_span_lessThan  | 
|
634  | 
halfspace_span_halfspace_ge open_span_halfspace  | 
|
635  | 
by auto  | 
|
636  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
637  | 
by (rule borel.sets_sigma_subset) auto  | 
|
638  | 
qed auto  | 
|
639  | 
||
640  | 
lemma borel_eq_greaterThanLessThan:  | 
|
641  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})\<rparr>)"
 | 
|
642  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 643  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 644  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
645  | 
by (rule borel.sets_sigma_subset) auto  | 
|
646  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
|
647  | 
proof -  | 
|
648  | 
have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto  | 
|
649  | 
then interpret sigma_algebra ?SIGMA .  | 
|
650  | 
    { fix M :: "'a set" assume "M \<in> open"
 | 
|
651  | 
then have "open M" by (simp add: mem_def)  | 
|
652  | 
have "M \<in> sets ?SIGMA"  | 
|
653  | 
apply (subst open_UNION[OF `open M`])  | 
|
654  | 
apply (safe intro!: countable_UN)  | 
|
655  | 
by (auto simp add: sigma_def intro!: sigma_sets.Basic) }  | 
|
656  | 
then show ?thesis  | 
|
657  | 
unfolding borel_def by (intro sets_sigma_subset) auto  | 
|
658  | 
qed  | 
|
| 38656 | 659  | 
qed auto  | 
660  | 
||
| 42862 | 661  | 
lemma borel_eq_atLeastLessThan:  | 
662  | 
  "borel = sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real})\<rparr>" (is "_ = ?S")
 | 
|
663  | 
proof (intro algebra.equality antisym)  | 
|
664  | 
interpret sigma_algebra ?S  | 
|
665  | 
by (rule sigma_algebra_sigma) auto  | 
|
666  | 
show "sets borel \<subseteq> sets ?S"  | 
|
667  | 
unfolding borel_eq_lessThan  | 
|
668  | 
proof (intro sets_sigma_subset subsetI)  | 
|
669  | 
have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto  | 
|
670  | 
fix A :: "real set" assume "A \<in> sets \<lparr>space = UNIV, sets = range lessThan\<rparr>"  | 
|
671  | 
    then obtain x where "A = {..< x}" by auto
 | 
|
672  | 
    then have "A = (\<Union>i::nat. {-real i ..< x})"
 | 
|
673  | 
by (auto simp: move_uminus real_arch_simple)  | 
|
674  | 
then show "A \<in> sets ?S"  | 
|
675  | 
by (auto simp: sets_sigma intro!: sigma_sets.intros)  | 
|
676  | 
qed simp  | 
|
677  | 
show "sets ?S \<subseteq> sets borel"  | 
|
678  | 
by (intro borel.sets_sigma_subset) auto  | 
|
679  | 
qed simp_all  | 
|
680  | 
||
| 40859 | 681  | 
lemma borel_eq_halfspace_le:  | 
682  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})\<rparr>)"
 | 
|
683  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 684  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 685  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
686  | 
using open_span_halfspace halfspace_span_halfspace_le by auto  | 
|
687  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
688  | 
by (rule borel.sets_sigma_subset) auto  | 
|
689  | 
qed auto  | 
|
690  | 
||
691  | 
lemma borel_eq_halfspace_less:  | 
|
692  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})\<rparr>)"
 | 
|
693  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 694  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 695  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
696  | 
using open_span_halfspace .  | 
|
697  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
698  | 
by (rule borel.sets_sigma_subset) auto  | 
|
| 38656 | 699  | 
qed auto  | 
700  | 
||
| 40859 | 701  | 
lemma borel_eq_halfspace_gt:  | 
702  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})\<rparr>)"
 | 
|
703  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 704  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 705  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
706  | 
using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto  | 
|
707  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
|
708  | 
by (rule borel.sets_sigma_subset) auto  | 
|
709  | 
qed auto  | 
|
| 38656 | 710  | 
|
| 40859 | 711  | 
lemma borel_eq_halfspace_ge:  | 
712  | 
  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})\<rparr>)"
 | 
|
713  | 
(is "_ = ?SIGMA")  | 
|
| 40869 | 714  | 
proof (intro algebra.equality antisym)  | 
| 40859 | 715  | 
show "sets borel \<subseteq> sets ?SIGMA"  | 
| 38656 | 716  | 
using halfspace_span_halfspace_ge open_span_halfspace by auto  | 
| 40859 | 717  | 
show "sets ?SIGMA \<subseteq> sets borel"  | 
718  | 
by (rule borel.sets_sigma_subset) auto  | 
|
719  | 
qed auto  | 
|
| 38656 | 720  | 
|
721  | 
lemma (in sigma_algebra) borel_measurable_halfspacesI:  | 
|
722  | 
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"  | 
|
| 40859 | 723  | 
assumes "borel = (sigma \<lparr>space=UNIV, sets=range F\<rparr>)"  | 
| 38656 | 724  | 
and "\<And>a i. S a i = f -` F (a,i) \<inter> space M"  | 
725  | 
  and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
 | 
|
726  | 
  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
 | 
|
727  | 
proof safe  | 
|
728  | 
  fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
 | 
|
729  | 
then show "S a i \<in> sets M" unfolding assms  | 
|
730  | 
by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def)  | 
|
731  | 
next  | 
|
732  | 
  assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
 | 
|
733  | 
  { fix a i have "S a i \<in> sets M"
 | 
|
734  | 
proof cases  | 
|
735  | 
      assume "i < DIM('c)"
 | 
|
736  | 
with a show ?thesis unfolding assms(2) by simp  | 
|
737  | 
next  | 
|
738  | 
      assume "\<not> i < DIM('c)"
 | 
|
739  | 
from assms(3)[OF this] show ?thesis .  | 
|
740  | 
qed }  | 
|
| 40859 | 741  | 
then have "f \<in> measurable M (sigma \<lparr>space=UNIV, sets=range F\<rparr>)"  | 
| 38656 | 742  | 
by (auto intro!: measurable_sigma simp: assms(2))  | 
743  | 
then show "f \<in> borel_measurable M" unfolding measurable_def  | 
|
744  | 
unfolding assms(1) by simp  | 
|
745  | 
qed  | 
|
746  | 
||
747  | 
lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:  | 
|
748  | 
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"  | 
|
749  | 
  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 | 750  | 
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto  | 
| 38656 | 751  | 
|
752  | 
lemma (in sigma_algebra) borel_measurable_iff_halfspace_less:  | 
|
753  | 
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"  | 
|
754  | 
  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 | 755  | 
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto  | 
| 38656 | 756  | 
|
757  | 
lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge:  | 
|
758  | 
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"  | 
|
759  | 
  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 | 760  | 
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto  | 
| 38656 | 761  | 
|
762  | 
lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater:  | 
|
763  | 
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"  | 
|
764  | 
  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 | 765  | 
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_gt]) auto  | 
| 38656 | 766  | 
|
767  | 
lemma (in sigma_algebra) borel_measurable_iff_le:  | 
|
768  | 
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
 | 
|
769  | 
using borel_measurable_iff_halfspace_le[where 'c=real] by simp  | 
|
770  | 
||
771  | 
lemma (in sigma_algebra) borel_measurable_iff_less:  | 
|
772  | 
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
 | 
|
773  | 
using borel_measurable_iff_halfspace_less[where 'c=real] by simp  | 
|
774  | 
||
775  | 
lemma (in sigma_algebra) borel_measurable_iff_ge:  | 
|
776  | 
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
 | 
|
777  | 
using borel_measurable_iff_halfspace_ge[where 'c=real] by simp  | 
|
778  | 
||
779  | 
lemma (in sigma_algebra) borel_measurable_iff_greater:  | 
|
780  | 
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
 | 
|
781  | 
using borel_measurable_iff_halfspace_greater[where 'c=real] by simp  | 
|
782  | 
||
| 41025 | 783  | 
lemma borel_measurable_euclidean_component:  | 
| 40859 | 784  | 
"(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel"  | 
785  | 
unfolding borel_def[where 'a=real]  | 
|
786  | 
proof (rule borel.measurable_sigma, simp_all)  | 
|
| 
39087
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
787  | 
fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def .  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
788  | 
from open_vimage_euclidean_component[OF this]  | 
| 40859 | 789  | 
show "(\<lambda>x. x $$ i) -` S \<in> sets borel"  | 
790  | 
by (auto intro: borel_open)  | 
|
791  | 
qed  | 
|
| 
39087
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
792  | 
|
| 41025 | 793  | 
lemma (in sigma_algebra) borel_measurable_euclidean_space:  | 
| 
39087
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
794  | 
fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
795  | 
  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: 
39083 
diff
changeset
 | 
796  | 
proof safe  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
797  | 
fix i assume "f \<in> borel_measurable M"  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
798  | 
then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M"  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
799  | 
using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def]  | 
| 41025 | 800  | 
by (auto intro: borel_measurable_euclidean_component)  | 
| 
39087
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
801  | 
next  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
802  | 
  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: 
39083 
diff
changeset
 | 
803  | 
then show "f \<in> borel_measurable M"  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
804  | 
unfolding borel_measurable_iff_halfspace_le by auto  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
805  | 
qed  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
806  | 
|
| 38656 | 807  | 
subsection "Borel measurable operators"  | 
808  | 
||
809  | 
lemma (in sigma_algebra) affine_borel_measurable_vector:  | 
|
810  | 
fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"  | 
|
811  | 
assumes "f \<in> borel_measurable M"  | 
|
812  | 
shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"  | 
|
813  | 
proof (rule borel_measurableI)  | 
|
814  | 
fix S :: "'x set" assume "open S"  | 
|
815  | 
show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"  | 
|
816  | 
proof cases  | 
|
817  | 
assume "b \<noteq> 0"  | 
|
818  | 
with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open")  | 
|
819  | 
by (auto intro!: open_affinity simp: scaleR.add_right mem_def)  | 
|
| 40859 | 820  | 
hence "?S \<in> sets borel"  | 
821  | 
unfolding borel_def by (auto simp: sigma_def intro!: sigma_sets.Basic)  | 
|
| 38656 | 822  | 
moreover  | 
823  | 
from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"  | 
|
824  | 
apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)  | 
|
| 40859 | 825  | 
ultimately show ?thesis using assms unfolding in_borel_measurable_borel  | 
| 38656 | 826  | 
by auto  | 
827  | 
qed simp  | 
|
828  | 
qed  | 
|
829  | 
||
830  | 
lemma (in sigma_algebra) affine_borel_measurable:  | 
|
831  | 
fixes g :: "'a \<Rightarrow> real"  | 
|
832  | 
assumes g: "g \<in> borel_measurable M"  | 
|
833  | 
shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"  | 
|
834  | 
using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)  | 
|
835  | 
||
836  | 
lemma (in sigma_algebra) borel_measurable_add[simp, intro]:  | 
|
837  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
838  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
839  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
840  | 
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
 | 
841  | 
proof -  | 
| 38656 | 842  | 
  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
 | 
843  | 
by auto  | 
| 38656 | 844  | 
have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"  | 
845  | 
by (rule affine_borel_measurable [OF g])  | 
|
846  | 
  then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
 | 
|
847  | 
by auto  | 
|
848  | 
  then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
 | 
|
849  | 
by (simp add: 1)  | 
|
850  | 
then show ?thesis  | 
|
851  | 
by (simp add: borel_measurable_iff_ge)  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
852  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
853  | 
|
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
854  | 
lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]:  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
855  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
856  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
857  | 
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
858  | 
proof cases  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
859  | 
assume "finite S"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
860  | 
thus ?thesis using assms by induct auto  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
861  | 
qed simp  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
862  | 
|
| 38656 | 863  | 
lemma (in sigma_algebra) borel_measurable_square:  | 
864  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
865  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
866  | 
shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
867  | 
proof -  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
868  | 
  {
 | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
869  | 
fix a  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
870  | 
    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
 | 
871  | 
proof (cases rule: linorder_cases [of a 0])  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
872  | 
case less  | 
| 38656 | 873  | 
      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
 | 
874  | 
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
 | 
875  | 
also have "... \<in> sets M"  | 
| 38656 | 876  | 
by (rule empty_sets)  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
877  | 
finally show ?thesis .  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
878  | 
next  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
879  | 
case equal  | 
| 38656 | 880  | 
      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
 | 
881  | 
             {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
 | 
882  | 
by auto  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
883  | 
also have "... \<in> sets M"  | 
| 38656 | 884  | 
apply (insert f)  | 
885  | 
apply (rule Int)  | 
|
886  | 
apply (simp add: borel_measurable_iff_le)  | 
|
887  | 
apply (simp add: borel_measurable_iff_ge)  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
888  | 
done  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
889  | 
finally show ?thesis .  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
890  | 
next  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
891  | 
case greater  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
892  | 
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
 | 
893  | 
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
 | 
894  | 
real_sqrt_le_iff real_sqrt_power)  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
895  | 
      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
 | 
| 38656 | 896  | 
             {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
 | 
897  | 
using greater by auto  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
898  | 
also have "... \<in> sets M"  | 
| 38656 | 899  | 
apply (insert f)  | 
900  | 
apply (rule Int)  | 
|
901  | 
apply (simp add: borel_measurable_iff_ge)  | 
|
902  | 
apply (simp add: borel_measurable_iff_le)  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
903  | 
done  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
904  | 
finally show ?thesis .  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
905  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
906  | 
}  | 
| 38656 | 907  | 
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
 | 
908  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
909  | 
|
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
910  | 
lemma times_eq_sum_squares:  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
911  | 
fixes x::real  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
912  | 
shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4"  | 
| 38656 | 913  | 
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
 | 
914  | 
|
| 38656 | 915  | 
lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]:  | 
916  | 
fixes g :: "'a \<Rightarrow> real"  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
917  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
918  | 
shows "(\<lambda>x. - g x) \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
919  | 
proof -  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
920  | 
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
 | 
921  | 
by simp  | 
| 38656 | 922  | 
also have "... \<in> borel_measurable M"  | 
923  | 
by (fast intro: affine_borel_measurable g)  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
924  | 
finally show ?thesis .  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
925  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
926  | 
|
| 38656 | 927  | 
lemma (in sigma_algebra) borel_measurable_times[simp, intro]:  | 
928  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
929  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
930  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
931  | 
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
 | 
932  | 
proof -  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
933  | 
have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M"  | 
| 38656 | 934  | 
using assms by (fast intro: affine_borel_measurable borel_measurable_square)  | 
935  | 
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
 | 
936  | 
(\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))"  | 
| 35582 | 937  | 
by (simp add: minus_divide_right)  | 
| 38656 | 938  | 
also have "... \<in> borel_measurable M"  | 
939  | 
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
 | 
940  | 
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
 | 
941  | 
show ?thesis  | 
| 38656 | 942  | 
apply (simp add: times_eq_sum_squares diff_minus)  | 
943  | 
using 1 2 by simp  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
944  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
945  | 
|
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
946  | 
lemma (in sigma_algebra) borel_measurable_setprod[simp, intro]:  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
947  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
948  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
949  | 
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
950  | 
proof cases  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
951  | 
assume "finite S"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
952  | 
thus ?thesis using assms by induct auto  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
953  | 
qed simp  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
954  | 
|
| 38656 | 955  | 
lemma (in sigma_algebra) borel_measurable_diff[simp, intro]:  | 
956  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
957  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
958  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
959  | 
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"  | 
| 38656 | 960  | 
unfolding diff_minus using assms by fast  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
961  | 
|
| 38656 | 962  | 
lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]:  | 
963  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 35692 | 964  | 
assumes "f \<in> borel_measurable M"  | 
965  | 
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"  | 
|
| 38656 | 966  | 
unfolding borel_measurable_iff_ge unfolding inverse_eq_divide  | 
967  | 
proof safe  | 
|
968  | 
fix a :: real  | 
|
969  | 
  have *: "{w \<in> space M. a \<le> 1 / f w} =
 | 
|
970  | 
      ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
 | 
|
971  | 
      ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
 | 
|
972  | 
      ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
 | 
|
973  | 
  show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
 | 
|
974  | 
by (auto intro!: Int Un)  | 
|
| 35692 | 975  | 
qed  | 
976  | 
||
| 38656 | 977  | 
lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:  | 
978  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
| 35692 | 979  | 
assumes "f \<in> borel_measurable M"  | 
980  | 
and "g \<in> borel_measurable M"  | 
|
981  | 
shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"  | 
|
982  | 
unfolding field_divide_inverse  | 
|
| 38656 | 983  | 
by (rule borel_measurable_inverse borel_measurable_times assms)+  | 
984  | 
||
985  | 
lemma (in sigma_algebra) borel_measurable_max[intro, simp]:  | 
|
986  | 
fixes f g :: "'a \<Rightarrow> real"  | 
|
987  | 
assumes "f \<in> borel_measurable M"  | 
|
988  | 
assumes "g \<in> borel_measurable M"  | 
|
989  | 
shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"  | 
|
990  | 
unfolding borel_measurable_iff_le  | 
|
991  | 
proof safe  | 
|
992  | 
fix a  | 
|
993  | 
  have "{x \<in> space M. max (g x) (f x) \<le> a} =
 | 
|
994  | 
    {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
 | 
|
995  | 
  thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
 | 
|
996  | 
using assms unfolding borel_measurable_iff_le  | 
|
997  | 
by (auto intro!: Int)  | 
|
998  | 
qed  | 
|
999  | 
||
1000  | 
lemma (in sigma_algebra) borel_measurable_min[intro, simp]:  | 
|
1001  | 
fixes f g :: "'a \<Rightarrow> real"  | 
|
1002  | 
assumes "f \<in> borel_measurable M"  | 
|
1003  | 
assumes "g \<in> borel_measurable M"  | 
|
1004  | 
shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"  | 
|
1005  | 
unfolding borel_measurable_iff_ge  | 
|
1006  | 
proof safe  | 
|
1007  | 
fix a  | 
|
1008  | 
  have "{x \<in> space M. a \<le> min (g x) (f x)} =
 | 
|
1009  | 
    {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
 | 
|
1010  | 
  thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
 | 
|
1011  | 
using assms unfolding borel_measurable_iff_ge  | 
|
1012  | 
by (auto intro!: Int)  | 
|
1013  | 
qed  | 
|
1014  | 
||
1015  | 
lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:  | 
|
1016  | 
assumes "f \<in> borel_measurable M"  | 
|
1017  | 
shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"  | 
|
1018  | 
proof -  | 
|
1019  | 
have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)  | 
|
1020  | 
show ?thesis unfolding * using assms by auto  | 
|
1021  | 
qed  | 
|
1022  | 
||
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1023  | 
lemma borel_measurable_nth[simp, intro]:  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1024  | 
"(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1025  | 
using borel_measurable_euclidean_component  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1026  | 
unfolding nth_conv_component by auto  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1027  | 
|
| 41830 | 1028  | 
lemma borel_measurable_continuous_on1:  | 
1029  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"  | 
|
1030  | 
assumes "continuous_on UNIV f"  | 
|
1031  | 
shows "f \<in> borel_measurable borel"  | 
|
1032  | 
apply(rule borel.borel_measurableI)  | 
|
1033  | 
using continuous_open_preimage[OF assms] unfolding vimage_def by auto  | 
|
1034  | 
||
1035  | 
lemma borel_measurable_continuous_on:  | 
|
1036  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"  | 
|
| 
42990
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1037  | 
assumes cont: "continuous_on A f" "open A"  | 
| 41830 | 1038  | 
shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _")  | 
1039  | 
proof (rule borel.borel_measurableI)  | 
|
1040  | 
fix S :: "'b set" assume "open S"  | 
|
| 
42990
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1041  | 
  then have "open {x\<in>A. f x \<in> S}"
 | 
| 41830 | 1042  | 
by (intro continuous_open_preimage[OF cont]) auto  | 
| 
42990
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1043  | 
  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: 
42950 
diff
changeset
 | 
1044  | 
have "?f -` S \<inter> space borel =  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1045  | 
    {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: 
42950 
diff
changeset
 | 
1046  | 
by (auto split: split_if_asm)  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1047  | 
also have "\<dots> \<in> sets borel"  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1048  | 
using * `open A` by (auto simp del: space_borel intro!: borel.Un)  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1049  | 
finally show "?f -` S \<inter> space borel \<in> sets borel" .  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1050  | 
qed  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1051  | 
|
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1052  | 
lemma (in sigma_algebra) convex_measurable:  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1053  | 
fixes a b :: real  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1054  | 
  assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> { a <..< b}"
 | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1055  | 
  assumes q: "convex_on { a <..< b} q"
 | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1056  | 
shows "q \<circ> X \<in> borel_measurable M"  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1057  | 
proof -  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1058  | 
  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: 
42950 
diff
changeset
 | 
1059  | 
proof (rule borel_measurable_continuous_on)  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1060  | 
    show "open {a<..<b}" by auto
 | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1061  | 
    from this q show "continuous_on {a<..<b} q"
 | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1062  | 
by (rule convex_on_continuous)  | 
| 41830 | 1063  | 
qed  | 
| 
42990
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1064  | 
  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: 
42950 
diff
changeset
 | 
1065  | 
using X by (intro measurable_comp) auto  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1066  | 
moreover have "?qX \<longleftrightarrow> q \<circ> X \<in> borel_measurable M"  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1067  | 
using X by (intro measurable_cong) auto  | 
| 
 
3706951a6421
composition of convex and measurable function is measurable
 
hoelzl 
parents: 
42950 
diff
changeset
 | 
1068  | 
ultimately show ?thesis by simp  | 
| 41830 | 1069  | 
qed  | 
1070  | 
||
1071  | 
lemma borel_measurable_borel_log: assumes "1 < b" shows "log b \<in> borel_measurable borel"  | 
|
1072  | 
proof -  | 
|
1073  | 
  { fix x :: real assume x: "x \<le> 0"
 | 
|
1074  | 
    { fix x::real assume "x \<le> 0" then have "\<And>u. exp u = x \<longleftrightarrow> False" by auto }
 | 
|
1075  | 
from this[of x] x this[of 0] have "log b 0 = log b x"  | 
|
1076  | 
by (auto simp: ln_def log_def) }  | 
|
1077  | 
note log_imp = this  | 
|
1078  | 
  have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) \<in> borel_measurable borel"
 | 
|
1079  | 
proof (rule borel_measurable_continuous_on)  | 
|
1080  | 
    show "continuous_on {0<..} (log b)"
 | 
|
1081  | 
by (auto intro!: continuous_at_imp_continuous_on DERIV_log DERIV_isCont  | 
|
1082  | 
simp: continuous_isCont[symmetric])  | 
|
1083  | 
    show "open ({0<..}::real set)" by auto
 | 
|
1084  | 
qed  | 
|
1085  | 
  also have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) = log b"
 | 
|
1086  | 
by (simp add: fun_eq_iff not_less log_imp)  | 
|
1087  | 
finally show ?thesis .  | 
|
1088  | 
qed  | 
|
1089  | 
||
1090  | 
lemma (in sigma_algebra) borel_measurable_log[simp,intro]:  | 
|
1091  | 
assumes f: "f \<in> borel_measurable M" and "1 < b"  | 
|
1092  | 
shows "(\<lambda>x. log b (f x)) \<in> borel_measurable M"  | 
|
1093  | 
using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]]  | 
|
1094  | 
by (simp add: comp_def)  | 
|
1095  | 
||
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1096  | 
subsection "Borel space on the extended reals"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1097  | 
|
| 43920 | 1098  | 
lemma borel_measurable_ereal_borel:  | 
1099  | 
"ereal \<in> borel_measurable borel"  | 
|
1100  | 
unfolding borel_def[where 'a=ereal]  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1101  | 
proof (rule borel.measurable_sigma)  | 
| 43920 | 1102  | 
fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1103  | 
then have "open X" by (auto simp: mem_def)  | 
| 43920 | 1104  | 
then have "open (ereal -` X \<inter> space borel)"  | 
1105  | 
by (simp add: open_ereal_vimage)  | 
|
1106  | 
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: 
41969 
diff
changeset
 | 
1107  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1108  | 
|
| 43920 | 1109  | 
lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]:  | 
1110  | 
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"  | 
|
1111  | 
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: 
41969 
diff
changeset
 | 
1112  | 
|
| 43920 | 1113  | 
lemma borel_measurable_real_of_ereal_borel:  | 
1114  | 
"(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1115  | 
unfolding borel_def[where 'a=real]  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1116  | 
proof (rule borel.measurable_sigma)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1117  | 
fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1118  | 
then have "open B" by (auto simp: mem_def)  | 
| 43920 | 1119  | 
  have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
 | 
1120  | 
  have open_real: "open (real -` (B - {0}) :: ereal set)"
 | 
|
1121  | 
unfolding open_ereal_def * using `open B` by auto  | 
|
1122  | 
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: 
41969 
diff
changeset
 | 
1123  | 
proof cases  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1124  | 
assume "0 \<in> B"  | 
| 43923 | 1125  | 
    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0::ereal}"
 | 
| 43920 | 1126  | 
by (auto simp add: real_of_ereal_eq_0)  | 
1127  | 
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: 
41969 
diff
changeset
 | 
1128  | 
using open_real by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1129  | 
next  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1130  | 
assume "0 \<notin> B"  | 
| 43920 | 1131  | 
    then have *: "(real -` B :: ereal set) = real -` (B - {0})"
 | 
1132  | 
by (auto simp add: real_of_ereal_eq_0)  | 
|
1133  | 
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: 
41969 
diff
changeset
 | 
1134  | 
using open_real by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1135  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1136  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1137  | 
|
| 43920 | 1138  | 
lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]:  | 
1139  | 
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M"  | 
|
1140  | 
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: 
41969 
diff
changeset
 | 
1141  | 
|
| 43920 | 1142  | 
lemma (in sigma_algebra) borel_measurable_ereal_iff:  | 
1143  | 
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: 
41969 
diff
changeset
 | 
1144  | 
proof  | 
| 43920 | 1145  | 
assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"  | 
1146  | 
from borel_measurable_real_of_ereal[OF this]  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1147  | 
show "f \<in> borel_measurable M" by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1148  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1149  | 
|
| 43920 | 1150  | 
lemma (in sigma_algebra) borel_measurable_ereal_iff_real:  | 
| 43923 | 1151  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
1152  | 
shows "f \<in> borel_measurable M \<longleftrightarrow>  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1153  | 
    ((\<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: 
41969 
diff
changeset
 | 
1154  | 
proof safe  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1155  | 
  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: 
41969 
diff
changeset
 | 
1156  | 
  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: 
41969 
diff
changeset
 | 
1157  | 
  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
 | 
| 43920 | 1158  | 
let "?f 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: 
41969 
diff
changeset
 | 
1159  | 
have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto  | 
| 43920 | 1160  | 
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: 
41969 
diff
changeset
 | 
1161  | 
finally show "f \<in> borel_measurable M" .  | 
| 43920 | 1162  | 
qed (auto intro: measurable_sets borel_measurable_real_of_ereal)  | 
| 41830 | 1163  | 
|
| 38656 | 1164  | 
lemma (in sigma_algebra) less_eq_ge_measurable:  | 
1165  | 
fixes f :: "'a \<Rightarrow> 'c::linorder"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1166  | 
  shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
 | 
| 38656 | 1167  | 
proof  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1168  | 
  assume "f -` {a <..} \<inter> space M \<in> sets M"
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1169  | 
  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: 
41969 
diff
changeset
 | 
1170  | 
  ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1171  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
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: 
41969 
diff
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: 
41969 
diff
changeset
 | 
1174  | 
  ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1175  | 
qed  | 
| 35692 | 1176  | 
|
| 38656 | 1177  | 
lemma (in sigma_algebra) greater_eq_le_measurable:  | 
1178  | 
fixes f :: "'a \<Rightarrow> 'c::linorder"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1179  | 
  shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
 | 
| 38656 | 1180  | 
proof  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1181  | 
  assume "f -` {a ..} \<inter> space M \<in> sets M"
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1182  | 
  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: 
41969 
diff
changeset
 | 
1183  | 
  ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1184  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
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: 
41969 
diff
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: 
41969 
diff
changeset
 | 
1187  | 
  ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
 | 
| 38656 | 1188  | 
qed  | 
1189  | 
||
| 43920 | 1190  | 
lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal:  | 
1191  | 
"(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1192  | 
proof (subst borel_def, rule borel.measurable_sigma)  | 
| 43920 | 1193  | 
fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1194  | 
then have "open X" by (simp add: mem_def)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1195  | 
have "uminus -` X = uminus ` X" by (force simp: image_iff)  | 
| 43920 | 1196  | 
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: 
41969 
diff
changeset
 | 
1197  | 
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: 
41969 
diff
changeset
 | 
1198  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1199  | 
|
| 43920 | 1200  | 
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: 
41969 
diff
changeset
 | 
1201  | 
assumes "f \<in> borel_measurable M"  | 
| 43920 | 1202  | 
shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"  | 
1203  | 
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: 
41969 
diff
changeset
 | 
1204  | 
|
| 43920 | 1205  | 
lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]:  | 
1206  | 
"(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")  | 
|
| 38656 | 1207  | 
proof  | 
| 43920 | 1208  | 
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: 
41969 
diff
changeset
 | 
1209  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1210  | 
|
| 43920 | 1211  | 
lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal:  | 
| 43923 | 1212  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
1213  | 
  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: 
41969 
diff
changeset
 | 
1214  | 
proof (intro iffI allI)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1215  | 
  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: 
41969 
diff
changeset
 | 
1216  | 
show "f \<in> borel_measurable M"  | 
| 43920 | 1217  | 
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: 
41969 
diff
changeset
 | 
1218  | 
proof (intro conjI allI)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1219  | 
fix a :: real  | 
| 43920 | 1220  | 
    { 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: 
41969 
diff
changeset
 | 
1221  | 
have "x = \<infinity>"  | 
| 43920 | 1222  | 
proof (rule ereal_top)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1223  | 
fix B from real_arch_lt[of B] guess n ..  | 
| 43920 | 1224  | 
then have "ereal B < real n" by auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1225  | 
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: 
41969 
diff
changeset
 | 
1226  | 
qed }  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1227  | 
    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: 
41969 
diff
changeset
 | 
1228  | 
by (auto simp: not_le)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1229  | 
    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: 
41969 
diff
changeset
 | 
1230  | 
moreover  | 
| 43923 | 1231  | 
    have "{-\<infinity>::ereal} = {..-\<infinity>}" by auto
 | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1232  | 
    then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
 | 
| 43920 | 1233  | 
    moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
 | 
1234  | 
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: 
41969 
diff
changeset
 | 
1235  | 
    moreover have "{w \<in> space M. real (f w) \<le> a} =
 | 
| 43920 | 1236  | 
      (if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
 | 
1237  | 
      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: 
41969 
diff
changeset
 | 
1238  | 
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: 
41969 
diff
changeset
 | 
1239  | 
    ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
 | 
| 35582 | 1240  | 
qed  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1241  | 
qed (simp add: measurable_sets)  | 
| 35582 | 1242  | 
|
| 43920 | 1243  | 
lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal:  | 
1244  | 
  "(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: 
41969 
diff
changeset
 | 
1245  | 
proof  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1246  | 
  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: 
41969 
diff
changeset
 | 
1247  | 
  moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
 | 
| 43920 | 1248  | 
by (auto simp: ereal_uminus_le_reorder)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1249  | 
ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"  | 
| 43920 | 1250  | 
unfolding borel_measurable_eq_atMost_ereal by auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1251  | 
then show "f \<in> borel_measurable M" by simp  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1252  | 
qed (simp add: measurable_sets)  | 
| 35582 | 1253  | 
|
| 43920 | 1254  | 
lemma (in sigma_algebra) borel_measurable_ereal_iff_less:  | 
1255  | 
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
 | 
|
1256  | 
unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..  | 
|
| 38656 | 1257  | 
|
| 43920 | 1258  | 
lemma (in sigma_algebra) borel_measurable_ereal_iff_ge:  | 
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_atMost_ereal less_eq_ge_measurable ..  | 
|
| 38656 | 1261  | 
|
| 43920 | 1262  | 
lemma (in sigma_algebra) borel_measurable_ereal_eq_const:  | 
1263  | 
fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"  | 
|
| 38656 | 1264  | 
  shows "{x\<in>space M. f x = c} \<in> sets M"
 | 
1265  | 
proof -  | 
|
1266  | 
  have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
 | 
|
1267  | 
then show ?thesis using assms by (auto intro!: measurable_sets)  | 
|
1268  | 
qed  | 
|
1269  | 
||
| 43920 | 1270  | 
lemma (in sigma_algebra) borel_measurable_ereal_neq_const:  | 
1271  | 
fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"  | 
|
| 38656 | 1272  | 
  shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 | 
1273  | 
proof -  | 
|
1274  | 
  have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
 | 
|
1275  | 
then show ?thesis using assms by (auto intro!: measurable_sets)  | 
|
1276  | 
qed  | 
|
1277  | 
||
| 43920 | 1278  | 
lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]:  | 
1279  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1280  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1281  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1282  | 
  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: 
41969 
diff
changeset
 | 
1283  | 
proof -  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1284  | 
  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: 
41969 
diff
changeset
 | 
1285  | 
    {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: 
41969 
diff
changeset
 | 
1286  | 
    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: 
41969 
diff
changeset
 | 
1287  | 
proof (intro set_eqI)  | 
| 43920 | 1288  | 
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: 
41969 
diff
changeset
 | 
1289  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1290  | 
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: 
41969 
diff
changeset
 | 
1291  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1292  | 
|
| 43920 | 1293  | 
lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]:  | 
1294  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1295  | 
assumes f: "f \<in> borel_measurable M"  | 
1296  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1297  | 
  shows "{x \<in> space M. f x < g x} \<in> sets M"
 | 
|
1298  | 
proof -  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1299  | 
  have "{x \<in> space M. f x < g x} = space M - {x \<in> space M. g x \<le> f x}" by auto
 | 
| 38656 | 1300  | 
then show ?thesis using g f by auto  | 
1301  | 
qed  | 
|
1302  | 
||
| 43920 | 1303  | 
lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]:  | 
1304  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1305  | 
assumes f: "f \<in> borel_measurable M"  | 
1306  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1307  | 
  shows "{w \<in> space M. f w = g w} \<in> sets M"
 | 
|
1308  | 
proof -  | 
|
1309  | 
  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
 | 
|
1310  | 
then show ?thesis using g f by auto  | 
|
1311  | 
qed  | 
|
1312  | 
||
| 43920 | 1313  | 
lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]:  | 
1314  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1315  | 
assumes f: "f \<in> borel_measurable M"  | 
1316  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1317  | 
  shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | 
|
| 35692 | 1318  | 
proof -  | 
| 38656 | 1319  | 
  have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
 | 
1320  | 
thus ?thesis using f g by auto  | 
|
1321  | 
qed  | 
|
1322  | 
||
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1323  | 
lemma (in sigma_algebra) split_sets:  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1324  | 
  "{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: 
41969 
diff
changeset
 | 
1325  | 
  "{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: 
41969 
diff
changeset
 | 
1326  | 
by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1327  | 
|
| 43920 | 1328  | 
lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]:  | 
1329  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
|
| 41025 | 1330  | 
assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 38656 | 1331  | 
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"  | 
1332  | 
proof -  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1333  | 
  { 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: 
41969 
diff
changeset
 | 
1334  | 
(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: 
41969 
diff
changeset
 | 
1335  | 
else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>  | 
| 43920 | 1336  | 
else ereal (real (f x) + real (g x)))"  | 
1337  | 
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: 
41969 
diff
changeset
 | 
1338  | 
with assms show ?thesis  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1339  | 
by (auto cong: measurable_cong simp: split_sets  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1340  | 
intro!: Un measurable_If measurable_sets)  | 
| 38656 | 1341  | 
qed  | 
1342  | 
||
| 43920 | 1343  | 
lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]:  | 
1344  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"  | 
|
| 41096 | 1345  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
1346  | 
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"  | 
|
1347  | 
proof cases  | 
|
1348  | 
assume "finite S"  | 
|
1349  | 
thus ?thesis using assms  | 
|
1350  | 
by induct auto  | 
|
1351  | 
qed (simp add: borel_measurable_const)  | 
|
1352  | 
||
| 43920 | 1353  | 
lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]:  | 
1354  | 
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: 
41969 
diff
changeset
 | 
1355  | 
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: 
41969 
diff
changeset
 | 
1356  | 
proof -  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1357  | 
  { 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: 
41969 
diff
changeset
 | 
1358  | 
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: 
41969 
diff
changeset
 | 
1359  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1360  | 
|
| 43920 | 1361  | 
lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]:  | 
1362  | 
fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
|
| 38656 | 1363  | 
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"  | 
1364  | 
proof -  | 
|
| 43920 | 1365  | 
  { fix f g :: "'a \<Rightarrow> ereal"
 | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1366  | 
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: 
41969 
diff
changeset
 | 
1367  | 
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: 
41969 
diff
changeset
 | 
1368  | 
    { 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: 
41969 
diff
changeset
 | 
1369  | 
else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>  | 
| 43920 | 1370  | 
else ereal (real (f x) * real (g x)))"  | 
1371  | 
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: 
41969 
diff
changeset
 | 
1372  | 
using pos[of x] by auto }  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1373  | 
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: 
41969 
diff
changeset
 | 
1374  | 
by (auto cong: measurable_cong simp: split_sets  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1375  | 
intro!: Un measurable_If measurable_sets) }  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1376  | 
note pos_times = this  | 
| 38656 | 1377  | 
have *: "(\<lambda>x. f x * g x) =  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1378  | 
(\<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: 
41969 
diff
changeset
 | 
1379  | 
by (auto simp: fun_eq_iff)  | 
| 38656 | 1380  | 
show ?thesis using assms unfolding *  | 
| 43920 | 1381  | 
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: 
41969 
diff
changeset
 | 
1382  | 
(auto simp: split_sets intro!: Int)  | 
| 38656 | 1383  | 
qed  | 
1384  | 
||
| 43920 | 1385  | 
lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]:  | 
1386  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"  | 
|
| 38656 | 1387  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
| 41096 | 1388  | 
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 38656 | 1389  | 
proof cases  | 
1390  | 
assume "finite S"  | 
|
| 41096 | 1391  | 
thus ?thesis using assms by induct auto  | 
1392  | 
qed simp  | 
|
| 38656 | 1393  | 
|
| 43920 | 1394  | 
lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]:  | 
1395  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1396  | 
assumes "f \<in> borel_measurable M"  | 
1397  | 
assumes "g \<in> borel_measurable M"  | 
|
1398  | 
shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"  | 
|
1399  | 
using assms unfolding min_def by (auto intro!: measurable_If)  | 
|
1400  | 
||
| 43920 | 1401  | 
lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]:  | 
1402  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1403  | 
assumes "f \<in> borel_measurable M"  | 
1404  | 
and "g \<in> borel_measurable M"  | 
|
1405  | 
shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"  | 
|
1406  | 
using assms unfolding max_def by (auto intro!: measurable_If)  | 
|
1407  | 
||
1408  | 
lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:  | 
|
| 43920 | 1409  | 
fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 38656 | 1410  | 
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: 
41096 
diff
changeset
 | 
1411  | 
shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")  | 
| 43920 | 1412  | 
unfolding borel_measurable_ereal_iff_ge  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1413  | 
proof  | 
| 38656 | 1414  | 
fix a  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1415  | 
  have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
 | 
| 
41083
 
c987429a1298
work around problems with eta-expansion of equations
 
haftmann 
parents: 
41080 
diff
changeset
 | 
1416  | 
by (auto simp: less_SUP_iff SUPR_apply)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1417  | 
  then show "?sup -` {a<..} \<inter> space M \<in> sets M"
 | 
| 38656 | 1418  | 
using assms by auto  | 
1419  | 
qed  | 
|
1420  | 
||
1421  | 
lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:  | 
|
| 43920 | 1422  | 
fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 38656 | 1423  | 
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: 
41096 
diff
changeset
 | 
1424  | 
shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")  | 
| 43920 | 1425  | 
unfolding borel_measurable_ereal_iff_less  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1426  | 
proof  | 
| 38656 | 1427  | 
fix a  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1428  | 
  have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
 | 
| 
41083
 
c987429a1298
work around problems with eta-expansion of equations
 
haftmann 
parents: 
41080 
diff
changeset
 | 
1429  | 
by (auto simp: INF_less_iff INFI_apply)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1430  | 
  then show "?inf -` {..<a} \<inter> space M \<in> sets M"
 | 
| 38656 | 1431  | 
using assms by auto  | 
1432  | 
qed  | 
|
1433  | 
||
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1434  | 
lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:  | 
| 43920 | 1435  | 
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1436  | 
assumes "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1437  | 
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: 
41969 
diff
changeset
 | 
1438  | 
unfolding liminf_SUPR_INFI using assms by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1439  | 
|
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1440  | 
lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:  | 
| 43920 | 1441  | 
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1442  | 
assumes "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1443  | 
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: 
41969 
diff
changeset
 | 
1444  | 
unfolding limsup_INFI_SUPR using assms by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1445  | 
|
| 43920 | 1446  | 
lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]:  | 
1447  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
|
| 38656 | 1448  | 
assumes "f \<in> borel_measurable M"  | 
1449  | 
assumes "g \<in> borel_measurable M"  | 
|
1450  | 
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"  | 
|
| 43920 | 1451  | 
unfolding minus_ereal_def using assms by auto  | 
| 35692 | 1452  | 
|
| 
40870
 
94427db32392
Tuned setup for borel_measurable with min, max and psuminf.
 
hoelzl 
parents: 
40869 
diff
changeset
 | 
1453  | 
lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:  | 
| 43920 | 1454  | 
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1455  | 
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: 
41969 
diff
changeset
 | 
1456  | 
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: 
41969 
diff
changeset
 | 
1457  | 
apply (subst measurable_cong)  | 
| 43920 | 1458  | 
apply (subst suminf_ereal_eq_SUPR)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1459  | 
apply (rule pos)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1460  | 
using assms by auto  | 
| 39092 | 1461  | 
|
1462  | 
section "LIMSEQ is borel measurable"  | 
|
1463  | 
||
1464  | 
lemma (in sigma_algebra) borel_measurable_LIMSEQ:  | 
|
1465  | 
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"  | 
|
1466  | 
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"  | 
|
1467  | 
and u: "\<And>i. u i \<in> borel_measurable M"  | 
|
1468  | 
shows "u' \<in> borel_measurable M"  | 
|
1469  | 
proof -  | 
|
| 43920 | 1470  | 
have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"  | 
1471  | 
using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_ereal)  | 
|
1472  | 
moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"  | 
|
| 39092 | 1473  | 
by auto  | 
| 43920 | 1474  | 
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)  | 
| 39092 | 1475  | 
qed  | 
1476  | 
||
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
1477  | 
end  |