author | wenzelm |
Mon, 19 Jan 2015 20:39:01 +0100 | |
changeset 59409 | b7cfe12acf2e |
parent 59361 | fd5da2434be4 |
child 59415 | 854fe701c984 |
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 |
|
58876 | 6 |
section {*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 |
|
59088
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
14 |
lemma topological_basis_trivial: "topological_basis {A. open A}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
15 |
by (auto simp: topological_basis_def) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
16 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
17 |
lemma open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
18 |
proof - |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
19 |
have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
20 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
21 |
then show ?thesis |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
22 |
by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
23 |
qed |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
24 |
|
56994 | 25 |
subsection {* Generic Borel spaces *} |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
26 |
|
47694 | 27 |
definition borel :: "'a::topological_space measure" where |
28 |
"borel = sigma UNIV {S. open S}" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
29 |
|
47694 | 30 |
abbreviation "borel_measurable M \<equiv> measurable M borel" |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
31 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
32 |
lemma in_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
33 |
"f \<in> borel_measurable M \<longleftrightarrow> |
47694 | 34 |
(\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)" |
40859 | 35 |
by (auto simp add: measurable_def borel_def) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
36 |
|
40859 | 37 |
lemma in_borel_measurable_borel: |
38656 | 38 |
"f \<in> borel_measurable M \<longleftrightarrow> |
40859 | 39 |
(\<forall>S \<in> sets borel. |
38656 | 40 |
f -` S \<inter> space M \<in> sets M)" |
40859 | 41 |
by (auto simp add: measurable_def borel_def) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
42 |
|
40859 | 43 |
lemma space_borel[simp]: "space borel = UNIV" |
44 |
unfolding borel_def by auto |
|
38656 | 45 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
46 |
lemma space_in_borel[measurable]: "UNIV \<in> sets borel" |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
47 |
unfolding borel_def by auto |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
48 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
49 |
lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
50 |
unfolding borel_def by (rule sets_measure_of) simp |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
51 |
|
50387 | 52 |
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
|
53 |
unfolding borel_def pred_def by auto |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
54 |
|
50003 | 55 |
lemma borel_open[measurable (raw generic)]: |
40859 | 56 |
assumes "open A" shows "A \<in> sets borel" |
38656 | 57 |
proof - |
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
58 |
have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms . |
47694 | 59 |
thus ?thesis unfolding borel_def by auto |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
60 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
61 |
|
50003 | 62 |
lemma borel_closed[measurable (raw generic)]: |
40859 | 63 |
assumes "closed A" shows "A \<in> sets borel" |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
64 |
proof - |
40859 | 65 |
have "space borel - (- A) \<in> sets borel" |
66 |
using assms unfolding closed_def by (blast intro: borel_open) |
|
38656 | 67 |
thus ?thesis by simp |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
68 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
69 |
|
50003 | 70 |
lemma borel_singleton[measurable]: |
71 |
"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
|
72 |
unfolding insert_def by (rule sets.Un) auto |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
73 |
|
50003 | 74 |
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
|
75 |
unfolding Compl_eq_Diff_UNIV by simp |
41830 | 76 |
|
47694 | 77 |
lemma borel_measurable_vimage: |
38656 | 78 |
fixes f :: "'a \<Rightarrow> 'x::t2_space" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
79 |
assumes borel[measurable]: "f \<in> borel_measurable M" |
38656 | 80 |
shows "f -` {x} \<inter> space M \<in> sets M" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
81 |
by simp |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
82 |
|
47694 | 83 |
lemma borel_measurableI: |
38656 | 84 |
fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space" |
85 |
assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" |
|
86 |
shows "f \<in> borel_measurable M" |
|
40859 | 87 |
unfolding borel_def |
47694 | 88 |
proof (rule measurable_measure_of, simp_all) |
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
89 |
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
|
90 |
using assms[of S] by simp |
40859 | 91 |
qed |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
92 |
|
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
93 |
lemma borel_measurable_const: |
38656 | 94 |
"(\<lambda>x. c) \<in> borel_measurable M" |
47694 | 95 |
by auto |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
96 |
|
50003 | 97 |
lemma borel_measurable_indicator: |
38656 | 98 |
assumes A: "A \<in> sets M" |
99 |
shows "indicator A \<in> borel_measurable M" |
|
46905 | 100 |
unfolding indicator_def [abs_def] using A |
47694 | 101 |
by (auto intro!: measurable_If_set) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
102 |
|
50096 | 103 |
lemma borel_measurable_count_space[measurable (raw)]: |
104 |
"f \<in> borel_measurable (count_space S)" |
|
105 |
unfolding measurable_def by auto |
|
106 |
||
107 |
lemma borel_measurable_indicator'[measurable (raw)]: |
|
108 |
assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M" |
|
109 |
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
|
110 |
unfolding indicator_def[abs_def] |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49774
diff
changeset
|
111 |
by (auto intro!: measurable_If) |
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents:
49774
diff
changeset
|
112 |
|
47694 | 113 |
lemma borel_measurable_indicator_iff: |
40859 | 114 |
"(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M" |
115 |
(is "?I \<in> borel_measurable M \<longleftrightarrow> _") |
|
116 |
proof |
|
117 |
assume "?I \<in> borel_measurable M" |
|
118 |
then have "?I -` {1} \<inter> space M \<in> sets M" |
|
119 |
unfolding measurable_def by auto |
|
120 |
also have "?I -` {1} \<inter> space M = A \<inter> space M" |
|
46905 | 121 |
unfolding indicator_def [abs_def] by auto |
40859 | 122 |
finally show "A \<inter> space M \<in> sets M" . |
123 |
next |
|
124 |
assume "A \<inter> space M \<in> sets M" |
|
125 |
moreover have "?I \<in> borel_measurable M \<longleftrightarrow> |
|
126 |
(indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M" |
|
127 |
by (intro measurable_cong) (auto simp: indicator_def) |
|
128 |
ultimately show "?I \<in> borel_measurable M" by auto |
|
129 |
qed |
|
130 |
||
47694 | 131 |
lemma borel_measurable_subalgebra: |
41545 | 132 |
assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N" |
39092 | 133 |
shows "f \<in> borel_measurable M" |
134 |
using assms unfolding measurable_def by auto |
|
135 |
||
57137 | 136 |
lemma borel_measurable_restrict_space_iff_ereal: |
137 |
fixes f :: "'a \<Rightarrow> ereal" |
|
138 |
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" |
|
139 |
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> |
|
140 |
(\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M" |
|
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
141 |
by (subst measurable_restrict_space_iff) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
142 |
(auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_cong) |
57137 | 143 |
|
144 |
lemma borel_measurable_restrict_space_iff: |
|
145 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
146 |
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" |
|
147 |
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> |
|
148 |
(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> borel_measurable M" |
|
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
149 |
by (subst measurable_restrict_space_iff) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57447
diff
changeset
|
150 |
(auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps cong del: if_cong) |
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
151 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
152 |
lemma cbox_borel[measurable]: "cbox a b \<in> sets borel" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
153 |
by (auto intro: borel_closed) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
154 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
155 |
lemma box_borel[measurable]: "box a b \<in> sets borel" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
156 |
by (auto intro: borel_open) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
157 |
|
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
158 |
lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
159 |
by (auto intro: borel_closed dest!: compact_imp_closed) |
57137 | 160 |
|
59088
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
161 |
lemma second_countable_borel_measurable: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
162 |
fixes X :: "'a::second_countable_topology set set" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
163 |
assumes eq: "open = generate_topology X" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
164 |
shows "borel = sigma UNIV X" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
165 |
unfolding borel_def |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
166 |
proof (intro sigma_eqI sigma_sets_eqI) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
167 |
interpret X: sigma_algebra UNIV "sigma_sets UNIV X" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
168 |
by (rule sigma_algebra_sigma_sets) simp |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
169 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
170 |
fix S :: "'a set" assume "S \<in> Collect open" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
171 |
then have "generate_topology X S" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
172 |
by (auto simp: eq) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
173 |
then show "S \<in> sigma_sets UNIV X" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
174 |
proof induction |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
175 |
case (UN K) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
176 |
then have K: "\<And>k. k \<in> K \<Longrightarrow> open k" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
177 |
unfolding eq by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
178 |
from ex_countable_basis obtain B :: "'a set set" where |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
179 |
B: "\<And>b. b \<in> B \<Longrightarrow> open b" "\<And>X. open X \<Longrightarrow> \<exists>b\<subseteq>B. (\<Union>b) = X" and "countable B" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
180 |
by (auto simp: topological_basis_def) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
181 |
from B(2)[OF K] obtain m where m: "\<And>k. k \<in> K \<Longrightarrow> m k \<subseteq> B" "\<And>k. k \<in> K \<Longrightarrow> (\<Union>m k) = k" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
182 |
by metis |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
183 |
def U \<equiv> "(\<Union>k\<in>K. m k)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
184 |
with m have "countable U" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
185 |
by (intro countable_subset[OF _ `countable B`]) auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
186 |
have "\<Union>U = (\<Union>A\<in>U. A)" by simp |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
187 |
also have "\<dots> = \<Union>K" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
188 |
unfolding U_def UN_simps by (simp add: m) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
189 |
finally have "\<Union>U = \<Union>K" . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
190 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
191 |
have "\<forall>b\<in>U. \<exists>k\<in>K. b \<subseteq> k" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
192 |
using m by (auto simp: U_def) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
193 |
then obtain u where u: "\<And>b. b \<in> U \<Longrightarrow> u b \<in> K" and "\<And>b. b \<in> U \<Longrightarrow> b \<subseteq> u b" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
194 |
by metis |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
195 |
then have "(\<Union>b\<in>U. u b) \<subseteq> \<Union>K" "\<Union>U \<subseteq> (\<Union>b\<in>U. u b)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
196 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
197 |
then have "\<Union>K = (\<Union>b\<in>U. u b)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
198 |
unfolding `\<Union>U = \<Union>K` by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
199 |
also have "\<dots> \<in> sigma_sets UNIV X" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
200 |
using u UN by (intro X.countable_UN' `countable U`) auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
201 |
finally show "\<Union>K \<in> sigma_sets UNIV X" . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
202 |
qed auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
203 |
qed (auto simp: eq intro: generate_topology.Basis) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
204 |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
205 |
lemma borel_measurable_continuous_on_restrict: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
206 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
207 |
assumes f: "continuous_on A f" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
208 |
shows "f \<in> borel_measurable (restrict_space borel A)" |
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
209 |
proof (rule borel_measurableI) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
210 |
fix S :: "'b set" assume "open S" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
211 |
with f obtain T where "f -` S \<inter> A = T \<inter> A" "open T" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
212 |
by (metis continuous_on_open_invariant) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
213 |
then show "f -` S \<inter> space (restrict_space borel A) \<in> sets (restrict_space borel A)" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
214 |
by (force simp add: sets_restrict_space space_restrict_space) |
57137 | 215 |
qed |
216 |
||
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
217 |
lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
218 |
by (drule borel_measurable_continuous_on_restrict) simp |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
219 |
|
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
220 |
lemma borel_measurable_continuous_on_if: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
221 |
assumes [measurable]: "A \<in> sets borel" and f: "continuous_on A f" and g: "continuous_on (- A) g" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
222 |
shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
223 |
by (rule measurable_piecewise_restrict[where |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
224 |
X="\<lambda>b. if b then A else - A" and I=UNIV and f="\<lambda>b x. if b then f x else g x"]) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
225 |
(auto intro: f g borel_measurable_continuous_on_restrict) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
226 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
227 |
lemma borel_measurable_continuous_countable_exceptions: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
228 |
fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
229 |
assumes X: "countable X" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
230 |
assumes "continuous_on (- X) f" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
231 |
shows "f \<in> borel_measurable borel" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
232 |
proof (rule measurable_discrete_difference[OF _ X]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
233 |
have "X \<in> sets borel" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
234 |
by (rule sets.countable[OF _ X]) auto |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
235 |
then show "(\<lambda>x. if x \<in> X then undefined else f x) \<in> borel_measurable borel" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
236 |
by (intro borel_measurable_continuous_on_if assms continuous_intros) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
237 |
qed auto |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
238 |
|
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
239 |
lemma borel_measurable_continuous_on: |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
240 |
assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
241 |
shows "(\<lambda>x. f (g x)) \<in> borel_measurable M" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
242 |
using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
243 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
244 |
lemma borel_measurable_continuous_on_open': |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
245 |
"continuous_on A f \<Longrightarrow> open A \<Longrightarrow> |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
246 |
(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
247 |
using borel_measurable_continuous_on_if[of A f "\<lambda>_. c"] by (auto intro: continuous_on_const) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
248 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
249 |
lemma borel_measurable_continuous_on_open: |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
250 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
251 |
assumes cont: "continuous_on A f" "open A" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
252 |
assumes g: "g \<in> borel_measurable M" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
253 |
shows "(\<lambda>x. if g x \<in> A then f (g x) else c) \<in> borel_measurable M" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
254 |
using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c] |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
255 |
by (simp add: comp_def) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
256 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
257 |
lemma borel_measurable_continuous_on_indicator: |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
258 |
fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
259 |
assumes A: "A \<in> sets borel" and f: "continuous_on A f" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
260 |
shows "(\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
261 |
using borel_measurable_continuous_on_if[OF assms, of "\<lambda>_. 0"] |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
262 |
by (simp add: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] continuous_on_const |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
263 |
cong del: if_cong) |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
264 |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
265 |
lemma borel_eq_countable_basis: |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
266 |
fixes B::"'a::topological_space set set" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
267 |
assumes "countable B" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
268 |
assumes "topological_basis B" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
269 |
shows "borel = sigma UNIV B" |
50087 | 270 |
unfolding borel_def |
271 |
proof (intro sigma_eqI sigma_sets_eqI, safe) |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
272 |
interpret countable_basis using assms by unfold_locales |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
273 |
fix X::"'a set" assume "open X" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
274 |
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
|
275 |
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
|
276 |
by (blast intro: sigma_sets_UNION `countable B` countable_subset) |
50087 | 277 |
next |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
278 |
fix b assume "b \<in> B" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
279 |
hence "open b" by (rule topological_basis_open[OF assms(2)]) |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
280 |
thus "b \<in> sigma_sets UNIV (Collect open)" by auto |
50087 | 281 |
qed simp_all |
282 |
||
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
|
283 |
lemma borel_measurable_Pair[measurable (raw)]: |
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
284 |
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
|
285 |
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
|
286 |
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
|
287 |
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
|
288 |
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
|
289 |
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
|
290 |
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
|
291 |
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
|
292 |
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
|
293 |
by (auto intro!: countable_basis topological_basis_prod is_basis) |
38656 | 294 |
|
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
|
295 |
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
|
296 |
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
|
297 |
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
|
298 |
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
|
299 |
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
|
300 |
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
|
301 |
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
|
302 |
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
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
qed auto |
39087
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
307 |
qed |
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
hoelzl
parents:
39083
diff
changeset
|
308 |
|
49774 | 309 |
lemma borel_measurable_continuous_Pair: |
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
310 |
fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology" |
50003 | 311 |
assumes [measurable]: "f \<in> borel_measurable M" |
312 |
assumes [measurable]: "g \<in> borel_measurable M" |
|
49774 | 313 |
assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))" |
314 |
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" |
|
315 |
proof - |
|
316 |
have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto |
|
317 |
show ?thesis |
|
318 |
unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto |
|
319 |
qed |
|
320 |
||
59088
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
321 |
subsection {* Borel spaces on order topologies *} |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
322 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
323 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
324 |
lemma borel_Iio: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
325 |
"borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
326 |
unfolding second_countable_borel_measurable[OF open_generated_order] |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
327 |
proof (intro sigma_eqI sigma_sets_eqI) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
328 |
from countable_dense_setE guess D :: "'a set" . note D = this |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
329 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
330 |
interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
331 |
by (rule sigma_algebra_sigma_sets) simp |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
332 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
333 |
fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
334 |
then obtain y where "A = {y <..} \<or> A = {..< y}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
335 |
by blast |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
336 |
then show "A \<in> sigma_sets UNIV (range lessThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
337 |
proof |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
338 |
assume A: "A = {y <..}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
339 |
show ?thesis |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
340 |
proof cases |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
341 |
assume "\<forall>x>y. \<exists>d. y < d \<and> d < x" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
342 |
with D(2)[of "{y <..< x}" for x] have "\<forall>x>y. \<exists>d\<in>D. y < d \<and> d < x" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
343 |
by (auto simp: set_eq_iff) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
344 |
then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. y < d}. {..< d})" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
345 |
by (auto simp: A) (metis less_asym) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
346 |
also have "\<dots> \<in> sigma_sets UNIV (range lessThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
347 |
using D(1) by (intro L.Diff L.top L.countable_INT'') auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
348 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
349 |
next |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
350 |
assume "\<not> (\<forall>x>y. \<exists>d. y < d \<and> d < x)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
351 |
then obtain x where "y < x" "\<And>d. y < d \<Longrightarrow> \<not> d < x" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
352 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
353 |
then have "A = UNIV - {..< x}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
354 |
unfolding A by (auto simp: not_less[symmetric]) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
355 |
also have "\<dots> \<in> sigma_sets UNIV (range lessThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
356 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
357 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
358 |
qed |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
359 |
qed auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
360 |
qed auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
361 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
362 |
lemma borel_Ioi: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
363 |
"borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
364 |
unfolding second_countable_borel_measurable[OF open_generated_order] |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
365 |
proof (intro sigma_eqI sigma_sets_eqI) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
366 |
from countable_dense_setE guess D :: "'a set" . note D = this |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
367 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
368 |
interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
369 |
by (rule sigma_algebra_sigma_sets) simp |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
370 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
371 |
fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
372 |
then obtain y where "A = {y <..} \<or> A = {..< y}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
373 |
by blast |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
374 |
then show "A \<in> sigma_sets UNIV (range greaterThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
375 |
proof |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
376 |
assume A: "A = {..< y}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
377 |
show ?thesis |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
378 |
proof cases |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
379 |
assume "\<forall>x<y. \<exists>d. x < d \<and> d < y" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
380 |
with D(2)[of "{x <..< y}" for x] have "\<forall>x<y. \<exists>d\<in>D. x < d \<and> d < y" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
381 |
by (auto simp: set_eq_iff) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
382 |
then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. d < y}. {d <..})" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
383 |
by (auto simp: A) (metis less_asym) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
384 |
also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
385 |
using D(1) by (intro L.Diff L.top L.countable_INT'') auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
386 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
387 |
next |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
388 |
assume "\<not> (\<forall>x<y. \<exists>d. x < d \<and> d < y)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
389 |
then obtain x where "x < y" "\<And>d. y > d \<Longrightarrow> x \<ge> d" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
390 |
by (auto simp: not_less[symmetric]) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
391 |
then have "A = UNIV - {x <..}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
392 |
unfolding A Compl_eq_Diff_UNIV[symmetric] by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
393 |
also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
394 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
395 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
396 |
qed |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
397 |
qed auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
398 |
qed auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
399 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
400 |
lemma borel_measurableI_less: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
401 |
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
402 |
shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
403 |
unfolding borel_Iio |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
404 |
by (rule measurable_measure_of) (auto simp: Int_def conj_commute) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
405 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
406 |
lemma borel_measurableI_greater: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
407 |
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
408 |
shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
409 |
unfolding borel_Ioi |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
410 |
by (rule measurable_measure_of) (auto simp: Int_def conj_commute) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
411 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
412 |
lemma borel_measurable_SUP[measurable (raw)]: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
413 |
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
414 |
assumes [simp]: "countable I" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
415 |
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
416 |
shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
417 |
by (rule borel_measurableI_greater) (simp add: less_SUP_iff) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
418 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
419 |
lemma borel_measurable_INF[measurable (raw)]: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
420 |
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
421 |
assumes [simp]: "countable I" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
422 |
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
423 |
shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
424 |
by (rule borel_measurableI_less) (simp add: INF_less_iff) |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
425 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
426 |
lemma borel_measurable_lfp[consumes 1, case_names continuity step]: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
427 |
fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
428 |
assumes "Order_Continuity.continuous F" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
429 |
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
430 |
shows "lfp F \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
431 |
proof - |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
432 |
{ fix i have "((F ^^ i) bot) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
433 |
by (induct i) (auto intro!: *) } |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
434 |
then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
435 |
by measurable |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
436 |
also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
437 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
438 |
also have "(SUP i. (F ^^ i) bot) = lfp F" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
439 |
by (rule continuous_lfp[symmetric]) fact |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
440 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
441 |
qed |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
442 |
|
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
443 |
lemma borel_measurable_gfp[consumes 1, case_names continuity step]: |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
444 |
fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
445 |
assumes "Order_Continuity.down_continuous F" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
446 |
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
447 |
shows "gfp F \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
448 |
proof - |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
449 |
{ fix i have "((F ^^ i) top) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
450 |
by (induct i) (auto intro!: * simp: bot_fun_def) } |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
451 |
then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> borel_measurable M" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
452 |
by measurable |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
453 |
also have "(\<lambda>x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
454 |
by auto |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
455 |
also have "\<dots> = gfp F" |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
456 |
by (rule down_continuous_gfp[symmetric]) fact |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
457 |
finally show ?thesis . |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
458 |
qed |
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents:
59000
diff
changeset
|
459 |
|
56994 | 460 |
subsection {* Borel spaces on euclidean spaces *} |
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
|
461 |
|
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 |
lemma borel_measurable_inner[measurable (raw)]: |
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
463 |
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
|
464 |
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
|
465 |
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
|
466 |
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
|
467 |
using assms |
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
468 |
by (rule borel_measurable_continuous_Pair) (intro continuous_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
|
469 |
|
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 |
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
|
471 |
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
|
472 |
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
|
473 |
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
|
474 |
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
|
475 |
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
|
476 |
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
|
477 |
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
|
478 |
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
|
479 |
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
|
480 |
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
|
481 |
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
|
482 |
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
|
483 |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
484 |
notation |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
485 |
eucl_less (infix "<e" 50) |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
486 |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
487 |
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
|
488 |
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
|
489 |
by auto |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
490 |
|
51683
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
491 |
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
|
492 |
fixes a b :: "'a\<Colon>ordered_euclidean_space" |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
493 |
shows "{x. x <e a} \<in> sets borel" |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
494 |
and "{x. a <e x} \<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
|
495 |
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
|
496 |
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
|
497 |
and "{a..b} \<in> sets borel" |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
498 |
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
|
499 |
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
|
500 |
unfolding box_oc box_co |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
501 |
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
|
502 |
|
51683
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
503 |
lemma open_Collect_less: |
53216 | 504 |
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
|
505 |
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
|
506 |
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
|
507 |
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
|
508 |
proof - |
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
509 |
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
|
510 |
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
|
511 |
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
|
512 |
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
|
513 |
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
|
514 |
qed |
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
515 |
|
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
516 |
lemma closed_Collect_le: |
53216 | 517 |
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
|
518 |
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
|
519 |
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
|
520 |
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
|
521 |
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
|
522 |
|
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
|
523 |
lemma borel_measurable_less[measurable]: |
53216 | 524 |
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
|
525 |
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
|
526 |
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
|
527 |
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
|
528 |
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
|
529 |
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
|
530 |
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
|
531 |
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
|
532 |
by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] |
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
533 |
continuous_intros) |
51683
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
534 |
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
|
535 |
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
|
536 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
537 |
lemma |
53216 | 538 |
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
|
539 |
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
|
540 |
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
|
541 |
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
|
542 |
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
|
543 |
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
|
544 |
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
|
545 |
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
|
546 |
|
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 |
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
|
548 |
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
|
549 |
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
|
550 |
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
|
551 |
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
|
552 |
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
|
553 |
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
|
554 |
|
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 |
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
|
556 |
|
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 |
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
|
558 |
"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
|
559 |
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
|
560 |
|
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 |
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
|
562 |
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
|
563 |
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
|
564 |
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
|
565 |
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
|
566 |
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
|
567 |
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
|
568 |
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
|
569 |
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
|
570 |
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
|
571 |
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
|
572 |
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
|
573 |
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
|
574 |
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
|
575 |
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
|
576 |
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
|
577 |
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
|
578 |
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
|
579 |
|
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 |
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
|
581 |
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
|
582 |
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
|
583 |
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
|
584 |
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
|
585 |
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
|
586 |
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
|
587 |
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
|
588 |
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
|
589 |
|
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 |
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
|
591 |
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
|
592 |
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
|
593 |
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
|
594 |
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
|
595 |
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
|
596 |
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
|
597 |
|
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 |
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
|
599 |
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
|
600 |
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
|
601 |
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
|
602 |
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
|
603 |
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
|
604 |
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
|
605 |
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
|
606 |
|
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 |
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
|
608 |
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
|
609 |
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
|
610 |
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
|
611 |
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
|
612 |
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
|
613 |
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
|
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_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
|
616 |
"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
|
617 |
(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
|
618 |
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
|
619 |
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
|
620 |
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
|
621 |
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
|
622 |
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
|
623 |
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
|
624 |
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
|
625 |
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
|
626 |
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
|
627 |
|
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 |
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
|
629 |
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
|
630 |
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
|
631 |
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
|
632 |
(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
|
633 |
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
|
634 |
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
|
635 |
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
|
636 |
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
|
637 |
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
|
638 |
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
|
639 |
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
|
640 |
obtain n where "a + 1 / real (Suc n) < x \<bullet> i" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
641 |
by (auto simp: field_simps) |
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
|
642 |
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
|
643 |
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
|
644 |
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
|
645 |
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
|
646 |
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
|
647 |
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
|
648 |
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
|
649 |
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
|
650 |
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
|
651 |
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
|
652 |
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
|
653 |
|
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 |
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
|
655 |
"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
|
656 |
(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
|
657 |
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
|
658 |
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
|
659 |
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
|
660 |
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
|
661 |
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
|
662 |
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
|
663 |
(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
|
664 |
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
|
665 |
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
|
666 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
667 |
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
|
668 |
"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
|
669 |
(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
|
670 |
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
|
671 |
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
|
672 |
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
|
673 |
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
|
674 |
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
|
675 |
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
|
676 |
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
|
677 |
obtain n where "x \<bullet> i < a - 1 / (real (Suc n))" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
678 |
by (auto simp: field_simps) |
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
|
679 |
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
|
680 |
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
|
681 |
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
|
682 |
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
|
683 |
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
|
684 |
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
|
685 |
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
|
686 |
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
|
687 |
show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding * |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
688 |
by (intro sets.countable_UN) (auto intro: i) |
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
|
689 |
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
|
690 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
691 |
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
|
692 |
"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
|
693 |
(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
|
694 |
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
|
695 |
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
|
696 |
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
|
697 |
show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding * |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
698 |
using i by (intro sets.compl_sets) auto |
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
|
699 |
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
|
700 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
701 |
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
|
702 |
"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
|
703 |
(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
|
704 |
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
|
705 |
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
|
706 |
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
|
707 |
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
|
708 |
show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding * |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
709 |
by (intro sets.compl_sets) (auto intro: i) |
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
|
710 |
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
|
711 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
712 |
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
|
713 |
"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
|
714 |
(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
|
715 |
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
|
716 |
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
|
717 |
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
|
718 |
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
|
719 |
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
|
720 |
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
|
721 |
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
|
722 |
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
|
723 |
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
|
724 |
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
|
725 |
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
|
726 |
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
|
727 |
show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding * |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
728 |
by (intro sets.countable_UN) auto |
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
|
729 |
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
|
730 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
731 |
lemma borel_eq_greaterThan: |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
732 |
"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
|
733 |
(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
|
734 |
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
|
735 |
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
|
736 |
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
|
737 |
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
|
738 |
also have *: "{x::'a. a < x\<bullet>i} = |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
739 |
(\<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
|
740 |
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
|
741 |
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
|
742 |
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
|
743 |
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
|
744 |
{ 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
|
745 |
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
|
746 |
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
|
747 |
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
|
748 |
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
|
749 |
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
|
750 |
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
|
751 |
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
|
752 |
apply (simp only:) |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
753 |
apply (intro sets.countable_UN sets.Diff) |
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
|
754 |
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
|
755 |
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
|
756 |
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
|
757 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
758 |
lemma borel_eq_lessThan: |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
759 |
"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
|
760 |
(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
|
761 |
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
|
762 |
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
|
763 |
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
|
764 |
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
|
765 |
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
|
766 |
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
|
767 |
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
|
768 |
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
|
769 |
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
|
770 |
{ 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
|
771 |
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
|
772 |
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
|
773 |
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
|
774 |
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
|
775 |
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
|
776 |
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
|
777 |
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
|
778 |
apply (simp only:) |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
779 |
apply (intro sets.countable_UN sets.Diff) |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
780 |
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
|
781 |
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
|
782 |
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
|
783 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
784 |
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
|
785 |
"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
|
786 |
(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
|
787 |
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
|
788 |
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
|
789 |
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
|
790 |
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
|
791 |
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
|
792 |
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
|
793 |
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
|
794 |
{ 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
|
795 |
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
|
796 |
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
|
797 |
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
|
798 |
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
|
799 |
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
|
800 |
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
|
801 |
show "{..a} \<in> ?SIGMA" unfolding * |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
802 |
by (intro sets.countable_UN) |
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
|
803 |
(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
|
804 |
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
|
805 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
806 |
lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
807 |
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
808 |
fix i :: real |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
809 |
have "{..i} = (\<Union>j::nat. {-j <.. i})" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
810 |
by (auto simp: minus_less_iff reals_Archimedean2) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
811 |
also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
812 |
by (intro sets.countable_nat_UN) auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
813 |
finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" . |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
814 |
qed simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
815 |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
816 |
lemma eucl_lessThan: "{x::real. x <e a} = lessThan a" |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
817 |
by (simp add: eucl_less_def lessThan_def) |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
818 |
|
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
|
819 |
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
|
820 |
"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
|
821 |
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
|
822 |
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
|
823 |
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
|
824 |
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
|
825 |
by (auto simp: move_uminus real_arch_simple) |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
826 |
then show "{y. y <e x} \<in> ?SIGMA" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
827 |
by (auto intro: sigma_sets.intros(2-) 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
|
828 |
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
|
829 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
830 |
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
|
831 |
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
|
832 |
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
|
833 |
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
|
834 |
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
|
835 |
also have "\<dots> \<in> sigma_sets UNIV (Collect closed)" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
836 |
by (force intro: sigma_sets.Compl simp: `open 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
|
837 |
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
|
838 |
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
|
839 |
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
|
840 |
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
|
841 |
also have "\<dots> \<in> sigma_sets UNIV (Collect open)" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
842 |
by (force intro: sigma_sets.Compl simp: `closed 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
|
843 |
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
|
844 |
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
|
845 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
846 |
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
|
847 |
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
|
848 |
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
|
849 |
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
|
850 |
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
|
851 |
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
|
852 |
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
|
853 |
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
|
854 |
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
|
855 |
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
|
856 |
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
|
857 |
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
|
858 |
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
|
859 |
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
|
860 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
861 |
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
|
862 |
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
|
863 |
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
|
864 |
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
|
865 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
866 |
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
|
867 |
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
|
868 |
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
|
869 |
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
|
870 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
871 |
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
|
872 |
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
|
873 |
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
|
874 |
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
|
875 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
876 |
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
|
877 |
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
|
878 |
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
|
879 |
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
|
880 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
881 |
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
|
882 |
"(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
|
883 |
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
|
884 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
885 |
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
|
886 |
"(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
|
887 |
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
|
888 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
889 |
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
|
890 |
"(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
|
891 |
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
|
892 |
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
|
893 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
894 |
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
|
895 |
"(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
|
896 |
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
|
897 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
898 |
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
|
899 |
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
|
900 |
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
|
901 |
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
|
902 |
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
|
903 |
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
|
904 |
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
|
905 |
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
|
906 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50419
diff
changeset
|
907 |
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
|
908 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
909 |
lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
910 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
911 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
912 |
lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
913 |
by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
914 |
(auto intro!: continuous_on_sgn continuous_on_id) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
915 |
|
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
|
916 |
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
|
917 |
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
|
918 |
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
|
919 |
shows "(\<lambda>x. - g x) \<in> borel_measurable M" |
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
920 |
by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_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
|
921 |
|
50003 | 922 |
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
|
923 |
fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" |
49774 | 924 |
assumes f: "f \<in> borel_measurable M" |
925 |
assumes g: "g \<in> borel_measurable M" |
|
926 |
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
927 |
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) |
49774 | 928 |
|
50003 | 929 |
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
|
930 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" |
49774 | 931 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
932 |
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" |
|
933 |
proof cases |
|
934 |
assume "finite S" |
|
935 |
thus ?thesis using assms by induct auto |
|
936 |
qed simp |
|
937 |
||
50003 | 938 |
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
|
939 |
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" |
49774 | 940 |
assumes f: "f \<in> borel_measurable M" |
941 |
assumes g: "g \<in> borel_measurable M" |
|
942 |
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
|
943 |
using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) |
49774 | 944 |
|
50003 | 945 |
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
|
946 |
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}" |
49774 | 947 |
assumes f: "f \<in> borel_measurable M" |
948 |
assumes g: "g \<in> borel_measurable M" |
|
949 |
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
950 |
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) |
51683
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
951 |
|
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
952 |
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
|
953 |
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
|
954 |
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
|
955 |
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
|
956 |
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
|
957 |
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
|
958 |
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
|
959 |
qed simp |
49774 | 960 |
|
50003 | 961 |
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
|
962 |
fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}" |
49774 | 963 |
assumes f: "f \<in> borel_measurable M" |
964 |
assumes g: "g \<in> borel_measurable M" |
|
965 |
shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M" |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
966 |
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) |
49774 | 967 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
968 |
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
|
969 |
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
|
970 |
assumes f: "f \<in> borel_measurable M" |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
971 |
assumes g: "g \<in> borel_measurable M" |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
972 |
shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M" |
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56212
diff
changeset
|
973 |
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
974 |
|
47694 | 975 |
lemma affine_borel_measurable_vector: |
38656 | 976 |
fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" |
977 |
assumes "f \<in> borel_measurable M" |
|
978 |
shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" |
|
979 |
proof (rule borel_measurableI) |
|
980 |
fix S :: "'x set" assume "open S" |
|
981 |
show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" |
|
982 |
proof cases |
|
983 |
assume "b \<noteq> 0" |
|
44537
c10485a6a7af
make HOL-Probability respect set/pred distinction
huffman
parents:
44282
diff
changeset
|
984 |
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
|
985 |
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
|
986 |
by (auto simp: algebra_simps) |
47694 | 987 |
hence "?S \<in> sets borel" by auto |
38656 | 988 |
moreover |
989 |
from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" |
|
990 |
apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) |
|
40859 | 991 |
ultimately show ?thesis using assms unfolding in_borel_measurable_borel |
38656 | 992 |
by auto |
993 |
qed simp |
|
994 |
qed |
|
995 |
||
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
996 |
lemma borel_measurable_const_scaleR[measurable (raw)]: |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
997 |
"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
|
998 |
using affine_borel_measurable_vector[of f M 0 b] by simp |
38656 | 999 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1000 |
lemma borel_measurable_const_add[measurable (raw)]: |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1001 |
"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
|
1002 |
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
|
1003 |
|
50003 | 1004 |
lemma borel_measurable_inverse[measurable (raw)]: |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1005 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" |
49774 | 1006 |
assumes f: "f \<in> borel_measurable M" |
35692 | 1007 |
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1008 |
apply (rule measurable_compose[OF f]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1009 |
apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1010 |
apply (auto intro!: continuous_on_inverse continuous_on_id) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1011 |
done |
35692 | 1012 |
|
50003 | 1013 |
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
|
1014 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1015 |
(\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1016 |
by (simp add: divide_inverse) |
38656 | 1017 |
|
50003 | 1018 |
lemma borel_measurable_max[measurable (raw)]: |
53216 | 1019 |
"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 | 1020 |
by (simp add: max_def) |
38656 | 1021 |
|
50003 | 1022 |
lemma borel_measurable_min[measurable (raw)]: |
53216 | 1023 |
"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 | 1024 |
by (simp add: min_def) |
38656 | 1025 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1026 |
lemma borel_measurable_Min[measurable (raw)]: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1027 |
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1028 |
proof (induct I rule: finite_induct) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1029 |
case (insert i I) then show ?case |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1030 |
by (cases "I = {}") auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1031 |
qed auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1032 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1033 |
lemma borel_measurable_Max[measurable (raw)]: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1034 |
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1035 |
proof (induct I rule: finite_induct) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1036 |
case (insert i I) then show ?case |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1037 |
by (cases "I = {}") auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1038 |
qed auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1039 |
|
50003 | 1040 |
lemma borel_measurable_abs[measurable (raw)]: |
1041 |
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" |
|
1042 |
unfolding abs_real_def by simp |
|
38656 | 1043 |
|
50003 | 1044 |
lemma borel_measurable_nth[measurable (raw)]: |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
1045 |
"(\<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
|
1046 |
by (simp add: cart_eq_inner_axis) |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41025
diff
changeset
|
1047 |
|
47694 | 1048 |
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
|
1049 |
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
|
1050 |
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
|
1051 |
assumes q: "convex_on A q" |
49774 | 1052 |
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
|
1053 |
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
|
1054 |
have "(\<lambda>x. if X x \<in> A then q (X x) else 0) \<in> borel_measurable M" (is "?qX") |
49774 | 1055 |
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
|
1056 |
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
|
1057 |
from this q show "continuous_on A q" |
42990
3706951a6421
composition of convex and measurable function is measurable
hoelzl
parents:
42950
diff
changeset
|
1058 |
by (rule convex_on_continuous) |
41830 | 1059 |
qed |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1060 |
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
|
1061 |
using X by (intro measurable_cong) auto |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1062 |
finally show ?thesis . |
41830 | 1063 |
qed |
1064 |
||
50003 | 1065 |
lemma borel_measurable_ln[measurable (raw)]: |
49774 | 1066 |
assumes f: "f \<in> borel_measurable M" |
1067 |
shows "(\<lambda>x. ln (f x)) \<in> borel_measurable M" |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1068 |
apply (rule measurable_compose[OF f]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1069 |
apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1070 |
apply (auto intro!: continuous_on_ln continuous_on_id) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57259
diff
changeset
|
1071 |
done |
41830 | 1072 |
|
50003 | 1073 |
lemma borel_measurable_log[measurable (raw)]: |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1074 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M" |
49774 | 1075 |
unfolding log_def by auto |
41830 | 1076 |
|
58656 | 1077 |
lemma borel_measurable_exp[measurable]: |
1078 |
"(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<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
|
1079 |
by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) |
50419 | 1080 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1081 |
lemma measurable_real_floor[measurable]: |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1082 |
"(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" |
47761 | 1083 |
proof - |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1084 |
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
|
1085 |
by (auto intro: floor_eq2) |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1086 |
then show ?thesis |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1087 |
by (auto simp: vimage_def measurable_count_space_eq2_countable) |
47761 | 1088 |
qed |
1089 |
||
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1090 |
lemma measurable_real_natfloor[measurable]: |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1091 |
"(natfloor :: real \<Rightarrow> nat) \<in> measurable borel (count_space UNIV)" |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1092 |
by (simp add: natfloor_def[abs_def]) |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1093 |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1094 |
lemma measurable_real_ceiling[measurable]: |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1095 |
"(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1096 |
unfolding ceiling_def[abs_def] by simp |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1097 |
|
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1098 |
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
|
1099 |
by simp |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1100 |
|
50003 | 1101 |
lemma borel_measurable_real_natfloor: |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1102 |
"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
|
1103 |
by simp |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1104 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1105 |
lemma borel_measurable_root [measurable]: "(root n) \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1106 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1107 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1108 |
lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1109 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1110 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1111 |
lemma borel_measurable_power [measurable (raw)]: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1112 |
fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1113 |
assumes f: "f \<in> borel_measurable M" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1114 |
shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1115 |
by (intro borel_measurable_continuous_on [OF _ f] continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1116 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1117 |
lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1118 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1119 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1120 |
lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1121 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1122 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1123 |
lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1124 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1125 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1126 |
lemma borel_measurable_sin [measurable]: "sin \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1127 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1128 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1129 |
lemma borel_measurable_cos [measurable]: "cos \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1130 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1131 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1132 |
lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1133 |
by (intro borel_measurable_continuous_on1 continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57138
diff
changeset
|
1134 |
|
57259
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1135 |
lemma borel_measurable_complex_iff: |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1136 |
"f \<in> borel_measurable M \<longleftrightarrow> |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1137 |
(\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M" |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1138 |
apply auto |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1139 |
apply (subst fun_complex_eq) |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1140 |
apply (intro borel_measurable_add) |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1141 |
apply auto |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1142 |
done |
3a448982a74a
add more derivative and continuity rules for complex-values functions
hoelzl
parents:
57235
diff
changeset
|
1143 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1144 |
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
|
1145 |
|
50003 | 1146 |
lemma borel_measurable_ereal[measurable (raw)]: |
43920 | 1147 |
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" |
49774 | 1148 |
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
|
1149 |
|
50003 | 1150 |
lemma borel_measurable_real_of_ereal[measurable (raw)]: |
49774 | 1151 |
fixes f :: "'a \<Rightarrow> ereal" |
1152 |
assumes f: "f \<in> borel_measurable M" |
|
1153 |
shows "(\<lambda>x. real (f x)) \<in> borel_measurable M" |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1154 |
apply (rule measurable_compose[OF f]) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1155 |
apply (rule borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"]) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1156 |
apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1157 |
done |
49774 | 1158 |
|
1159 |
lemma borel_measurable_ereal_cases: |
|
1160 |
fixes f :: "'a \<Rightarrow> ereal" |
|
1161 |
assumes f: "f \<in> borel_measurable M" |
|
1162 |
assumes H: "(\<lambda>x. H (ereal (real (f x)))) \<in> borel_measurable M" |
|
1163 |
shows "(\<lambda>x. H (f x)) \<in> borel_measurable M" |
|
1164 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1165 |
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 | 1166 |
{ 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
|
1167 |
with f H show ?thesis by simp |
47694 | 1168 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1169 |
|
49774 | 1170 |
lemma |
50003 | 1171 |
fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M" |
1172 |
shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" |
|
1173 |
and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M" |
|
1174 |
and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M" |
|
49774 | 1175 |
by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) |
1176 |
||
1177 |
lemma borel_measurable_uminus_eq_ereal[simp]: |
|
1178 |
"(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") |
|
1179 |
proof |
|
1180 |
assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp |
|
1181 |
qed auto |
|
1182 |
||
1183 |
lemma set_Collect_ereal2: |
|
1184 |
fixes f g :: "'a \<Rightarrow> ereal" |
|
1185 |
assumes f: "f \<in> borel_measurable M" |
|
1186 |
assumes g: "g \<in> borel_measurable M" |
|
1187 |
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
|
1188 |
"{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
|
1189 |
"{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
|
1190 |
"{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
|
1191 |
"{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel" |
49774 | 1192 |
shows "{x \<in> space M. H (f x) (g x)} \<in> sets M" |
1193 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1194 |
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
|
1195 |
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 | 1196 |
{ 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
|
1197 |
note * = this |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1198 |
from assms show ?thesis |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1199 |
by (subst *) (simp del: space_borel split del: split_if) |
49774 | 1200 |
qed |
1201 |
||
47694 | 1202 |
lemma borel_measurable_ereal_iff: |
43920 | 1203 |
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
|
1204 |
proof |
43920 | 1205 |
assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" |
1206 |
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
|
1207 |
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
|
1208 |
qed auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1209 |
|
59353
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59088
diff
changeset
|
1210 |
lemma borel_measurable_erealD[measurable_dest]: |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59088
diff
changeset
|
1211 |
"(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N" |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59088
diff
changeset
|
1212 |
unfolding borel_measurable_ereal_iff by simp |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59088
diff
changeset
|
1213 |
|
47694 | 1214 |
lemma borel_measurable_ereal_iff_real: |
43923 | 1215 |
fixes f :: "'a \<Rightarrow> ereal" |
1216 |
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
|
1217 |
((\<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
|
1218 |
proof safe |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1219 |
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
|
1220 |
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
|
1221 |
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 | 1222 |
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
|
1223 |
have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto |
43920 | 1224 |
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
|
1225 |
finally show "f \<in> borel_measurable M" . |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1226 |
qed simp_all |
41830 | 1227 |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1228 |
lemma borel_measurable_ereal_iff_Iio: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1229 |
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1230 |
by (auto simp: borel_Iio measurable_iff_measure_of) |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1231 |
|
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1232 |
lemma borel_measurable_ereal_iff_Ioi: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1233 |
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1234 |
by (auto simp: borel_Ioi measurable_iff_measure_of) |
35582 | 1235 |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1236 |
lemma vimage_sets_compl_iff: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1237 |
"f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1238 |
proof - |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1239 |
{ fix A assume "f -` A \<inter> space M \<in> sets M" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1240 |
moreover have "f -` (- A) \<inter> space M = space M - f -` A \<inter> space M" by auto |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1241 |
ultimately have "f -` (- A) \<inter> space M \<in> sets M" by auto } |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1242 |
from this[of A] this[of "-A"] show ?thesis |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1243 |
by (metis double_complement) |
49774 | 1244 |
qed |
1245 |
||
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1246 |
lemma borel_measurable_iff_Iic_ereal: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1247 |
"(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1248 |
unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp |
38656 | 1249 |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1250 |
lemma borel_measurable_iff_Ici_ereal: |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1251 |
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)" |
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1252 |
unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp |
38656 | 1253 |
|
49774 | 1254 |
lemma borel_measurable_ereal2: |
1255 |
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
|
1256 |
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
|
1257 |
assumes g: "g \<in> borel_measurable M" |
49774 | 1258 |
assumes H: "(\<lambda>x. H (ereal (real (f x))) (ereal (real (g x)))) \<in> borel_measurable M" |
1259 |
"(\<lambda>x. H (-\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M" |
|
1260 |
"(\<lambda>x. H (\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M" |
|
1261 |
"(\<lambda>x. H (ereal (real (f x))) (-\<infinity>)) \<in> borel_measurable M" |
|
1262 |
"(\<lambda>x. H (ereal (real (f x))) (\<infinity>)) \<in> borel_measurable M" |
|
1263 |
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
|
1264 |
proof - |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1265 |
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
|
1266 |
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 | 1267 |
{ 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
|
1268 |
note * = this |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1269 |
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
|
1270 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41969
diff
changeset
|
1271 |
|
49774 | 1272 |
lemma |
1273 |
fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M" |
|
1274 |
shows borel_measurable_ereal_eq_const: "{x\<in>space M. f x = c} \<in> sets M" |
|
1275 |
and borel_measurable_ereal_neq_const: "{x\<in>space M. f x \<noteq> c} \<in> sets M" |
|
1276 |
using f by auto |
|
38656 | 1277 |
|
50003 | 1278 |
lemma [measurable(raw)]: |
43920 | 1279 |
fixes f :: "'a \<Rightarrow> ereal" |
50003 | 1280 |
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
|
1281 |
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
|
1282 |
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
|
1283 |
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
|
1284 |
and borel_measurable_ereal_max: "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M" |
50003 | 1285 |
by (simp_all add: borel_measurable_ereal2 min_def max_def) |
49774 | 1286 |
|
50003 | 1287 |
lemma [measurable(raw)]: |
49774 | 1288 |
fixes f g :: "'a \<Rightarrow> ereal" |
1289 |
assumes "f \<in> borel_measurable M" |
|
1290 |
assumes "g \<in> borel_measurable M" |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1291 |
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
|
1292 |
and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M" |
50003 | 1293 |
using assms by (simp_all add: minus_ereal_def divide_ereal_def) |
38656 | 1294 |
|
50003 | 1295 |
lemma borel_measurable_ereal_setsum[measurable (raw)]: |
43920 | 1296 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" |
41096 | 1297 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
1298 |
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" |
|
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1299 |
using assms by (induction S rule: infinite_finite_induct) auto |
38656 | 1300 |
|
50003 | 1301 |
lemma borel_measurable_ereal_setprod[measurable (raw)]: |
43920 | 1302 |
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" |
38656 | 1303 |
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" |
41096 | 1304 |
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" |
59361
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents:
59353
diff
changeset
|
1305 |
using assms by (induction S rule: infinite_finite_induct) auto |
38656 | 1306 |
|
50003 | 1307 |
lemma [measurable (raw)]: |
43920 | 1308 |
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
|
1309 |
assumes "\<And>i. f i \<in> borel_measurable M" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1310 |
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
|
1311 |
and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M" |
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54775
diff
changeset
|
1312 |
unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto |
35692 | 1313 |
|
50104 | 1314 |
lemma sets_Collect_eventually_sequentially[measurable]: |
50003 | 1315 |
"(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M" |
1316 |
unfolding eventually_sequentially by simp |
|
1317 |
||
1318 |
lemma sets_Collect_ereal_convergent[measurable]: |
|
1319 |
fixes f :: "nat \<Rightarrow> 'a => ereal" |
|
1320 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
|
1321 |
shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M" |
|
1322 |
unfolding convergent_ereal by auto |
|
1323 |
||
1324 |
lemma borel_measurable_extreal_lim[measurable (raw)]: |
|
1325 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
|
1326 |
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" |
|
1327 |
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" |
|
1328 |
proof - |
|
1329 |
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 | 1330 |
by (simp add: lim_def convergent_def convergent_limsup_cl) |
50003 | 1331 |
then show ?thesis |
1332 |
by simp |
|
1333 |
qed |
|
1334 |
||
49774 | 1335 |
lemma borel_measurable_ereal_LIMSEQ: |
1336 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
|
1337 |
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" |
|
1338 |
and u: "\<And>i. u i \<in> borel_measurable M" |
|
1339 |
shows "u' \<in> borel_measurable M" |
|
47694 | 1340 |
proof - |
49774 | 1341 |
have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)" |
1342 |
using u' by (simp add: lim_imp_Liminf[symmetric]) |
|
50003 | 1343 |
with u show ?thesis by (simp cong: measurable_cong) |
47694 | 1344 |
qed |
1345 |
||
50003 | 1346 |
lemma borel_measurable_extreal_suminf[measurable (raw)]: |
43920 | 1347 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" |
50003 | 1348 |
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
|
1349 |
shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M" |
50003 | 1350 |
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp |
39092 | 1351 |
|
56994 | 1352 |
subsection {* LIMSEQ is borel measurable *} |
39092 | 1353 |
|
47694 | 1354 |
lemma borel_measurable_LIMSEQ: |
39092 | 1355 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" |
1356 |
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" |
|
1357 |
and u: "\<And>i. u i \<in> borel_measurable M" |
|
1358 |
shows "u' \<in> borel_measurable M" |
|
1359 |
proof - |
|
43920 | 1360 |
have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)" |
46731 | 1361 |
using u' by (simp add: lim_imp_Liminf) |
43920 | 1362 |
moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M" |
39092 | 1363 |
by auto |
43920 | 1364 |
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) |
39092 | 1365 |
qed |
1366 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1367 |
lemma borel_measurable_LIMSEQ_metric: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1368 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1369 |
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1370 |
assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) ----> g x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1371 |
shows "g \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1372 |
unfolding borel_eq_closed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1373 |
proof (safe intro!: measurable_measure_of) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1374 |
fix A :: "'b set" assume "closed A" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1375 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1376 |
have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1377 |
proof (rule borel_measurable_LIMSEQ) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1378 |
show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) ----> infdist (g x) A" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1379 |
by (intro tendsto_infdist lim) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1380 |
show "\<And>i. (\<lambda>x. infdist (f i x) A) \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1381 |
by (intro borel_measurable_continuous_on[where f="\<lambda>x. infdist x A"] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1382 |
continuous_at_imp_continuous_on ballI continuous_infdist isCont_ident) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1383 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1384 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1385 |
show "g -` A \<inter> space M \<in> sets M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1386 |
proof cases |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1387 |
assume "A \<noteq> {}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1388 |
then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1389 |
using `closed A` by (simp add: in_closed_iff_infdist_zero) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1390 |
then have "g -` A \<inter> space M = {x\<in>space M. infdist (g x) A = 0}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1391 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1392 |
also have "\<dots> \<in> sets M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1393 |
by measurable |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1394 |
finally show ?thesis . |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1395 |
qed simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1396 |
qed auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56371
diff
changeset
|
1397 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1398 |
lemma sets_Collect_Cauchy[measurable]: |
57036 | 1399 |
fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1400 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
49774 | 1401 |
shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M" |
57036 | 1402 |
unfolding metric_Cauchy_iff2 using f by auto |
49774 | 1403 |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1404 |
lemma borel_measurable_lim[measurable (raw)]: |
57036 | 1405 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1406 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
49774 | 1407 |
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" |
1408 |
proof - |
|
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1409 |
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
|
1410 |
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 | 1411 |
by (auto simp: lim_def convergent_eq_cauchy[symmetric]) |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1412 |
have "u' \<in> borel_measurable M" |
57036 | 1413 |
proof (rule borel_measurable_LIMSEQ_metric) |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1414 |
fix x |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1415 |
have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" |
49774 | 1416 |
by (cases "Cauchy (\<lambda>i. f i x)") |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1417 |
(auto simp add: convergent_eq_cauchy[symmetric] convergent_def) |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1418 |
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
|
1419 |
unfolding u'_def |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1420 |
by (rule convergent_LIMSEQ_iff[THEN iffD1]) |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1421 |
qed measurable |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1422 |
then show ?thesis |
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1423 |
unfolding * by measurable |
49774 | 1424 |
qed |
1425 |
||
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1426 |
lemma borel_measurable_suminf[measurable (raw)]: |
57036 | 1427 |
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
50001
diff
changeset
|
1428 |
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" |
49774 | 1429 |
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
|
1430 |
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp |
49774 | 1431 |
|
59000 | 1432 |
lemma borel_measurable_sup[measurable (raw)]: |
1433 |
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> |
|
1434 |
(\<lambda>x. sup (f x) (g x)::ereal) \<in> borel_measurable M" |
|
1435 |
by simp |
|
1436 |
||
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1437 |
(* Proof by Jeremy Avigad and Luke Serafin *) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1438 |
lemma isCont_borel: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1439 |
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1440 |
shows "{x. isCont f x} \<in> sets borel" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1441 |
proof - |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1442 |
let ?I = "\<lambda>j. inverse(real (Suc j))" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1443 |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1444 |
{ fix x |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1445 |
have "isCont f x = (\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1446 |
unfolding continuous_at_eps_delta |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1447 |
proof safe |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1448 |
fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1449 |
moreover have "0 < ?I i / 2" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1450 |
by simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1451 |
ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1452 |
by (metis dist_commute) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1453 |
then obtain j where j: "?I j < d" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1454 |
by (metis reals_Archimedean) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1455 |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1456 |
show "\<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1457 |
proof (safe intro!: exI[where x=j]) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1458 |
fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1459 |
have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1460 |
by (rule dist_triangle2) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1461 |
also have "\<dots> < ?I i / 2 + ?I i / 2" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1462 |
by (intro add_strict_mono d less_trans[OF _ j] *) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1463 |
also have "\<dots> \<le> ?I i" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1464 |
by (simp add: field_simps real_of_nat_Suc) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1465 |
finally show "dist (f y) (f z) \<le> ?I i" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1466 |
by simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1467 |
qed |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1468 |
next |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1469 |
fix e::real assume "0 < e" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1470 |
then obtain n where n: "?I n < e" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1471 |
by (metis reals_Archimedean) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1472 |
assume "\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1473 |
from this[THEN spec, of "Suc n"] |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1474 |
obtain j where j: "\<And>y z. dist x y < ?I j \<Longrightarrow> dist x z < ?I j \<Longrightarrow> dist (f y) (f z) \<le> ?I (Suc n)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1475 |
by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1476 |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1477 |
show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1478 |
proof (safe intro!: exI[of _ "?I j"]) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1479 |
fix y assume "dist y x < ?I j" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1480 |
then have "dist (f y) (f x) \<le> ?I (Suc n)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1481 |
by (intro j) (auto simp: dist_commute) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1482 |
also have "?I (Suc n) < ?I n" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1483 |
by simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1484 |
also note n |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1485 |
finally show "dist (f y) (f x) < e" . |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1486 |
qed simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1487 |
qed } |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1488 |
note * = this |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1489 |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1490 |
have **: "\<And>e y. open {x. dist x y < e}" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1491 |
using open_ball by (simp_all add: ball_def dist_commute) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1492 |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1493 |
have "{x\<in>space borel. isCont f x } \<in> sets borel" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1494 |
unfolding * |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1495 |
apply (intro sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1496 |
apply (simp add: Collect_all_eq) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1497 |
apply (intro borel_closed closed_INT ballI closed_Collect_imp open_Collect_conj **) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1498 |
apply auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1499 |
done |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1500 |
then show ?thesis |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1501 |
by simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1502 |
qed |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57275
diff
changeset
|
1503 |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
1504 |
no_notation |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
1505 |
eucl_less (infix "<e" 50) |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54230
diff
changeset
|
1506 |
|
51683
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents:
51478
diff
changeset
|
1507 |
end |