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