| author | wenzelm |
| Tue, 20 Nov 2012 22:52:04 +0100 | |
| changeset 50136 | a96bd08258a2 |
| parent 50104 | de19856feb54 |
| child 50244 | de72bbe42190 |
| permissions | -rw-r--r-- |
| 42150 | 1 |
(* Title: HOL/Probability/Borel_Space.thy |
| 42067 | 2 |
Author: Johannes Hölzl, TU München |
3 |
Author: Armin Heller, TU München |
|
4 |
*) |
|
| 38656 | 5 |
|
6 |
header {*Borel spaces*}
|
|
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
7 |
|
| 40859 | 8 |
theory Borel_Space |
| 45288 | 9 |
imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
10 |
begin |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
11 |
|
| 38656 | 12 |
section "Generic Borel spaces" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
13 |
|
| 47694 | 14 |
definition borel :: "'a::topological_space measure" where |
15 |
"borel = sigma UNIV {S. open S}"
|
|
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
16 |
|
| 47694 | 17 |
abbreviation "borel_measurable M \<equiv> measurable M borel" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
18 |
|
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
19 |
lemma in_borel_measurable: |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
20 |
"f \<in> borel_measurable M \<longleftrightarrow> |
| 47694 | 21 |
(\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
|
| 40859 | 22 |
by (auto simp add: measurable_def borel_def) |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
23 |
|
| 40859 | 24 |
lemma in_borel_measurable_borel: |
| 38656 | 25 |
"f \<in> borel_measurable M \<longleftrightarrow> |
| 40859 | 26 |
(\<forall>S \<in> sets borel. |
| 38656 | 27 |
f -` S \<inter> space M \<in> sets M)" |
| 40859 | 28 |
by (auto simp add: measurable_def borel_def) |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
29 |
|
| 40859 | 30 |
lemma space_borel[simp]: "space borel = UNIV" |
31 |
unfolding borel_def by auto |
|
| 38656 | 32 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
33 |
lemma space_in_borel[measurable]: "UNIV \<in> sets borel" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
34 |
unfolding borel_def by auto |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
35 |
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
36 |
lemma pred_Collect_borel[measurable (raw)]: "Sigma_Algebra.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
37 |
unfolding borel_def pred_def by auto |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
38 |
|
| 50003 | 39 |
lemma borel_open[measurable (raw generic)]: |
| 40859 | 40 |
assumes "open A" shows "A \<in> sets borel" |
| 38656 | 41 |
proof - |
|
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
42 |
have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
|
| 47694 | 43 |
thus ?thesis unfolding borel_def by auto |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
44 |
qed |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
45 |
|
| 50003 | 46 |
lemma borel_closed[measurable (raw generic)]: |
| 40859 | 47 |
assumes "closed A" shows "A \<in> sets borel" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
48 |
proof - |
| 40859 | 49 |
have "space borel - (- A) \<in> sets borel" |
50 |
using assms unfolding closed_def by (blast intro: borel_open) |
|
| 38656 | 51 |
thus ?thesis by simp |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
52 |
qed |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
53 |
|
| 50003 | 54 |
lemma borel_singleton[measurable]: |
55 |
"A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)" |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
56 |
unfolding insert_def by (rule Un) auto |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
57 |
|
| 50003 | 58 |
lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
59 |
unfolding Compl_eq_Diff_UNIV by simp |
| 41830 | 60 |
|
| 47694 | 61 |
lemma borel_measurable_vimage: |
| 38656 | 62 |
fixes f :: "'a \<Rightarrow> 'x::t2_space" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
63 |
assumes borel[measurable]: "f \<in> borel_measurable M" |
| 38656 | 64 |
shows "f -` {x} \<inter> space M \<in> sets M"
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
65 |
by simp |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
66 |
|
| 47694 | 67 |
lemma borel_measurableI: |
| 38656 | 68 |
fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space" |
69 |
assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" |
|
70 |
shows "f \<in> borel_measurable M" |
|
| 40859 | 71 |
unfolding borel_def |
| 47694 | 72 |
proof (rule measurable_measure_of, simp_all) |
|
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
73 |
fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M" |
|
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
74 |
using assms[of S] by simp |
| 40859 | 75 |
qed |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
76 |
|
|
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
77 |
lemma borel_measurable_const: |
| 38656 | 78 |
"(\<lambda>x. c) \<in> borel_measurable M" |
| 47694 | 79 |
by auto |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
80 |
|
| 50003 | 81 |
lemma borel_measurable_indicator: |
| 38656 | 82 |
assumes A: "A \<in> sets M" |
83 |
shows "indicator A \<in> borel_measurable M" |
|
| 46905 | 84 |
unfolding indicator_def [abs_def] using A |
| 47694 | 85 |
by (auto intro!: measurable_If_set) |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
86 |
|
| 50096 | 87 |
lemma borel_measurable_count_space[measurable (raw)]: |
88 |
"f \<in> borel_measurable (count_space S)" |
|
89 |
unfolding measurable_def by auto |
|
90 |
||
91 |
lemma borel_measurable_indicator'[measurable (raw)]: |
|
92 |
assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
|
|
93 |
shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M" |
|
|
50001
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49774
diff
changeset
|
94 |
unfolding indicator_def[abs_def] |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49774
diff
changeset
|
95 |
by (auto intro!: measurable_If) |
|
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49774
diff
changeset
|
96 |
|
| 47694 | 97 |
lemma borel_measurable_indicator_iff: |
| 40859 | 98 |
"(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
|
99 |
(is "?I \<in> borel_measurable M \<longleftrightarrow> _") |
|
100 |
proof |
|
101 |
assume "?I \<in> borel_measurable M" |
|
102 |
then have "?I -` {1} \<inter> space M \<in> sets M"
|
|
103 |
unfolding measurable_def by auto |
|
104 |
also have "?I -` {1} \<inter> space M = A \<inter> space M"
|
|
| 46905 | 105 |
unfolding indicator_def [abs_def] by auto |
| 40859 | 106 |
finally show "A \<inter> space M \<in> sets M" . |
107 |
next |
|
108 |
assume "A \<inter> space M \<in> sets M" |
|
109 |
moreover have "?I \<in> borel_measurable M \<longleftrightarrow> |
|
110 |
(indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M" |
|
111 |
by (intro measurable_cong) (auto simp: indicator_def) |
|
112 |
ultimately show "?I \<in> borel_measurable M" by auto |
|
113 |
qed |
|
114 |
||
| 47694 | 115 |
lemma borel_measurable_subalgebra: |
| 41545 | 116 |
assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N" |
| 39092 | 117 |
shows "f \<in> borel_measurable M" |
118 |
using assms unfolding measurable_def by auto |
|
119 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
120 |
lemma borel_measurable_continuous_on1: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
121 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
122 |
assumes "continuous_on UNIV f" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
123 |
shows "f \<in> borel_measurable borel" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
124 |
apply(rule borel_measurableI) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
125 |
using continuous_open_preimage[OF assms] unfolding vimage_def by auto |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
126 |
|
| 38656 | 127 |
section "Borel spaces on euclidean spaces" |
128 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
129 |
lemma borel_measurable_euclidean_component'[measurable]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
130 |
"(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
131 |
by (intro continuous_on_euclidean_component continuous_on_id borel_measurable_continuous_on1) |
| 38656 | 132 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
133 |
lemma borel_measurable_euclidean_component: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
134 |
"(f :: 'a \<Rightarrow> 'b::euclidean_space) \<in> borel_measurable M \<Longrightarrow>(\<lambda>x. f x $$ i) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
135 |
by simp |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
136 |
|
| 50003 | 137 |
lemma [measurable]: |
| 38656 | 138 |
fixes a b :: "'a\<Colon>ordered_euclidean_space" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
139 |
shows lessThan_borel: "{..< a} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
140 |
and greaterThan_borel: "{a <..} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
141 |
and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
142 |
and atMost_borel: "{..a} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
143 |
and atLeast_borel: "{a..} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
144 |
and atLeastAtMost_borel: "{a..b} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
145 |
and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
146 |
and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
147 |
unfolding greaterThanAtMost_def atLeastLessThan_def |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
148 |
by (blast intro: borel_open borel_closed)+ |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
149 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
150 |
lemma |
| 50003 | 151 |
shows hafspace_less_borel: "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel"
|
152 |
and hafspace_greater_borel: "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel"
|
|
153 |
and hafspace_less_eq_borel: "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel"
|
|
154 |
and hafspace_greater_eq_borel: "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel"
|
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
155 |
by simp_all |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
156 |
|
| 50003 | 157 |
lemma borel_measurable_less[measurable]: |
| 38656 | 158 |
fixes f :: "'a \<Rightarrow> real" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
159 |
assumes f: "f \<in> borel_measurable M" |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
160 |
assumes g: "g \<in> borel_measurable M" |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
161 |
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
|
162 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
163 |
have "{w \<in> space M. f w < g w} = {x \<in> space M. \<exists>r. f x < of_rat r \<and> of_rat r < g x}"
|
| 38656 | 164 |
using Rats_dense_in_real by (auto simp add: Rats_def) |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
165 |
with f g show ?thesis |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
166 |
by simp |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
167 |
qed |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
168 |
|
| 50003 | 169 |
lemma |
| 38656 | 170 |
fixes f :: "'a \<Rightarrow> real" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
171 |
assumes f[measurable]: "f \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
172 |
assumes g[measurable]: "g \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
173 |
shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
174 |
and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
175 |
and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
|
|
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
176 |
unfolding eq_iff not_less[symmetric] |
|
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
177 |
by measurable |
| 38656 | 178 |
|
179 |
subsection "Borel space equals sigma algebras over intervals" |
|
180 |
||
| 47694 | 181 |
lemma borel_sigma_sets_subset: |
182 |
"A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel" |
|
183 |
using sigma_sets_subset[of A borel] by simp |
|
184 |
||
185 |
lemma borel_eq_sigmaI1: |
|
186 |
fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" |
|
187 |
assumes borel_eq: "borel = sigma UNIV X" |
|
188 |
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range F))" |
|
189 |
assumes F: "\<And>i. F i \<in> sets borel" |
|
190 |
shows "borel = sigma UNIV (range F)" |
|
191 |
unfolding borel_def |
|
192 |
proof (intro sigma_eqI antisym) |
|
193 |
have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
|
|
194 |
unfolding borel_def by simp |
|
195 |
also have "\<dots> = sigma_sets UNIV X" |
|
196 |
unfolding borel_eq by simp |
|
197 |
also have "\<dots> \<subseteq> sigma_sets UNIV (range F)" |
|
198 |
using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto |
|
199 |
finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (range F)" .
|
|
200 |
show "sigma_sets UNIV (range F) \<subseteq> sigma_sets UNIV {S. open S}"
|
|
201 |
unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto |
|
202 |
qed auto |
|
| 38656 | 203 |
|
| 47694 | 204 |
lemma borel_eq_sigmaI2: |
205 |
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" |
|
206 |
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" |
|
207 |
assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))" |
|
208 |
assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))" |
|
209 |
assumes F: "\<And>i j. F i j \<in> sets borel" |
|
210 |
shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))" |
|
211 |
using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F="(\<lambda>(i, j). F i j)"]) auto |
|
212 |
||
213 |
lemma borel_eq_sigmaI3: |
|
214 |
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" |
|
215 |
assumes borel_eq: "borel = sigma UNIV X" |
|
216 |
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))" |
|
217 |
assumes F: "\<And>i j. F i j \<in> sets borel" |
|
218 |
shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))" |
|
219 |
using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto |
|
220 |
||
221 |
lemma borel_eq_sigmaI4: |
|
222 |
fixes F :: "'i \<Rightarrow> 'a::topological_space set" |
|
223 |
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" |
|
224 |
assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))" |
|
225 |
assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range F))" |
|
226 |
assumes F: "\<And>i. F i \<in> sets borel" |
|
227 |
shows "borel = sigma UNIV (range F)" |
|
228 |
using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F=F]) auto |
|
229 |
||
230 |
lemma borel_eq_sigmaI5: |
|
231 |
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set" |
|
232 |
assumes borel_eq: "borel = sigma UNIV (range G)" |
|
233 |
assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))" |
|
234 |
assumes F: "\<And>i j. F i j \<in> sets borel" |
|
235 |
shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))" |
|
236 |
using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto |
|
| 38656 | 237 |
|
238 |
lemma halfspace_gt_in_halfspace: |
|
| 47694 | 239 |
"{x\<Colon>'a. a < x $$ i} \<in> sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))"
|
240 |
(is "?set \<in> ?SIGMA") |
|
| 38656 | 241 |
proof - |
| 47694 | 242 |
interpret sigma_algebra UNIV ?SIGMA |
243 |
by (intro sigma_algebra_sigma_sets) simp_all |
|
244 |
have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
|
|
| 38656 | 245 |
proof (safe, simp_all add: not_less) |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
246 |
fix x :: 'a assume "a < x $$ i" |
| 38656 | 247 |
with reals_Archimedean[of "x $$ i - a"] |
248 |
obtain n where "a + 1 / real (Suc n) < x $$ i" |
|
249 |
by (auto simp: inverse_eq_divide field_simps) |
|
250 |
then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i" |
|
251 |
by (blast intro: less_imp_le) |
|
252 |
next |
|
253 |
fix x n |
|
254 |
have "a < a + 1 / real (Suc n)" by auto |
|
255 |
also assume "\<dots> \<le> x" |
|
256 |
finally show "a < x" . |
|
257 |
qed |
|
| 47694 | 258 |
show "?set \<in> ?SIGMA" unfolding * |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
259 |
by (auto del: Diff intro!: Diff) |
| 40859 | 260 |
qed |
| 38656 | 261 |
|
| 47694 | 262 |
lemma borel_eq_halfspace_less: |
263 |
"borel = sigma UNIV (range (\<lambda>(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))"
|
|
264 |
(is "_ = ?SIGMA") |
|
265 |
proof (rule borel_eq_sigmaI3[OF borel_def]) |
|
266 |
fix S :: "'a set" assume "S \<in> {S. open S}"
|
|
267 |
then have "open S" by simp |
|
268 |
from open_UNION[OF this] |
|
269 |
obtain I where *: "S = |
|
270 |
(\<Union>(a, b)\<in>I. |
|
271 |
(\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
|
|
272 |
(\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
|
|
273 |
unfolding greaterThanLessThan_def |
|
274 |
unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] |
|
275 |
unfolding eucl_lessThan_eq_halfspaces[where 'a='a] |
|
276 |
by blast |
|
277 |
show "S \<in> ?SIGMA" |
|
278 |
unfolding * |
|
279 |
by (safe intro!: countable_UN Int countable_INT) (auto intro!: halfspace_gt_in_halfspace) |
|
280 |
qed auto |
|
| 38656 | 281 |
|
| 47694 | 282 |
lemma borel_eq_halfspace_le: |
283 |
"borel = sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i \<le> a}))"
|
|
284 |
(is "_ = ?SIGMA") |
|
285 |
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) |
|
286 |
fix a i |
|
287 |
have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
|
|
288 |
proof (safe, simp_all) |
|
289 |
fix x::'a assume *: "x$$i < a" |
|
290 |
with reals_Archimedean[of "a - x$$i"] |
|
291 |
obtain n where "x $$ i < a - 1 / (real (Suc n))" |
|
292 |
by (auto simp: field_simps inverse_eq_divide) |
|
293 |
then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))" |
|
294 |
by (blast intro: less_imp_le) |
|
295 |
next |
|
296 |
fix x::'a and n |
|
297 |
assume "x$$i \<le> a - 1 / real (Suc n)" |
|
298 |
also have "\<dots> < a" by auto |
|
299 |
finally show "x$$i < a" . |
|
300 |
qed |
|
301 |
show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
|
|
302 |
by (safe intro!: countable_UN) auto |
|
303 |
qed auto |
|
| 38656 | 304 |
|
| 47694 | 305 |
lemma borel_eq_halfspace_ge: |
306 |
"borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i}))"
|
|
307 |
(is "_ = ?SIGMA") |
|
308 |
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) |
|
309 |
fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
|
|
310 |
show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
|
|
311 |
by (safe intro!: compl_sets) auto |
|
312 |
qed auto |
|
| 38656 | 313 |
|
| 47694 | 314 |
lemma borel_eq_halfspace_greater: |
315 |
"borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a < x $$ i}))"
|
|
316 |
(is "_ = ?SIGMA") |
|
317 |
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) |
|
318 |
fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
|
|
319 |
show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
|
|
320 |
by (safe intro!: compl_sets) auto |
|
321 |
qed auto |
|
322 |
||
323 |
lemma borel_eq_atMost: |
|
324 |
"borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
|
|
325 |
(is "_ = ?SIGMA") |
|
326 |
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) |
|
327 |
fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
|
|
| 38656 | 328 |
proof cases |
| 47694 | 329 |
assume "i < DIM('a)"
|
| 38656 | 330 |
then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
|
331 |
proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) |
|
332 |
fix x |
|
333 |
from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
|
|
334 |
then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
|
|
335 |
by (subst (asm) Max_le_iff) auto |
|
336 |
then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
|
|
337 |
by (auto intro!: exI[of _ k]) |
|
338 |
qed |
|
| 47694 | 339 |
show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
|
340 |
by (safe intro!: countable_UN) auto |
|
341 |
qed (auto intro: sigma_sets_top sigma_sets.Empty) |
|
342 |
qed auto |
|
| 38656 | 343 |
|
| 47694 | 344 |
lemma borel_eq_greaterThan: |
345 |
"borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {a<..}))"
|
|
346 |
(is "_ = ?SIGMA") |
|
347 |
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) |
|
348 |
fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
|
|
| 38656 | 349 |
proof cases |
| 47694 | 350 |
assume "i < DIM('a)"
|
351 |
have "{x::'a. x$$i \<le> a} = UNIV - {x::'a. a < x$$i}" by auto
|
|
| 38656 | 352 |
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)`
|
353 |
proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) |
|
354 |
fix x |
|
| 44666 | 355 |
from reals_Archimedean2[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
|
| 38656 | 356 |
guess k::nat .. note k = this |
357 |
{ fix i assume "i < DIM('a)"
|
|
358 |
then have "-x$$i < real k" |
|
359 |
using k by (subst (asm) Max_less_iff) auto |
|
360 |
then have "- real k < x$$i" by simp } |
|
361 |
then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
|
|
362 |
by (auto intro!: exI[of _ k]) |
|
363 |
qed |
|
| 47694 | 364 |
finally show "{x. x$$i \<le> a} \<in> ?SIGMA"
|
| 38656 | 365 |
apply (simp only:) |
366 |
apply (safe intro!: countable_UN Diff) |
|
| 47694 | 367 |
apply (auto intro: sigma_sets_top) |
| 46731 | 368 |
done |
| 47694 | 369 |
qed (auto intro: sigma_sets_top sigma_sets.Empty) |
370 |
qed auto |
|
| 40859 | 371 |
|
| 47694 | 372 |
lemma borel_eq_lessThan: |
373 |
"borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {..<a}))"
|
|
374 |
(is "_ = ?SIGMA") |
|
375 |
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge]) |
|
376 |
fix a i show "{x. a \<le> x$$i} \<in> ?SIGMA"
|
|
| 40859 | 377 |
proof cases |
378 |
fix a i assume "i < DIM('a)"
|
|
| 47694 | 379 |
have "{x::'a. a \<le> x$$i} = UNIV - {x::'a. x$$i < a}" by auto
|
| 40859 | 380 |
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)`
|
381 |
proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) |
|
382 |
fix x |
|
| 44666 | 383 |
from reals_Archimedean2[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
|
| 40859 | 384 |
guess k::nat .. note k = this |
385 |
{ fix i assume "i < DIM('a)"
|
|
386 |
then have "x$$i < real k" |
|
387 |
using k by (subst (asm) Max_less_iff) auto |
|
388 |
then have "x$$i < real k" by simp } |
|
389 |
then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
|
|
390 |
by (auto intro!: exI[of _ k]) |
|
391 |
qed |
|
| 47694 | 392 |
finally show "{x. a \<le> x$$i} \<in> ?SIGMA"
|
| 40859 | 393 |
apply (simp only:) |
394 |
apply (safe intro!: countable_UN Diff) |
|
| 47694 | 395 |
apply (auto intro: sigma_sets_top) |
| 46731 | 396 |
done |
| 47694 | 397 |
qed (auto intro: sigma_sets_top sigma_sets.Empty) |
| 40859 | 398 |
qed auto |
399 |
||
400 |
lemma borel_eq_atLeastAtMost: |
|
| 47694 | 401 |
"borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
|
402 |
(is "_ = ?SIGMA") |
|
403 |
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) |
|
404 |
fix a::'a |
|
405 |
have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
|
|
406 |
proof (safe, simp_all add: eucl_le[where 'a='a]) |
|
407 |
fix x |
|
408 |
from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
|
|
409 |
guess k::nat .. note k = this |
|
410 |
{ fix i assume "i < DIM('a)"
|
|
411 |
with k have "- x$$i \<le> real k" |
|
412 |
by (subst (asm) Max_le_iff) (auto simp: field_simps) |
|
413 |
then have "- real k \<le> x$$i" by simp } |
|
414 |
then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
|
|
415 |
by (auto intro!: exI[of _ k]) |
|
416 |
qed |
|
417 |
show "{..a} \<in> ?SIGMA" unfolding *
|
|
418 |
by (safe intro!: countable_UN) |
|
419 |
(auto intro!: sigma_sets_top) |
|
| 40859 | 420 |
qed auto |
421 |
||
422 |
lemma borel_eq_greaterThanLessThan: |
|
| 47694 | 423 |
"borel = sigma UNIV (range (\<lambda> (a, b). {a <..< b} :: 'a \<Colon> ordered_euclidean_space set))"
|
| 40859 | 424 |
(is "_ = ?SIGMA") |
| 47694 | 425 |
proof (rule borel_eq_sigmaI1[OF borel_def]) |
426 |
fix M :: "'a set" assume "M \<in> {S. open S}"
|
|
427 |
then have "open M" by simp |
|
428 |
show "M \<in> ?SIGMA" |
|
429 |
apply (subst open_UNION[OF `open M`]) |
|
430 |
apply (safe intro!: countable_UN) |
|
431 |
apply auto |
|
432 |
done |
|
| 38656 | 433 |
qed auto |
434 |
||
| 42862 | 435 |
lemma borel_eq_atLeastLessThan: |
| 47694 | 436 |
"borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
|
437 |
proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) |
|
438 |
have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto |
|
439 |
fix x :: real |
|
440 |
have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
|
|
441 |
by (auto simp: move_uminus real_arch_simple) |
|
442 |
then show "{..< x} \<in> ?SIGMA"
|
|
443 |
by (auto intro: sigma_sets.intros) |
|
| 40859 | 444 |
qed auto |
445 |
||
| 50087 | 446 |
lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" |
447 |
unfolding borel_def |
|
448 |
proof (intro sigma_eqI sigma_sets_eqI, safe) |
|
449 |
fix x :: "'a set" assume "open x" |
|
450 |
hence "x = UNIV - (UNIV - x)" by auto |
|
451 |
also have "\<dots> \<in> sigma_sets UNIV (Collect closed)" |
|
452 |
by (rule sigma_sets.Compl) |
|
453 |
(auto intro!: sigma_sets.Basic simp: `open x`) |
|
454 |
finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp |
|
455 |
next |
|
456 |
fix x :: "'a set" assume "closed x" |
|
457 |
hence "x = UNIV - (UNIV - x)" by auto |
|
458 |
also have "\<dots> \<in> sigma_sets UNIV (Collect open)" |
|
459 |
by (rule sigma_sets.Compl) |
|
460 |
(auto intro!: sigma_sets.Basic simp: `closed x`) |
|
461 |
finally show "x \<in> sigma_sets UNIV (Collect open)" by simp |
|
462 |
qed simp_all |
|
463 |
||
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
464 |
lemma borel_eq_enumerated_basis: |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
465 |
fixes f::"nat\<Rightarrow>'a::topological_space set" |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
466 |
assumes "topological_basis (range f)" |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
467 |
shows "borel = sigma UNIV (range f)" |
| 50087 | 468 |
unfolding borel_def |
469 |
proof (intro sigma_eqI sigma_sets_eqI, safe) |
|
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
470 |
interpret enumerates_basis proof qed (rule assms) |
| 50087 | 471 |
fix x::"'a set" assume "open x" |
472 |
from open_enumerable_basisE[OF this] guess N . |
|
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
473 |
hence x: "x = (\<Union>n. if n \<in> N then f n else {})" by (auto split: split_if_asm)
|
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
474 |
also have "\<dots> \<in> sigma_sets UNIV (range f)" by (auto intro!: sigma_sets.Empty Union) |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
475 |
finally show "x \<in> sigma_sets UNIV (range f)" . |
| 50087 | 476 |
next |
477 |
fix n |
|
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
478 |
have "open (f n)" by (rule topological_basis_open[OF assms]) simp |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
479 |
thus "f n \<in> sigma_sets UNIV (Collect open)" by auto |
| 50087 | 480 |
qed simp_all |
481 |
||
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
482 |
lemma borel_eq_enum_basis: |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
483 |
"borel = sigma UNIV (range enum_basis)" |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
484 |
by (rule borel_eq_enumerated_basis[OF enum_basis_basis]) |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
485 |
|
| 47694 | 486 |
lemma borel_measurable_halfspacesI: |
| 38656 | 487 |
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" |
| 47694 | 488 |
assumes F: "borel = sigma UNIV (range F)" |
489 |
and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M" |
|
490 |
and S: "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
|
|
| 38656 | 491 |
shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
|
492 |
proof safe |
|
493 |
fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
|
|
494 |
then show "S a i \<in> sets M" unfolding assms |
|
| 47694 | 495 |
by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1)) |
| 38656 | 496 |
next |
497 |
assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
|
|
498 |
{ fix a i have "S a i \<in> sets M"
|
|
499 |
proof cases |
|
500 |
assume "i < DIM('c)"
|
|
501 |
with a show ?thesis unfolding assms(2) by simp |
|
502 |
next |
|
503 |
assume "\<not> i < DIM('c)"
|
|
| 47694 | 504 |
from S[OF this] show ?thesis . |
| 38656 | 505 |
qed } |
| 47694 | 506 |
then show "f \<in> borel_measurable M" |
507 |
by (auto intro!: measurable_measure_of simp: S_eq F) |
|
| 38656 | 508 |
qed |
509 |
||
| 47694 | 510 |
lemma borel_measurable_iff_halfspace_le: |
| 38656 | 511 |
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" |
512 |
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 | 513 |
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto |
| 38656 | 514 |
|
| 47694 | 515 |
lemma borel_measurable_iff_halfspace_less: |
| 38656 | 516 |
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" |
517 |
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 | 518 |
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto |
| 38656 | 519 |
|
| 47694 | 520 |
lemma borel_measurable_iff_halfspace_ge: |
| 38656 | 521 |
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" |
522 |
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 | 523 |
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto |
| 38656 | 524 |
|
| 47694 | 525 |
lemma borel_measurable_iff_halfspace_greater: |
| 38656 | 526 |
fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space" |
527 |
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
|
|
| 47694 | 528 |
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto |
| 38656 | 529 |
|
| 47694 | 530 |
lemma borel_measurable_iff_le: |
| 38656 | 531 |
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
|
532 |
using borel_measurable_iff_halfspace_le[where 'c=real] by simp |
|
533 |
||
| 47694 | 534 |
lemma borel_measurable_iff_less: |
| 38656 | 535 |
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
|
536 |
using borel_measurable_iff_halfspace_less[where 'c=real] by simp |
|
537 |
||
| 47694 | 538 |
lemma borel_measurable_iff_ge: |
| 38656 | 539 |
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
540 |
using borel_measurable_iff_halfspace_ge[where 'c=real] |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
541 |
by simp |
| 38656 | 542 |
|
| 47694 | 543 |
lemma borel_measurable_iff_greater: |
| 38656 | 544 |
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
|
545 |
using borel_measurable_iff_halfspace_greater[where 'c=real] by simp |
|
546 |
||
| 47694 | 547 |
lemma borel_measurable_euclidean_space: |
|
39087
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
548 |
fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space" |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
549 |
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
|
550 |
proof safe |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
551 |
fix i assume "f \<in> borel_measurable M" |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
552 |
then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M" |
| 41025 | 553 |
by (auto intro: borel_measurable_euclidean_component) |
|
39087
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
554 |
next |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
555 |
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
|
556 |
then show "f \<in> borel_measurable M" |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
557 |
unfolding borel_measurable_iff_halfspace_le by auto |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
558 |
qed |
|
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
559 |
|
| 38656 | 560 |
subsection "Borel measurable operators" |
561 |
||
| 49774 | 562 |
lemma borel_measurable_continuous_on: |
563 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
564 |
assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M" |
|
565 |
shows "(\<lambda>x. f (g x)) \<in> borel_measurable M" |
|
566 |
using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def) |
|
567 |
||
568 |
lemma borel_measurable_continuous_on_open': |
|
569 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space" |
|
570 |
assumes cont: "continuous_on A f" "open A" |
|
571 |
shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _") |
|
572 |
proof (rule borel_measurableI) |
|
573 |
fix S :: "'b set" assume "open S" |
|
574 |
then have "open {x\<in>A. f x \<in> S}"
|
|
575 |
by (intro continuous_open_preimage[OF cont]) auto |
|
576 |
then have *: "{x\<in>A. f x \<in> S} \<in> sets borel" by auto
|
|
577 |
have "?f -` S \<inter> space borel = |
|
578 |
{x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
|
|
579 |
by (auto split: split_if_asm) |
|
580 |
also have "\<dots> \<in> sets borel" |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
581 |
using * `open A` by auto |
| 49774 | 582 |
finally show "?f -` S \<inter> space borel \<in> sets borel" . |
583 |
qed |
|
584 |
||
585 |
lemma borel_measurable_continuous_on_open: |
|
586 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space" |
|
587 |
assumes cont: "continuous_on A f" "open A" |
|
588 |
assumes g: "g \<in> borel_measurable M" |
|
589 |
shows "(\<lambda>x. if g x \<in> A then f (g x) else c) \<in> borel_measurable M" |
|
590 |
using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c] |
|
591 |
by (simp add: comp_def) |
|
592 |
||
| 50003 | 593 |
lemma borel_measurable_uminus[measurable (raw)]: |
| 49774 | 594 |
fixes g :: "'a \<Rightarrow> real" |
595 |
assumes g: "g \<in> borel_measurable M" |
|
596 |
shows "(\<lambda>x. - g x) \<in> borel_measurable M" |
|
597 |
by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id) |
|
598 |
||
599 |
lemma euclidean_component_prod: |
|
600 |
fixes x :: "'a :: euclidean_space \<times> 'b :: euclidean_space" |
|
601 |
shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))"
|
|
602 |
unfolding euclidean_component_def basis_prod_def inner_prod_def by auto |
|
603 |
||
| 50003 | 604 |
lemma borel_measurable_Pair[measurable (raw)]: |
| 49774 | 605 |
fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space" |
606 |
assumes f: "f \<in> borel_measurable M" |
|
607 |
assumes g: "g \<in> borel_measurable M" |
|
608 |
shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" |
|
609 |
proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI) |
|
610 |
fix i and a :: real assume i: "i < DIM('b \<times> 'c)"
|
|
611 |
have [simp]: "\<And>P A B C. {w. (P \<longrightarrow> A w \<and> B w) \<and> (\<not> P \<longrightarrow> A w \<and> C w)} =
|
|
612 |
{w. A w \<and> (P \<longrightarrow> B w) \<and> (\<not> P \<longrightarrow> C w)}" by auto
|
|
613 |
from i f g show "{w \<in> space M. (f w, g w) $$ i \<le> a} \<in> sets M"
|
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
614 |
by (auto simp: euclidean_component_prod) |
| 49774 | 615 |
qed |
616 |
||
617 |
lemma continuous_on_fst: "continuous_on UNIV fst" |
|
618 |
proof - |
|
619 |
have [simp]: "range fst = UNIV" by (auto simp: image_iff) |
|
620 |
show ?thesis |
|
621 |
using closed_vimage_fst |
|
622 |
by (auto simp: continuous_on_closed closed_closedin vimage_def) |
|
623 |
qed |
|
624 |
||
625 |
lemma continuous_on_snd: "continuous_on UNIV snd" |
|
626 |
proof - |
|
627 |
have [simp]: "range snd = UNIV" by (auto simp: image_iff) |
|
628 |
show ?thesis |
|
629 |
using closed_vimage_snd |
|
630 |
by (auto simp: continuous_on_closed closed_closedin vimage_def) |
|
631 |
qed |
|
632 |
||
633 |
lemma borel_measurable_continuous_Pair: |
|
634 |
fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space" |
|
| 50003 | 635 |
assumes [measurable]: "f \<in> borel_measurable M" |
636 |
assumes [measurable]: "g \<in> borel_measurable M" |
|
| 49774 | 637 |
assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))" |
638 |
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" |
|
639 |
proof - |
|
640 |
have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto |
|
641 |
show ?thesis |
|
642 |
unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto |
|
643 |
qed |
|
644 |
||
| 50003 | 645 |
lemma borel_measurable_add[measurable (raw)]: |
| 49774 | 646 |
fixes f g :: "'a \<Rightarrow> 'c::ordered_euclidean_space" |
647 |
assumes f: "f \<in> borel_measurable M" |
|
648 |
assumes g: "g \<in> borel_measurable M" |
|
649 |
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" |
|
650 |
using f g |
|
651 |
by (rule borel_measurable_continuous_Pair) |
|
652 |
(auto intro: continuous_on_fst continuous_on_snd continuous_on_add) |
|
653 |
||
| 50003 | 654 |
lemma borel_measurable_setsum[measurable (raw)]: |
| 49774 | 655 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real" |
656 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
|
657 |
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" |
|
658 |
proof cases |
|
659 |
assume "finite S" |
|
660 |
thus ?thesis using assms by induct auto |
|
661 |
qed simp |
|
662 |
||
| 50003 | 663 |
lemma borel_measurable_diff[measurable (raw)]: |
| 49774 | 664 |
fixes f :: "'a \<Rightarrow> real" |
665 |
assumes f: "f \<in> borel_measurable M" |
|
666 |
assumes g: "g \<in> borel_measurable M" |
|
667 |
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" |
|
| 50003 | 668 |
unfolding diff_minus using assms by simp |
| 49774 | 669 |
|
| 50003 | 670 |
lemma borel_measurable_times[measurable (raw)]: |
| 49774 | 671 |
fixes f :: "'a \<Rightarrow> real" |
672 |
assumes f: "f \<in> borel_measurable M" |
|
673 |
assumes g: "g \<in> borel_measurable M" |
|
674 |
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" |
|
675 |
using f g |
|
676 |
by (rule borel_measurable_continuous_Pair) |
|
677 |
(auto intro: continuous_on_fst continuous_on_snd continuous_on_mult) |
|
678 |
||
679 |
lemma continuous_on_dist: |
|
680 |
fixes f :: "'a :: t2_space \<Rightarrow> 'b :: metric_space" |
|
681 |
shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. dist (f x) (g x))" |
|
682 |
unfolding continuous_on_eq_continuous_within by (auto simp: continuous_dist) |
|
683 |
||
| 50003 | 684 |
lemma borel_measurable_dist[measurable (raw)]: |
| 49774 | 685 |
fixes g f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" |
686 |
assumes f: "f \<in> borel_measurable M" |
|
687 |
assumes g: "g \<in> borel_measurable M" |
|
688 |
shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M" |
|
689 |
using f g |
|
690 |
by (rule borel_measurable_continuous_Pair) |
|
691 |
(intro continuous_on_dist continuous_on_fst continuous_on_snd) |
|
692 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
693 |
lemma borel_measurable_scaleR[measurable (raw)]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
694 |
fixes g :: "'a \<Rightarrow> 'b::ordered_euclidean_space" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
695 |
assumes f: "f \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
696 |
assumes g: "g \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
697 |
shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
698 |
by (rule borel_measurable_continuous_Pair[OF f g]) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
699 |
(auto intro!: continuous_on_scaleR continuous_on_fst continuous_on_snd) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
700 |
|
| 47694 | 701 |
lemma affine_borel_measurable_vector: |
| 38656 | 702 |
fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" |
703 |
assumes "f \<in> borel_measurable M" |
|
704 |
shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" |
|
705 |
proof (rule borel_measurableI) |
|
706 |
fix S :: "'x set" assume "open S" |
|
707 |
show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" |
|
708 |
proof cases |
|
709 |
assume "b \<noteq> 0" |
|
|
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
710 |
with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") |
|
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
711 |
by (auto intro!: open_affinity simp: scaleR_add_right) |
| 47694 | 712 |
hence "?S \<in> sets borel" by auto |
| 38656 | 713 |
moreover |
714 |
from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" |
|
715 |
apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) |
|
| 40859 | 716 |
ultimately show ?thesis using assms unfolding in_borel_measurable_borel |
| 38656 | 717 |
by auto |
718 |
qed simp |
|
719 |
qed |
|
720 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
721 |
lemma borel_measurable_const_scaleR[measurable (raw)]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
722 |
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
723 |
using affine_borel_measurable_vector[of f M 0 b] by simp |
| 38656 | 724 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
725 |
lemma borel_measurable_const_add[measurable (raw)]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
726 |
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
727 |
using affine_borel_measurable_vector[of f M a 1] by simp |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
728 |
|
| 50003 | 729 |
lemma borel_measurable_setprod[measurable (raw)]: |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
730 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real" |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
731 |
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
|
732 |
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
|
733 |
proof cases |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
734 |
assume "finite S" |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
735 |
thus ?thesis using assms by induct auto |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
736 |
qed simp |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
737 |
|
| 50003 | 738 |
lemma borel_measurable_inverse[measurable (raw)]: |
| 38656 | 739 |
fixes f :: "'a \<Rightarrow> real" |
| 49774 | 740 |
assumes f: "f \<in> borel_measurable M" |
| 35692 | 741 |
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M" |
| 49774 | 742 |
proof - |
| 50003 | 743 |
have "(\<lambda>x::real. if x \<in> UNIV - {0} then inverse x else 0) \<in> borel_measurable borel"
|
744 |
by (intro borel_measurable_continuous_on_open' continuous_on_inverse continuous_on_id) auto |
|
745 |
also have "(\<lambda>x::real. if x \<in> UNIV - {0} then inverse x else 0) = inverse" by (intro ext) auto
|
|
746 |
finally show ?thesis using f by simp |
|
| 35692 | 747 |
qed |
748 |
||
| 50003 | 749 |
lemma borel_measurable_divide[measurable (raw)]: |
750 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. f x / g x::real) \<in> borel_measurable M" |
|
751 |
by (simp add: field_divide_inverse) |
|
| 38656 | 752 |
|
| 50003 | 753 |
lemma borel_measurable_max[measurable (raw)]: |
754 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: real) \<in> borel_measurable M" |
|
755 |
by (simp add: max_def) |
|
| 38656 | 756 |
|
| 50003 | 757 |
lemma borel_measurable_min[measurable (raw)]: |
758 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: real) \<in> borel_measurable M" |
|
759 |
by (simp add: min_def) |
|
| 38656 | 760 |
|
| 50003 | 761 |
lemma borel_measurable_abs[measurable (raw)]: |
762 |
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" |
|
763 |
unfolding abs_real_def by simp |
|
| 38656 | 764 |
|
| 50003 | 765 |
lemma borel_measurable_nth[measurable (raw)]: |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
766 |
"(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel" |
| 50003 | 767 |
by (simp add: nth_conv_component) |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
768 |
|
| 47694 | 769 |
lemma convex_measurable: |
|
42990
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
770 |
fixes a b :: real |
|
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
771 |
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
|
772 |
assumes q: "convex_on { a <..< b} q"
|
| 49774 | 773 |
shows "(\<lambda>x. q (X x)) \<in> borel_measurable M" |
|
42990
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
774 |
proof - |
| 49774 | 775 |
have "(\<lambda>x. if X x \<in> {a <..< b} then q (X x) else 0) \<in> borel_measurable M" (is "?qX")
|
776 |
proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)]) |
|
|
42990
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
777 |
show "open {a<..<b}" by auto
|
|
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
778 |
from this q show "continuous_on {a<..<b} q"
|
|
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
779 |
by (rule convex_on_continuous) |
| 41830 | 780 |
qed |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
781 |
also have "?qX \<longleftrightarrow> (\<lambda>x. q (X x)) \<in> borel_measurable M" |
|
42990
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
782 |
using X by (intro measurable_cong) auto |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
783 |
finally show ?thesis . |
| 41830 | 784 |
qed |
785 |
||
| 50003 | 786 |
lemma borel_measurable_ln[measurable (raw)]: |
| 49774 | 787 |
assumes f: "f \<in> borel_measurable M" |
788 |
shows "(\<lambda>x. ln (f x)) \<in> borel_measurable M" |
|
| 41830 | 789 |
proof - |
790 |
{ fix x :: real assume x: "x \<le> 0"
|
|
791 |
{ fix x::real assume "x \<le> 0" then have "\<And>u. exp u = x \<longleftrightarrow> False" by auto }
|
|
| 49774 | 792 |
from this[of x] x this[of 0] have "ln 0 = ln x" |
793 |
by (auto simp: ln_def) } |
|
794 |
note ln_imp = this |
|
795 |
have "(\<lambda>x. if f x \<in> {0<..} then ln (f x) else ln 0) \<in> borel_measurable M"
|
|
796 |
proof (rule borel_measurable_continuous_on_open[OF _ _ f]) |
|
797 |
show "continuous_on {0<..} ln"
|
|
798 |
by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont |
|
| 41830 | 799 |
simp: continuous_isCont[symmetric]) |
800 |
show "open ({0<..}::real set)" by auto
|
|
801 |
qed |
|
| 49774 | 802 |
also have "(\<lambda>x. if x \<in> {0<..} then ln x else ln 0) = ln"
|
803 |
by (simp add: fun_eq_iff not_less ln_imp) |
|
| 41830 | 804 |
finally show ?thesis . |
805 |
qed |
|
806 |
||
| 50003 | 807 |
lemma borel_measurable_log[measurable (raw)]: |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
808 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M" |
| 49774 | 809 |
unfolding log_def by auto |
| 41830 | 810 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
811 |
lemma measurable_count_space_eq2_countable: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
812 |
fixes f :: "'a => 'c::countable" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
813 |
shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
814 |
proof - |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
815 |
{ fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
816 |
then have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
817 |
by auto |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
818 |
moreover assume "\<And>a. a\<in>A \<Longrightarrow> f -` {a} \<inter> space M \<in> sets M"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
819 |
ultimately have "f -` X \<inter> space M \<in> sets M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
820 |
using `X \<subseteq> A` by (simp add: subset_eq del: UN_simps) } |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
821 |
then show ?thesis |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
822 |
unfolding measurable_def by auto |
| 47761 | 823 |
qed |
824 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
825 |
lemma measurable_real_floor[measurable]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
826 |
"(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" |
| 47761 | 827 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
828 |
have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real a \<le> x \<and> x < real (a + 1))" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
829 |
by (auto intro: floor_eq2) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
830 |
then show ?thesis |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
831 |
by (auto simp: vimage_def measurable_count_space_eq2_countable) |
| 47761 | 832 |
qed |
833 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
834 |
lemma measurable_real_natfloor[measurable]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
835 |
"(natfloor :: real \<Rightarrow> nat) \<in> measurable borel (count_space UNIV)" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
836 |
by (simp add: natfloor_def[abs_def]) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
837 |
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
838 |
lemma measurable_real_ceiling[measurable]: |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
839 |
"(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
840 |
unfolding ceiling_def[abs_def] by simp |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
841 |
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
842 |
lemma borel_measurable_real_floor: "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
843 |
by simp |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
844 |
|
| 50003 | 845 |
lemma borel_measurable_real_natfloor: |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
846 |
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
847 |
by simp |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
848 |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
849 |
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
|
850 |
|
| 50003 | 851 |
lemma borel_measurable_ereal[measurable (raw)]: |
| 43920 | 852 |
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" |
| 49774 | 853 |
using continuous_on_ereal f by (rule borel_measurable_continuous_on) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
854 |
|
| 50003 | 855 |
lemma borel_measurable_real_of_ereal[measurable (raw)]: |
| 49774 | 856 |
fixes f :: "'a \<Rightarrow> ereal" |
857 |
assumes f: "f \<in> borel_measurable M" |
|
858 |
shows "(\<lambda>x. real (f x)) \<in> borel_measurable M" |
|
859 |
proof - |
|
860 |
have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) \<in> borel_measurable M"
|
|
861 |
using continuous_on_real |
|
862 |
by (rule borel_measurable_continuous_on_open[OF _ _ f]) auto |
|
863 |
also have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) = (\<lambda>x. real (f x))"
|
|
864 |
by auto |
|
865 |
finally show ?thesis . |
|
866 |
qed |
|
867 |
||
868 |
lemma borel_measurable_ereal_cases: |
|
869 |
fixes f :: "'a \<Rightarrow> ereal" |
|
870 |
assumes f: "f \<in> borel_measurable M" |
|
871 |
assumes H: "(\<lambda>x. H (ereal (real (f x)))) \<in> borel_measurable M" |
|
872 |
shows "(\<lambda>x. H (f x)) \<in> borel_measurable M" |
|
873 |
proof - |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
874 |
let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real (f x)))" |
| 49774 | 875 |
{ fix x have "H (f x) = ?F x" by (cases "f x") auto }
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
876 |
with f H show ?thesis by simp |
| 47694 | 877 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
878 |
|
| 49774 | 879 |
lemma |
| 50003 | 880 |
fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M" |
881 |
shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" |
|
882 |
and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M" |
|
883 |
and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M" |
|
| 49774 | 884 |
by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) |
885 |
||
886 |
lemma borel_measurable_uminus_eq_ereal[simp]: |
|
887 |
"(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") |
|
888 |
proof |
|
889 |
assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp |
|
890 |
qed auto |
|
891 |
||
892 |
lemma set_Collect_ereal2: |
|
893 |
fixes f g :: "'a \<Rightarrow> ereal" |
|
894 |
assumes f: "f \<in> borel_measurable M" |
|
895 |
assumes g: "g \<in> borel_measurable M" |
|
896 |
assumes H: "{x \<in> space M. H (ereal (real (f x))) (ereal (real (g x)))} \<in> sets M"
|
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
897 |
"{x \<in> space borel. H (-\<infinity>) (ereal x)} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
898 |
"{x \<in> space borel. H (\<infinity>) (ereal x)} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
899 |
"{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
900 |
"{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
|
| 49774 | 901 |
shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
|
902 |
proof - |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
903 |
let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real (g x)))" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
904 |
let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x" |
| 49774 | 905 |
{ fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
906 |
note * = this |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
907 |
from assms show ?thesis |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
908 |
by (subst *) (simp del: space_borel split del: split_if) |
| 49774 | 909 |
qed |
910 |
||
| 50003 | 911 |
lemma [measurable]: |
| 49774 | 912 |
fixes f g :: "'a \<Rightarrow> ereal" |
913 |
assumes f: "f \<in> borel_measurable M" |
|
914 |
assumes g: "g \<in> borel_measurable M" |
|
| 50003 | 915 |
shows borel_measurable_ereal_le: "{x \<in> space M. f x \<le> g x} \<in> sets M"
|
916 |
and borel_measurable_ereal_less: "{x \<in> space M. f x < g x} \<in> sets M"
|
|
917 |
and borel_measurable_ereal_eq: "{w \<in> space M. f w = g w} \<in> sets M"
|
|
918 |
using f g by (simp_all add: set_Collect_ereal2) |
|
919 |
||
920 |
lemma borel_measurable_ereal_neq: |
|
921 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> {w \<in> space M. f w \<noteq> (g w :: ereal)} \<in> sets M"
|
|
922 |
by simp |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
923 |
|
| 47694 | 924 |
lemma borel_measurable_ereal_iff: |
| 43920 | 925 |
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
|
926 |
proof |
| 43920 | 927 |
assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" |
928 |
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
|
929 |
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
|
930 |
qed auto |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
931 |
|
| 47694 | 932 |
lemma borel_measurable_ereal_iff_real: |
| 43923 | 933 |
fixes f :: "'a \<Rightarrow> ereal" |
934 |
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
|
935 |
((\<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
|
936 |
proof safe |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
937 |
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
|
938 |
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
|
939 |
with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
|
| 46731 | 940 |
let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real (f x))" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
941 |
have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto |
| 43920 | 942 |
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
|
943 |
finally show "f \<in> borel_measurable M" . |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
944 |
qed simp_all |
| 41830 | 945 |
|
| 47694 | 946 |
lemma borel_measurable_eq_atMost_ereal: |
| 43923 | 947 |
fixes f :: "'a \<Rightarrow> ereal" |
948 |
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
|
949 |
proof (intro iffI allI) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
950 |
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
|
951 |
show "f \<in> borel_measurable M" |
| 43920 | 952 |
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
|
953 |
proof (intro conjI allI) |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
954 |
fix a :: real |
| 43920 | 955 |
{ 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
|
956 |
have "x = \<infinity>" |
| 43920 | 957 |
proof (rule ereal_top) |
| 44666 | 958 |
fix B from reals_Archimedean2[of B] guess n .. |
| 43920 | 959 |
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
|
960 |
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
|
961 |
qed } |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
962 |
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
|
963 |
by (auto simp: not_le) |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
964 |
then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos
|
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
965 |
by (auto simp del: UN_simps) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
966 |
moreover |
| 43923 | 967 |
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
|
968 |
then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
|
| 43920 | 969 |
moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
|
970 |
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
|
971 |
moreover have "{w \<in> space M. real (f w) \<le> a} =
|
| 43920 | 972 |
(if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
|
973 |
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
|
974 |
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
|
975 |
ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
|
| 35582 | 976 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
977 |
qed (simp add: measurable_sets) |
| 35582 | 978 |
|
| 47694 | 979 |
lemma borel_measurable_eq_atLeast_ereal: |
| 43920 | 980 |
"(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
|
981 |
proof |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
982 |
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
|
983 |
moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
|
| 43920 | 984 |
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
|
985 |
ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M" |
| 43920 | 986 |
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
|
987 |
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
|
988 |
qed (simp add: measurable_sets) |
| 35582 | 989 |
|
| 49774 | 990 |
lemma greater_eq_le_measurable: |
991 |
fixes f :: "'a \<Rightarrow> 'c::linorder" |
|
992 |
shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
|
|
993 |
proof |
|
994 |
assume "f -` {a ..} \<inter> space M \<in> sets M"
|
|
995 |
moreover have "f -` {..< a} \<inter> space M = space M - f -` {a ..} \<inter> space M" by auto
|
|
996 |
ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
|
|
997 |
next |
|
998 |
assume "f -` {..< a} \<inter> space M \<in> sets M"
|
|
999 |
moreover have "f -` {a ..} \<inter> space M = space M - f -` {..< a} \<inter> space M" by auto
|
|
1000 |
ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
|
|
1001 |
qed |
|
1002 |
||
| 47694 | 1003 |
lemma borel_measurable_ereal_iff_less: |
| 43920 | 1004 |
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
|
1005 |
unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable .. |
|
| 38656 | 1006 |
|
| 49774 | 1007 |
lemma less_eq_ge_measurable: |
1008 |
fixes f :: "'a \<Rightarrow> 'c::linorder" |
|
1009 |
shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
|
|
1010 |
proof |
|
1011 |
assume "f -` {a <..} \<inter> space M \<in> sets M"
|
|
1012 |
moreover have "f -` {..a} \<inter> space M = space M - f -` {a <..} \<inter> space M" by auto
|
|
1013 |
ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
|
|
1014 |
next |
|
1015 |
assume "f -` {..a} \<inter> space M \<in> sets M"
|
|
1016 |
moreover have "f -` {a <..} \<inter> space M = space M - f -` {..a} \<inter> space M" by auto
|
|
1017 |
ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
|
|
1018 |
qed |
|
1019 |
||
| 47694 | 1020 |
lemma borel_measurable_ereal_iff_ge: |
| 43920 | 1021 |
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
|
1022 |
unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable .. |
|
| 38656 | 1023 |
|
| 49774 | 1024 |
lemma borel_measurable_ereal2: |
1025 |
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
|
1026 |
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
|
1027 |
assumes g: "g \<in> borel_measurable M" |
| 49774 | 1028 |
assumes H: "(\<lambda>x. H (ereal (real (f x))) (ereal (real (g x)))) \<in> borel_measurable M" |
1029 |
"(\<lambda>x. H (-\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M" |
|
1030 |
"(\<lambda>x. H (\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M" |
|
1031 |
"(\<lambda>x. H (ereal (real (f x))) (-\<infinity>)) \<in> borel_measurable M" |
|
1032 |
"(\<lambda>x. H (ereal (real (f x))) (\<infinity>)) \<in> borel_measurable M" |
|
1033 |
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" |
|
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1034 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1035 |
let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real (g x)))" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1036 |
let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x" |
| 49774 | 1037 |
{ fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1038 |
note * = this |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1039 |
from assms show ?thesis unfolding * by simp |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1040 |
qed |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1041 |
|
| 49774 | 1042 |
lemma |
1043 |
fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M" |
|
1044 |
shows borel_measurable_ereal_eq_const: "{x\<in>space M. f x = c} \<in> sets M"
|
|
1045 |
and borel_measurable_ereal_neq_const: "{x\<in>space M. f x \<noteq> c} \<in> sets M"
|
|
1046 |
using f by auto |
|
| 38656 | 1047 |
|
| 50003 | 1048 |
lemma [measurable(raw)]: |
| 43920 | 1049 |
fixes f :: "'a \<Rightarrow> ereal" |
| 50003 | 1050 |
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1051 |
shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1052 |
and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1053 |
and borel_measurable_ereal_min: "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1054 |
and borel_measurable_ereal_max: "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" |
| 50003 | 1055 |
by (simp_all add: borel_measurable_ereal2 min_def max_def) |
| 49774 | 1056 |
|
| 50003 | 1057 |
lemma [measurable(raw)]: |
| 49774 | 1058 |
fixes f g :: "'a \<Rightarrow> ereal" |
1059 |
assumes "f \<in> borel_measurable M" |
|
1060 |
assumes "g \<in> borel_measurable M" |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1061 |
shows borel_measurable_ereal_diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1062 |
and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M" |
| 50003 | 1063 |
using assms by (simp_all add: minus_ereal_def divide_ereal_def) |
| 38656 | 1064 |
|
| 50003 | 1065 |
lemma borel_measurable_ereal_setsum[measurable (raw)]: |
| 43920 | 1066 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" |
| 41096 | 1067 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
1068 |
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" |
|
1069 |
proof cases |
|
1070 |
assume "finite S" |
|
1071 |
thus ?thesis using assms |
|
1072 |
by induct auto |
|
| 49774 | 1073 |
qed simp |
| 38656 | 1074 |
|
| 50003 | 1075 |
lemma borel_measurable_ereal_setprod[measurable (raw)]: |
| 43920 | 1076 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" |
| 38656 | 1077 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
| 41096 | 1078 |
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" |
| 38656 | 1079 |
proof cases |
1080 |
assume "finite S" |
|
| 41096 | 1081 |
thus ?thesis using assms by induct auto |
1082 |
qed simp |
|
| 38656 | 1083 |
|
| 50003 | 1084 |
lemma borel_measurable_SUP[measurable (raw)]: |
| 43920 | 1085 |
fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal" |
| 38656 | 1086 |
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
|
1087 |
shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M") |
| 43920 | 1088 |
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
|
1089 |
proof |
| 38656 | 1090 |
fix a |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1091 |
have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
|
| 46884 | 1092 |
by (auto simp: less_SUP_iff) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1093 |
then show "?sup -` {a<..} \<inter> space M \<in> sets M"
|
| 38656 | 1094 |
using assms by auto |
1095 |
qed |
|
1096 |
||
| 50003 | 1097 |
lemma borel_measurable_INF[measurable (raw)]: |
| 43920 | 1098 |
fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal" |
| 38656 | 1099 |
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
|
1100 |
shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M") |
| 43920 | 1101 |
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
|
1102 |
proof |
| 38656 | 1103 |
fix a |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1104 |
have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
|
| 46884 | 1105 |
by (auto simp: INF_less_iff) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1106 |
then show "?inf -` {..<a} \<inter> space M \<in> sets M"
|
| 38656 | 1107 |
using assms by auto |
1108 |
qed |
|
1109 |
||
| 50003 | 1110 |
lemma [measurable (raw)]: |
| 43920 | 1111 |
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
|
1112 |
assumes "\<And>i. f i \<in> borel_measurable M" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1113 |
shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1114 |
and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M" |
| 49774 | 1115 |
unfolding liminf_SUPR_INFI limsup_INFI_SUPR using assms by auto |
| 35692 | 1116 |
|
| 50104 | 1117 |
lemma sets_Collect_eventually_sequentially[measurable]: |
| 50003 | 1118 |
"(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
|
1119 |
unfolding eventually_sequentially by simp |
|
1120 |
||
1121 |
lemma sets_Collect_ereal_convergent[measurable]: |
|
1122 |
fixes f :: "nat \<Rightarrow> 'a => ereal" |
|
1123 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
|
1124 |
shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
|
|
1125 |
unfolding convergent_ereal by auto |
|
1126 |
||
1127 |
lemma borel_measurable_extreal_lim[measurable (raw)]: |
|
1128 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
|
1129 |
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" |
|
1130 |
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" |
|
1131 |
proof - |
|
1132 |
have "\<And>x. lim (\<lambda>i. f i x) = (if convergent (\<lambda>i. f i x) then limsup (\<lambda>i. f i x) else (THE i. False))" |
|
1133 |
using convergent_ereal_limsup by (simp add: lim_def convergent_def) |
|
1134 |
then show ?thesis |
|
1135 |
by simp |
|
1136 |
qed |
|
1137 |
||
| 49774 | 1138 |
lemma borel_measurable_ereal_LIMSEQ: |
1139 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
|
1140 |
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" |
|
1141 |
and u: "\<And>i. u i \<in> borel_measurable M" |
|
1142 |
shows "u' \<in> borel_measurable M" |
|
| 47694 | 1143 |
proof - |
| 49774 | 1144 |
have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)" |
1145 |
using u' by (simp add: lim_imp_Liminf[symmetric]) |
|
| 50003 | 1146 |
with u show ?thesis by (simp cong: measurable_cong) |
| 47694 | 1147 |
qed |
1148 |
||
| 50003 | 1149 |
lemma borel_measurable_extreal_suminf[measurable (raw)]: |
| 43920 | 1150 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
| 50003 | 1151 |
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1152 |
shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M" |
| 50003 | 1153 |
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp |
| 39092 | 1154 |
|
1155 |
section "LIMSEQ is borel measurable" |
|
1156 |
||
| 47694 | 1157 |
lemma borel_measurable_LIMSEQ: |
| 39092 | 1158 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" |
1159 |
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" |
|
1160 |
and u: "\<And>i. u i \<in> borel_measurable M" |
|
1161 |
shows "u' \<in> borel_measurable M" |
|
1162 |
proof - |
|
| 43920 | 1163 |
have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)" |
| 46731 | 1164 |
using u' by (simp add: lim_imp_Liminf) |
| 43920 | 1165 |
moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M" |
| 39092 | 1166 |
by auto |
| 43920 | 1167 |
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) |
| 39092 | 1168 |
qed |
1169 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1170 |
lemma sets_Collect_Cauchy[measurable]: |
| 49774 | 1171 |
fixes f :: "nat \<Rightarrow> 'a => real" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1172 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
| 49774 | 1173 |
shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1174 |
unfolding Cauchy_iff2 using f by auto |
| 49774 | 1175 |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1176 |
lemma borel_measurable_lim[measurable (raw)]: |
| 49774 | 1177 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1178 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
| 49774 | 1179 |
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" |
1180 |
proof - |
|
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1181 |
def u' \<equiv> "\<lambda>x. lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1182 |
then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))" |
| 49774 | 1183 |
by (auto simp: lim_def convergent_eq_cauchy[symmetric]) |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1184 |
have "u' \<in> borel_measurable M" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1185 |
proof (rule borel_measurable_LIMSEQ) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1186 |
fix x |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1187 |
have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" |
| 49774 | 1188 |
by (cases "Cauchy (\<lambda>i. f i x)") |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1189 |
(auto simp add: convergent_eq_cauchy[symmetric] convergent_def) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1190 |
then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) ----> u' x" |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1191 |
unfolding u'_def |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1192 |
by (rule convergent_LIMSEQ_iff[THEN iffD1]) |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1193 |
qed measurable |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1194 |
then show ?thesis |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1195 |
unfolding * by measurable |
| 49774 | 1196 |
qed |
1197 |
||
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1198 |
lemma borel_measurable_suminf[measurable (raw)]: |
| 49774 | 1199 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1200 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
| 49774 | 1201 |
shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1202 |
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp |
| 49774 | 1203 |
|
1204 |
end |