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