| author | wenzelm | 
| Thu, 13 Dec 2018 21:36:09 +0100 | |
| changeset 69467 | e8893c893241 | 
| parent 69260 | 0a9688695a1b | 
| child 69517 | dc20f278e8f3 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Borel_Space.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 38656 | 5 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 6 | section%important \<open>Borel spaces\<close> | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 7 | |
| 40859 | 8 | theory Borel_Space | 
| 50387 | 9 | imports | 
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63566diff
changeset | 10 | Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 11 | begin | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 12 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 13 | lemma%unimportant sets_Collect_eventually_sequentially[measurable]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 14 |   "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 15 | unfolding eventually_sequentially by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 16 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 17 | lemma%unimportant topological_basis_trivial: "topological_basis {A. open A}"
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 18 | 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 | 19 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 20 | lemma%important open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
 | 
| 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 21 | proof%unimportant - | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 22 |   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 | 23 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 24 | then show ?thesis | 
| 62372 | 25 | by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 26 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 27 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 28 | definition%important "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s" | 
| 62083 | 29 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 30 | lemma%unimportant mono_onI: | 
| 62083 | 31 | "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A" | 
| 32 | unfolding mono_on_def by simp | |
| 33 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 34 | lemma%unimportant mono_onD: | 
| 62083 | 35 | "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s" | 
| 36 | unfolding mono_on_def by simp | |
| 37 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 38 | lemma%unimportant mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A" | 
| 62083 | 39 | unfolding mono_def mono_on_def by auto | 
| 40 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 41 | lemma%unimportant mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B" | 
| 62083 | 42 | unfolding mono_on_def by auto | 
| 43 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 44 | definition%important "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s" | 
| 62083 | 45 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 46 | lemma%unimportant strict_mono_onI: | 
| 62083 | 47 | "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A" | 
| 48 | unfolding strict_mono_on_def by simp | |
| 49 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 50 | lemma%unimportant strict_mono_onD: | 
| 62083 | 51 | "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s" | 
| 52 | unfolding strict_mono_on_def by simp | |
| 53 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 54 | lemma%unimportant mono_on_greaterD: | 
| 62083 | 55 | assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)" | 
| 56 | shows "x > y" | |
| 57 | proof (rule ccontr) | |
| 58 | assume "\<not>x > y" | |
| 59 | hence "x \<le> y" by (simp add: not_less) | |
| 60 | from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD) | |
| 61 | with assms(4) show False by simp | |
| 62 | qed | |
| 63 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 64 | lemma%unimportant strict_mono_inv: | 
| 62083 | 65 |   fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
 | 
| 66 | assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x" | |
| 67 | shows "strict_mono g" | |
| 68 | proof | |
| 69 | fix x y :: 'b assume "x < y" | |
| 70 | from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast | |
| 71 | with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less) | |
| 72 | with inv show "g x < g y" by simp | |
| 73 | qed | |
| 74 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 75 | lemma%unimportant strict_mono_on_imp_inj_on: | 
| 62083 | 76 | assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A" | 
| 77 | shows "inj_on f A" | |
| 78 | proof (rule inj_onI) | |
| 79 | fix x y assume "x \<in> A" "y \<in> A" "f x = f y" | |
| 80 | thus "x = y" | |
| 81 | by (cases x y rule: linorder_cases) | |
| 62372 | 82 | (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) | 
| 62083 | 83 | qed | 
| 84 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 85 | lemma%unimportant strict_mono_on_leD: | 
| 62083 | 86 | assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y" | 
| 87 | shows "f x \<le> f y" | |
| 88 | proof (insert le_less_linear[of y x], elim disjE) | |
| 89 | assume "x < y" | |
| 90 | with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all | |
| 91 | thus ?thesis by (rule less_imp_le) | |
| 92 | qed (insert assms, simp) | |
| 93 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 94 | lemma%unimportant strict_mono_on_eqD: | 
| 62083 | 95 | fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)" | 
| 96 | assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A" | |
| 97 | shows "y = x" | |
| 98 | using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD) | |
| 99 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 100 | lemma%important mono_on_imp_deriv_nonneg: | 
| 62083 | 101 | assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)" | 
| 102 | assumes "x \<in> interior A" | |
| 103 | shows "D \<ge> 0" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 104 | proof%unimportant (rule tendsto_lowerbound) | 
| 62083 | 105 | let ?A' = "(\<lambda>y. y - x) ` interior A" | 
| 106 | from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)" | |
| 107 | by (simp add: field_has_derivative_at has_field_derivative_def) | |
| 108 | from mono have mono': "mono_on f (interior A)" by (rule mono_on_subset) (rule interior_subset) | |
| 109 | ||
| 110 | show "eventually (\<lambda>h. (f (x + h) - f x) / h \<ge> 0) (at 0)" | |
| 111 | proof (subst eventually_at_topological, intro exI conjI ballI impI) | |
| 112 | have "open (interior A)" by simp | |
| 67399 | 113 | hence "open ((+) (-x) ` interior A)" by (rule open_translation) | 
| 114 | also have "((+) (-x) ` interior A) = ?A'" by auto | |
| 62083 | 115 | finally show "open ?A'" . | 
| 116 | next | |
| 117 | from \<open>x \<in> interior A\<close> show "0 \<in> ?A'" by auto | |
| 118 | next | |
| 119 | fix h assume "h \<in> ?A'" | |
| 120 | hence "x + h \<in> interior A" by auto | |
| 121 | with mono' and \<open>x \<in> interior A\<close> show "(f (x + h) - f x) / h \<ge> 0" | |
| 122 | by (cases h rule: linorder_cases[of _ 0]) | |
| 123 | (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps) | |
| 124 | qed | |
| 125 | qed simp | |
| 126 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 127 | lemma%unimportant strict_mono_on_imp_mono_on: | 
| 62083 | 128 | "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A" | 
| 129 | by (rule mono_onI, rule strict_mono_on_leD) | |
| 130 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 131 | lemma%important mono_on_ctble_discont: | 
| 62083 | 132 | fixes f :: "real \<Rightarrow> real" | 
| 133 | fixes A :: "real set" | |
| 134 | assumes "mono_on f A" | |
| 135 |   shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 136 | proof%unimportant - | 
| 62083 | 137 | have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" | 
| 63167 | 138 | using \<open>mono_on f A\<close> by (simp add: mono_on_def) | 
| 62083 | 139 |   have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
 | 
| 140 | (fst q = 0 \<and> of_rat (snd q) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd q))) \<or> | |
| 141 | (fst q = 1 \<and> of_rat (snd q) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd q)))" | |
| 142 | proof (clarsimp simp del: One_nat_def) | |
| 143 | fix a assume "a \<in> A" assume "\<not> continuous (at a within A) f" | |
| 144 | thus "\<exists>q1 q2. | |
| 145 | q1 = 0 \<and> real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2) \<or> | |
| 146 | q1 = 1 \<and> f a < real_of_rat q2 \<and> (\<forall>x\<in>A. a < x \<longrightarrow> real_of_rat q2 < f x)" | |
| 147 | proof (auto simp add: continuous_within order_tendsto_iff eventually_at) | |
| 148 | fix l assume "l < f a" | |
| 149 | then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a" | |
| 150 | using of_rat_dense by blast | |
| 151 | assume * [rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> l < f x" | |
| 152 | from q2 have "real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2)" | |
| 153 | proof auto | |
| 154 | fix x assume "x \<in> A" "x < a" | |
| 155 | with q2 *[of "a - x"] show "f x < real_of_rat q2" | |
| 156 | apply (auto simp add: dist_real_def not_less) | |
| 157 | apply (subgoal_tac "f x \<le> f xa") | |
| 158 | by (auto intro: mono) | |
| 62372 | 159 | qed | 
| 62083 | 160 | thus ?thesis by auto | 
| 161 | next | |
| 162 | fix u assume "u > f a" | |
| 163 | then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u" | |
| 164 | using of_rat_dense by blast | |
| 165 | assume *[rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> u > f x" | |
| 166 | from q2 have "real_of_rat q2 > f a \<and> (\<forall>x\<in>A. x > a \<longrightarrow> f x > real_of_rat q2)" | |
| 167 | proof auto | |
| 168 | fix x assume "x \<in> A" "x > a" | |
| 169 | with q2 *[of "x - a"] show "f x > real_of_rat q2" | |
| 170 | apply (auto simp add: dist_real_def) | |
| 171 | apply (subgoal_tac "f x \<ge> f xa") | |
| 172 | by (auto intro: mono) | |
| 62372 | 173 | qed | 
| 62083 | 174 | thus ?thesis by auto | 
| 175 | qed | |
| 176 | qed | |
| 62372 | 177 |   hence "\<exists>g :: real \<Rightarrow> nat \<times> rat . \<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}.
 | 
| 62083 | 178 | (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd (g a)))) | | 
| 179 | (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd (g a))))" | |
| 180 | by (rule bchoice) | |
| 181 | then guess g .. | |
| 182 | hence g: "\<And>a x. a \<in> A \<Longrightarrow> \<not> continuous (at a within A) f \<Longrightarrow> x \<in> A \<Longrightarrow> | |
| 183 | (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (x < a \<longrightarrow> f x < of_rat (snd (g a)))) | | |
| 184 | (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (x > a \<longrightarrow> f x > of_rat (snd (g a))))" | |
| 185 | by auto | |
| 186 |   have "inj_on g {a\<in>A. \<not> continuous (at a within A) f}"
 | |
| 187 | proof (auto simp add: inj_on_def) | |
| 188 | fix w z | |
| 189 | assume 1: "w \<in> A" and 2: "\<not> continuous (at w within A) f" and | |
| 190 | 3: "z \<in> A" and 4: "\<not> continuous (at z within A) f" and | |
| 191 | 5: "g w = g z" | |
| 62372 | 192 | from g [OF 1 2 3] g [OF 3 4 1] 5 | 
| 62083 | 193 | show "w = z" by auto | 
| 194 | qed | |
| 62372 | 195 | thus ?thesis | 
| 196 | by (rule countableI') | |
| 62083 | 197 | qed | 
| 198 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 199 | lemma%important mono_on_ctble_discont_open: | 
| 62083 | 200 | fixes f :: "real \<Rightarrow> real" | 
| 201 | fixes A :: "real set" | |
| 202 | assumes "open A" "mono_on f A" | |
| 203 |   shows "countable {a\<in>A. \<not>isCont f a}"
 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 204 | proof%unimportant - | 
| 62083 | 205 |   have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
 | 
| 63167 | 206 | by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>]) | 
| 62083 | 207 | thus ?thesis | 
| 208 | apply (elim ssubst) | |
| 209 | by (rule mono_on_ctble_discont, rule assms) | |
| 210 | qed | |
| 211 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 212 | lemma%important mono_ctble_discont: | 
| 62083 | 213 | fixes f :: "real \<Rightarrow> real" | 
| 214 | assumes "mono f" | |
| 215 |   shows "countable {a. \<not> isCont f a}"
 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 216 | using%unimportant assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto | 
| 62083 | 217 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 218 | lemma%important has_real_derivative_imp_continuous_on: | 
| 62083 | 219 | assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)" | 
| 220 | shows "continuous_on A f" | |
| 221 | apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def) | |
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 222 | using assms differentiable_at_withinI real_differentiable_def by blast | 
| 62083 | 223 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 224 | lemma%important closure_contains_Sup: | 
| 62083 | 225 | fixes S :: "real set" | 
| 226 |   assumes "S \<noteq> {}" "bdd_above S"
 | |
| 227 | shows "Sup S \<in> closure S" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 228 | proof%unimportant - | 
| 62372 | 229 | have "Inf (uminus ` S) \<in> closure (uminus ` S)" | 
| 62083 | 230 | using assms by (intro closure_contains_Inf) auto | 
| 231 | also have "Inf (uminus ` S) = -Sup S" by (simp add: Inf_real_def) | |
| 232 | also have "closure (uminus ` S) = uminus ` closure S" | |
| 233 | by (rule sym, intro closure_injective_linear_image) (auto intro: linearI) | |
| 234 | finally show ?thesis by auto | |
| 235 | qed | |
| 236 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 237 | lemma%unimportant closed_contains_Sup: | 
| 62083 | 238 | fixes S :: "real set" | 
| 239 |   shows "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> closed S \<Longrightarrow> Sup S \<in> S"
 | |
| 240 | by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) | |
| 241 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 242 | lemma%unimportant closed_subset_contains_Sup: | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67399diff
changeset | 243 | fixes A C :: "real set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67399diff
changeset | 244 |   shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> Sup A \<in> C"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67399diff
changeset | 245 | by (metis closure_contains_Sup closure_minimal subset_eq) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67399diff
changeset | 246 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 247 | lemma%important deriv_nonneg_imp_mono: | 
| 62083 | 248 |   assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
 | 
| 249 |   assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
 | |
| 250 | assumes ab: "a \<le> b" | |
| 251 | shows "g a \<le> g b" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 252 | proof%unimportant (cases "a < b") | 
| 62083 | 253 | assume "a < b" | 
| 68635 | 254 | from deriv have "\<And>x. \<lbrakk>x \<ge> a; x \<le> b\<rbrakk> \<Longrightarrow> (g has_real_derivative g' x) (at x)" by simp | 
| 255 | with MVT2[OF \<open>a < b\<close>] and deriv | |
| 62083 | 256 | obtain \<xi> where \<xi>_ab: "\<xi> > a" "\<xi> < b" and g_ab: "g b - g a = (b - a) * g' \<xi>" by blast | 
| 257 | from \<xi>_ab ab nonneg have "(b - a) * g' \<xi> \<ge> 0" by simp | |
| 258 | with g_ab show ?thesis by simp | |
| 259 | qed (insert ab, simp) | |
| 260 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 261 | lemma%important continuous_interval_vimage_Int: | 
| 62083 | 262 |   assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
 | 
| 263 |   assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
 | |
| 264 |   obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 265 | proof%unimportant- | 
| 63040 | 266 |   let ?A = "{a..b} \<inter> g -` {c..d}"
 | 
| 267 | from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5) | |
| 268 | obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto | |
| 269 | from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5) | |
| 270 | obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto | |
| 271 |   hence [simp]: "?A \<noteq> {}" by blast
 | |
| 62083 | 272 | |
| 63040 | 273 | define c' where "c' = Inf ?A" | 
| 274 | define d' where "d' = Sup ?A" | |
| 275 |   have "?A \<subseteq> {c'..d'}" unfolding c'_def d'_def
 | |
| 276 | by (intro subsetI) (auto intro: cInf_lower cSup_upper) | |
| 277 | moreover from assms have "closed ?A" | |
| 278 |     using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp
 | |
| 279 | hence c'd'_in_set: "c' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def | |
| 280 | by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ | |
| 281 |   hence "{c'..d'} \<subseteq> ?A" using assms
 | |
| 282 | by (intro subsetI) | |
| 283 | (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] | |
| 284 | intro!: mono) | |
| 285 | moreover have "c' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto | |
| 286 | moreover have "g c' \<le> c" "g d' \<ge> d" | |
| 287 | apply (insert c'' d'' c'd'_in_set) | |
| 288 | apply (subst c''(2)[symmetric]) | |
| 289 | apply (auto simp: c'_def intro!: mono cInf_lower c'') [] | |
| 290 | apply (subst d''(2)[symmetric]) | |
| 291 | apply (auto simp: d'_def intro!: mono cSup_upper d'') [] | |
| 292 | done | |
| 293 | with c'd'_in_set have "g c' = c" "g d' = d" by auto | |
| 294 | ultimately show ?thesis using that by blast | |
| 62083 | 295 | qed | 
| 296 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 297 | subsection%important \<open>Generic Borel spaces\<close> | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 298 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 299 | definition%important (in topological_space) borel :: "'a measure" where | 
| 47694 | 300 |   "borel = sigma UNIV {S. open S}"
 | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 301 | |
| 47694 | 302 | abbreviation "borel_measurable M \<equiv> measurable M borel" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 303 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 304 | lemma%important in_borel_measurable: | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 305 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 47694 | 306 |     (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 307 | by%unimportant (auto simp add: measurable_def borel_def) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 308 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 309 | lemma%important in_borel_measurable_borel: | 
| 38656 | 310 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 40859 | 311 | (\<forall>S \<in> sets borel. | 
| 38656 | 312 | f -` S \<inter> space M \<in> sets M)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 313 | by%unimportant (auto simp add: measurable_def borel_def) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 314 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 315 | lemma%unimportant space_borel[simp]: "space borel = UNIV" | 
| 40859 | 316 | unfolding borel_def by auto | 
| 38656 | 317 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 318 | lemma%unimportant space_in_borel[measurable]: "UNIV \<in> sets borel" | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 319 | unfolding borel_def by auto | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 320 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 321 | lemma%unimportant sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
 | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 322 | 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 | 323 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 324 | lemma%unimportant measurable_sets_borel: | 
| 62083 | 325 | "\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel" | 
| 326 | by (drule (1) measurable_sets) simp | |
| 327 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 328 | lemma%unimportant 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 | 329 | unfolding borel_def pred_def by auto | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 330 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 331 | lemma%unimportant borel_open[measurable (raw generic)]: | 
| 40859 | 332 | assumes "open A" shows "A \<in> sets borel" | 
| 38656 | 333 | proof - | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 334 |   have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
 | 
| 47694 | 335 | thus ?thesis unfolding borel_def by auto | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 336 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 337 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 338 | lemma%unimportant borel_closed[measurable (raw generic)]: | 
| 40859 | 339 | assumes "closed A" shows "A \<in> sets borel" | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 340 | proof - | 
| 40859 | 341 | have "space borel - (- A) \<in> sets borel" | 
| 342 | using assms unfolding closed_def by (blast intro: borel_open) | |
| 38656 | 343 | thus ?thesis by simp | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 344 | qed | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 345 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 346 | lemma%unimportant borel_singleton[measurable]: | 
| 50003 | 347 | "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 | 348 | unfolding insert_def by (rule sets.Un) auto | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 349 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 350 | lemma%unimportant sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
 | 
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 351 | proof - | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 352 |   have "(\<Union>a\<in>A. {a}) \<in> sets borel" for A :: "'a set"
 | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 353 | by (intro sets.countable_UN') auto | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 354 | then show ?thesis | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 355 | by auto | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 356 | qed | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64287diff
changeset | 357 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 358 | lemma%unimportant 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 | 359 | unfolding Compl_eq_Diff_UNIV by simp | 
| 41830 | 360 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 361 | lemma%unimportant borel_measurable_vimage: | 
| 38656 | 362 | fixes f :: "'a \<Rightarrow> 'x::t2_space" | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 363 | assumes borel[measurable]: "f \<in> borel_measurable M" | 
| 38656 | 364 |   shows "f -` {x} \<inter> space M \<in> sets M"
 | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 365 | by simp | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 366 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 367 | lemma%unimportant borel_measurableI: | 
| 61076 | 368 | fixes f :: "'a \<Rightarrow> 'x::topological_space" | 
| 38656 | 369 | assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" | 
| 370 | shows "f \<in> borel_measurable M" | |
| 40859 | 371 | unfolding borel_def | 
| 47694 | 372 | proof (rule measurable_measure_of, simp_all) | 
| 44537 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 huffman parents: 
44282diff
changeset | 373 | 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 | 374 | using assms[of S] by simp | 
| 40859 | 375 | qed | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 376 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 377 | lemma%unimportant borel_measurable_const: | 
| 38656 | 378 | "(\<lambda>x. c) \<in> borel_measurable M" | 
| 47694 | 379 | by auto | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 380 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 381 | lemma%unimportant borel_measurable_indicator: | 
| 38656 | 382 | assumes A: "A \<in> sets M" | 
| 383 | shows "indicator A \<in> borel_measurable M" | |
| 46905 | 384 | unfolding indicator_def [abs_def] using A | 
| 47694 | 385 | by (auto intro!: measurable_If_set) | 
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: diff
changeset | 386 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 387 | lemma%unimportant borel_measurable_count_space[measurable (raw)]: | 
| 50096 | 388 | "f \<in> borel_measurable (count_space S)" | 
| 389 | unfolding measurable_def by auto | |
| 390 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 391 | lemma%unimportant borel_measurable_indicator'[measurable (raw)]: | 
| 50096 | 392 |   assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
 | 
| 393 | 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 | 394 | unfolding indicator_def[abs_def] | 
| 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49774diff
changeset | 395 | by (auto intro!: measurable_If) | 
| 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49774diff
changeset | 396 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 397 | lemma%important borel_measurable_indicator_iff: | 
| 40859 | 398 |   "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
 | 
| 399 | (is "?I \<in> borel_measurable M \<longleftrightarrow> _") | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 400 | proof%unimportant | 
| 40859 | 401 | assume "?I \<in> borel_measurable M" | 
| 402 |   then have "?I -` {1} \<inter> space M \<in> sets M"
 | |
| 403 | unfolding measurable_def by auto | |
| 404 |   also have "?I -` {1} \<inter> space M = A \<inter> space M"
 | |
| 46905 | 405 | unfolding indicator_def [abs_def] by auto | 
| 40859 | 406 | finally show "A \<inter> space M \<in> sets M" . | 
| 407 | next | |
| 408 | assume "A \<inter> space M \<in> sets M" | |
| 409 | moreover have "?I \<in> borel_measurable M \<longleftrightarrow> | |
| 410 | (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M" | |
| 411 | by (intro measurable_cong) (auto simp: indicator_def) | |
| 412 | ultimately show "?I \<in> borel_measurable M" by auto | |
| 413 | qed | |
| 414 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 415 | lemma%unimportant borel_measurable_subalgebra: | 
| 41545 | 416 | assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N" | 
| 39092 | 417 | shows "f \<in> borel_measurable M" | 
| 418 | using assms unfolding measurable_def by auto | |
| 419 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 420 | lemma%unimportant borel_measurable_restrict_space_iff_ereal: | 
| 57137 | 421 | fixes f :: "'a \<Rightarrow> ereal" | 
| 422 | assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" | |
| 423 | shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> | |
| 424 | (\<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 | 425 | by (subst measurable_restrict_space_iff) | 
| 63566 | 426 | (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong) | 
| 57137 | 427 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 428 | lemma%unimportant borel_measurable_restrict_space_iff_ennreal: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 429 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 430 | assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 431 | shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 432 | (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 433 | by (subst measurable_restrict_space_iff) | 
| 63566 | 434 | (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 435 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 436 | lemma%unimportant borel_measurable_restrict_space_iff: | 
| 57137 | 437 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 438 | assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" | |
| 439 | shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> | |
| 440 | (\<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 | 441 | by (subst measurable_restrict_space_iff) | 
| 63566 | 442 | (auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps | 
| 443 | cong del: if_weak_cong) | |
| 57138 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 444 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 445 | lemma%unimportant cbox_borel[measurable]: "cbox a b \<in> sets borel" | 
| 57138 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 446 | by (auto intro: borel_closed) | 
| 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 447 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 448 | lemma%unimportant box_borel[measurable]: "box a b \<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 | 449 | 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 | 450 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 451 | lemma%unimportant borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel" | 
| 57138 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 452 | by (auto intro: borel_closed dest!: compact_imp_closed) | 
| 57137 | 453 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 454 | lemma%unimportant borel_sigma_sets_subset: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 455 | "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 456 | using sets.sigma_sets_subset[of A borel] by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 457 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 458 | lemma%important borel_eq_sigmaI1: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 459 | fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 460 | assumes borel_eq: "borel = sigma UNIV X" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 461 | assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 462 | assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 463 | shows "borel = sigma UNIV (F ` A)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 464 | unfolding borel_def | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 465 | proof%unimportant (intro sigma_eqI antisym) | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 466 |   have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 467 | unfolding borel_def by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 468 | also have "\<dots> = sigma_sets UNIV X" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 469 | unfolding borel_eq by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 470 | also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 471 | using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 472 |   finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 473 |   show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 474 | unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 475 | qed auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 476 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 477 | lemma%unimportant borel_eq_sigmaI2: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 478 | fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 479 | and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 480 | assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 481 | assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 482 | assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 483 | shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 484 | using assms | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 485 | by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 486 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 487 | lemma%unimportant borel_eq_sigmaI3: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 488 | fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 489 | assumes borel_eq: "borel = sigma UNIV X" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 490 | assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 491 | assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 492 | shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 493 | using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 494 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 495 | lemma%unimportant borel_eq_sigmaI4: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 496 | fixes F :: "'i \<Rightarrow> 'a::topological_space set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 497 | and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 498 | assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 499 | assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 500 | assumes F: "\<And>i. F i \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 501 | shows "borel = sigma UNIV (range F)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 502 | using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 503 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 504 | lemma%unimportant borel_eq_sigmaI5: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 505 | fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 506 | assumes borel_eq: "borel = sigma UNIV (range G)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 507 | assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 508 | assumes F: "\<And>i j. F i j \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 509 | shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 510 | using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 511 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 512 | lemma%important second_countable_borel_measurable: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 513 | 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 | 514 | 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 | 515 | 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 | 516 | unfolding borel_def | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 517 | proof%unimportant (intro sigma_eqI sigma_sets_eqI) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 518 | 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 | 519 | 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 | 520 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 521 | 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 | 522 | 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 | 523 | by (auto simp: eq) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 524 | 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 | 525 | proof induction | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 526 | case (UN K) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 527 | 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 | 528 | unfolding eq by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 529 | 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 | 530 | 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 | 531 | 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 | 532 | 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 | 533 | by metis | 
| 63040 | 534 | define U where "U = (\<Union>k\<in>K. m k)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 535 | with m have "countable U" | 
| 61808 | 536 | by (intro countable_subset[OF _ \<open>countable B\<close>]) auto | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 537 | 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 | 538 | 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 | 539 | 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 | 540 | 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 | 541 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 542 | 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 | 543 | 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 | 544 | 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 | 545 | by metis | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 546 | 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 | 547 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 548 | then have "\<Union>K = (\<Union>b\<in>U. u b)" | 
| 61808 | 549 | unfolding \<open>\<Union>U = \<Union>K\<close> by auto | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 550 | also have "\<dots> \<in> sigma_sets UNIV X" | 
| 61808 | 551 | using u UN by (intro X.countable_UN' \<open>countable U\<close>) auto | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 552 | 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 | 553 | qed auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 554 | 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 | 555 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 556 | lemma%unimportant borel_eq_closed: "borel = sigma UNIV (Collect closed)" | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 557 | unfolding borel_def | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 558 | proof (intro sigma_eqI sigma_sets_eqI, safe) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 559 | fix x :: "'a set" assume "open x" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 560 | hence "x = UNIV - (UNIV - x)" by auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 561 | also have "\<dots> \<in> sigma_sets UNIV (Collect closed)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 562 | by (force intro: sigma_sets.Compl simp: \<open>open x\<close>) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 563 | finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 564 | next | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 565 | fix x :: "'a set" assume "closed x" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 566 | hence "x = UNIV - (UNIV - x)" by auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 567 | also have "\<dots> \<in> sigma_sets UNIV (Collect open)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 568 | by (force intro: sigma_sets.Compl simp: \<open>closed x\<close>) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 569 | finally show "x \<in> sigma_sets UNIV (Collect open)" by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 570 | qed simp_all | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 571 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 572 | lemma%important borel_eq_countable_basis: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 573 | fixes B::"'a::topological_space set set" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 574 | assumes "countable B" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 575 | assumes "topological_basis B" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 576 | shows "borel = sigma UNIV B" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 577 | unfolding borel_def | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 578 | proof%unimportant (intro sigma_eqI sigma_sets_eqI, safe) | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 579 | interpret countable_basis using assms by unfold_locales | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 580 | fix X::"'a set" assume "open X" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 581 | from open_countable_basisE[OF this] guess B' . note B' = this | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 582 | then show "X \<in> sigma_sets UNIV B" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 583 | by (blast intro: sigma_sets_UNION \<open>countable B\<close> countable_subset) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 584 | next | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 585 | fix b assume "b \<in> B" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 586 | hence "open b" by (rule topological_basis_open[OF assms(2)]) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 587 | thus "b \<in> sigma_sets UNIV (Collect open)" by auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 588 | qed simp_all | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 589 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 590 | lemma%unimportant borel_measurable_continuous_on_restrict: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 591 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 592 | assumes f: "continuous_on A f" | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 593 | 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 | 594 | proof (rule borel_measurableI) | 
| 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 595 | fix S :: "'b set" assume "open S" | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 596 | 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 | 597 | by (metis continuous_on_open_invariant) | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 598 | 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 | 599 | by (force simp add: sets_restrict_space space_restrict_space) | 
| 57137 | 600 | qed | 
| 601 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 602 | lemma%unimportant borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel" | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 603 | by (drule borel_measurable_continuous_on_restrict) simp | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 604 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 605 | lemma%unimportant borel_measurable_continuous_on_if: | 
| 59415 | 606 | "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow> | 
| 607 | (\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel" | |
| 608 | by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq | |
| 609 | intro!: borel_measurable_continuous_on_restrict) | |
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 610 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 611 | lemma%unimportant borel_measurable_continuous_countable_exceptions: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57259diff
changeset | 612 | 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 | 613 | 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 | 614 | 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 | 615 | 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 | 616 | 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 | 617 | 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 | 618 | 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 | 619 | 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 | 620 | 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 | 621 | qed auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57259diff
changeset | 622 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 623 | lemma%unimportant borel_measurable_continuous_on: | 
| 57138 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 624 | 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 | 625 | 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 | 626 | 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 | 627 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 628 | lemma%unimportant borel_measurable_continuous_on_indicator: | 
| 57138 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 hoelzl parents: 
57137diff
changeset | 629 | fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 59415 | 630 | shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel" | 
| 631 | by (subst borel_measurable_restrict_space_iff[symmetric]) | |
| 632 | (auto intro: borel_measurable_continuous_on_restrict) | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 633 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 634 | lemma%important borel_measurable_Pair[measurable (raw)]: | 
| 50881 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 hoelzl parents: 
50526diff
changeset | 635 | 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 | 636 | 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 | 637 | 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 | 638 | shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 639 | proof%unimportant (subst borel_eq_countable_basis) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 640 | 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 | 641 | 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 | 642 | 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 | 643 | 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 | 644 | by (auto intro!: countable_basis topological_basis_prod is_basis) | 
| 38656 | 645 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 646 | 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 | 647 | 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 | 648 | 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 | 649 | 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 | 650 | 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 | 651 | 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 | 652 | 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 | 653 | 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 | 654 | 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 | 655 | 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 | 656 | 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 | 657 | qed auto | 
| 39087 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 658 | qed | 
| 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 hoelzl parents: 
39083diff
changeset | 659 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 660 | lemma%important borel_measurable_continuous_Pair: | 
| 50881 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 hoelzl parents: 
50526diff
changeset | 661 | fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology" | 
| 50003 | 662 | assumes [measurable]: "f \<in> borel_measurable M" | 
| 663 | assumes [measurable]: "g \<in> borel_measurable M" | |
| 49774 | 664 | assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))" | 
| 665 | shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 666 | proof%unimportant - | 
| 49774 | 667 | have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto | 
| 668 | show ?thesis | |
| 669 | unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto | |
| 670 | qed | |
| 671 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 672 | subsection%important \<open>Borel spaces on order topologies\<close> | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 673 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 674 | lemma%unimportant [measurable]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 675 | fixes a b :: "'a::linorder_topology" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 676 |   shows lessThan_borel: "{..< a} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 677 |     and greaterThan_borel: "{a <..} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 678 |     and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 679 |     and atMost_borel: "{..a} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 680 |     and atLeast_borel: "{a..} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 681 |     and atLeastAtMost_borel: "{a..b} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 682 |     and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 683 |     and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 684 | unfolding greaterThanAtMost_def atLeastLessThan_def | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 685 | by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 686 | closed_atMost closed_atLeast closed_atLeastAtMost)+ | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 687 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 688 | lemma%unimportant borel_Iio: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 689 |   "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 | 690 | 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 | 691 | 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 | 692 | 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 | 693 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 694 | 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 | 695 | 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 | 696 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 697 | 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 | 698 |   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 | 699 | by blast | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 700 | 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 | 701 | proof | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 702 |     assume A: "A = {y <..}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 703 | show ?thesis | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 704 | proof cases | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 705 | 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 | 706 |       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 | 707 | 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 | 708 |       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 | 709 | 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 | 710 | 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 | 711 | 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 | 712 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 713 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 714 | 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 | 715 | 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 | 716 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 717 |       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 | 718 | 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 | 719 | 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 | 720 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 721 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 722 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 723 | qed auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 724 | qed auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 725 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 726 | lemma%unimportant borel_Ioi: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 727 |   "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 | 728 | 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 | 729 | 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 | 730 | 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 | 731 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 732 | 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 | 733 | 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 | 734 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 735 | 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 | 736 |   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 | 737 | by blast | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 738 | 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 | 739 | proof | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 740 |     assume A: "A = {..< y}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 741 | show ?thesis | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 742 | proof cases | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 743 | 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 | 744 |       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 | 745 | 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 | 746 |       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 | 747 | 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 | 748 | 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 | 749 | 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 | 750 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 751 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 752 | 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 | 753 | 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 | 754 | 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 | 755 |       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 | 756 | 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 | 757 | 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 | 758 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 759 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 760 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 761 | qed auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 762 | qed auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 763 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 764 | lemma%unimportant borel_measurableI_less: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 765 |   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 | 766 |   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 | 767 | unfolding borel_Iio | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 768 | 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 | 769 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 770 | lemma%important borel_measurableI_greater: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 771 |   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 | 772 |   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 | 773 | unfolding borel_Ioi | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 774 | by%unimportant (rule measurable_measure_of) (auto simp: Int_def conj_commute) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 775 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 776 | lemma%unimportant borel_measurableI_le: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 777 |   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 778 |   shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 779 | by (rule borel_measurableI_greater) (auto simp: not_le[symmetric]) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 780 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 781 | lemma%unimportant borel_measurableI_ge: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 782 |   fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 783 |   shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 784 | by (rule borel_measurableI_less) (auto simp: not_le[symmetric]) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 785 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 786 | lemma%unimportant borel_measurable_less[measurable]: | 
| 63332 | 787 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 788 | assumes "f \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 789 | assumes "g \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 790 |   shows "{w \<in> space M. f w < g w} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 791 | proof - | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 792 |   have "{w \<in> space M. f w < g w} = (\<lambda>x. (f x, g x)) -` {x. fst x < snd x} \<inter> space M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 793 | by auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 794 | also have "\<dots> \<in> sets M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 795 | by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 796 | continuous_intros) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 797 | finally show ?thesis . | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 798 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 799 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 800 | lemma%important | 
| 63332 | 801 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 802 | assumes f[measurable]: "f \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 803 | assumes g[measurable]: "g \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 804 |   shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 805 |     and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 806 |     and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 807 | unfolding eq_iff not_less[symmetric] | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 808 | by%unimportant measurable | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 809 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 810 | lemma%important borel_measurable_SUP[measurable (raw)]: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 811 |   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 | 812 | assumes [simp]: "countable I" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 813 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 814 | shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 815 | by%unimportant (rule borel_measurableI_greater) (simp add: less_SUP_iff) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 816 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 817 | lemma%unimportant borel_measurable_INF[measurable (raw)]: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 818 |   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 | 819 | assumes [simp]: "countable I" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 820 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 821 | shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 822 | 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 | 823 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 824 | lemma%unimportant borel_measurable_cSUP[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 825 |   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 826 | assumes [simp]: "countable I" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 827 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 828 | assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_above ((\<lambda>i. F i x) ` I)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 829 | shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> borel_measurable M" | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 830 | proof cases | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 831 |   assume "I = {}" then show ?thesis
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 832 |     unfolding \<open>I = {}\<close> image_empty by simp
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 833 | next | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 834 |   assume "I \<noteq> {}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 835 | show ?thesis | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 836 | proof (rule borel_measurableI_le) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 837 | fix y | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 838 |     have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 839 | by measurable | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 840 |     also have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} = {x \<in> space M. (SUP i\<in>I. F i x) \<le> y}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 841 |       by (simp add: cSUP_le_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 842 |     finally show "{x \<in> space M. (SUP i\<in>I. F i x) \<le>  y} \<in> sets M"  .
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 843 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 844 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 845 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 846 | lemma%important borel_measurable_cINF[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 847 |   fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 848 | assumes [simp]: "countable I" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 849 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 850 | assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 851 | shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 852 | proof%unimportant cases | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 853 |   assume "I = {}" then show ?thesis
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 854 |     unfolding \<open>I = {}\<close> image_empty by simp
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 855 | next | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 856 |   assume "I \<noteq> {}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 857 | show ?thesis | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 858 | proof (rule borel_measurableI_ge) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 859 | fix y | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 860 |     have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 861 | by measurable | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 862 |     also have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} = {x \<in> space M. y \<le> (INF i\<in>I. F i x)}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 863 |       by (simp add: le_cINF_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 864 |     finally show "{x \<in> space M. y \<le> (INF i\<in>I. F i x)} \<in> sets M"  .
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 865 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 866 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 867 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 868 | lemma%unimportant borel_measurable_lfp[consumes 1, case_names continuity step]: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 869 |   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 | 870 | 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 | 871 | 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 | 872 | 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 | 873 | proof - | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 874 |   { 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 | 875 | 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 | 876 | 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 | 877 | by measurable | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 878 | 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 | 879 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 880 | 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 | 881 | 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 | 882 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 883 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 884 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 885 | lemma%unimportant borel_measurable_gfp[consumes 1, case_names continuity step]: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 886 |   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 | 887 | 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 | 888 | 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 | 889 | 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 | 890 | proof - | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 891 |   { 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 | 892 | 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 | 893 | 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 | 894 | by measurable | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 895 | 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 | 896 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 897 | 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 | 898 | 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 | 899 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 900 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59000diff
changeset | 901 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 902 | lemma%unimportant borel_measurable_max[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 903 |   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 904 | by (rule borel_measurableI_less) simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 905 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 906 | lemma%unimportant borel_measurable_min[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 907 |   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 908 | by (rule borel_measurableI_greater) simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 909 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 910 | lemma%unimportant borel_measurable_Min[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 911 |   "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, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 912 | proof (induct I rule: finite_induct) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 913 | case (insert i I) then show ?case | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 914 |     by (cases "I = {}") auto
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 915 | qed auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 916 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 917 | lemma%unimportant borel_measurable_Max[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 918 |   "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, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 919 | proof (induct I rule: finite_induct) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 920 | case (insert i I) then show ?case | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 921 |     by (cases "I = {}") auto
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 922 | qed auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 923 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 924 | lemma%unimportant borel_measurable_sup[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 925 |   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 926 | unfolding sup_max by measurable | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 927 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 928 | lemma%unimportant borel_measurable_inf[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 929 |   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 930 | unfolding inf_min by measurable | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 931 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 932 | lemma%unimportant [measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 933 |   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 934 | assumes "\<And>i. f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 935 | shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 936 | and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 937 | unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 938 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 939 | lemma%unimportant measurable_convergent[measurable (raw)]: | 
| 63332 | 940 |   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 941 | assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 942 | shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 943 | unfolding convergent_ereal by measurable | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 944 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 945 | lemma%unimportant sets_Collect_convergent[measurable]: | 
| 63332 | 946 |   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 947 | assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 948 |   shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 949 | by measurable | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 950 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 951 | lemma%unimportant borel_measurable_lim[measurable (raw)]: | 
| 63332 | 952 |   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 953 | assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 954 | shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 955 | proof - | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 956 | 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))" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 957 | by (simp add: lim_def convergent_def convergent_limsup_cl) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 958 | then show ?thesis | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 959 | by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 960 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 961 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 962 | lemma%unimportant borel_measurable_LIMSEQ_order: | 
| 63332 | 963 |   fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 964 | assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 965 | and u: "\<And>i. u i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 966 | shows "u' \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 967 | proof - | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 968 | have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 969 | using u' by (simp add: lim_imp_Liminf[symmetric]) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 970 | with u show ?thesis by (simp cong: measurable_cong) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 971 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 972 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 973 | subsection%important \<open>Borel spaces on topological monoids\<close> | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 974 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 975 | lemma%unimportant borel_measurable_add[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 976 |   fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 977 | assumes f: "f \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 978 | assumes g: "g \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 979 | shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 980 | using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 981 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 982 | lemma%unimportant borel_measurable_sum[measurable (raw)]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 983 |   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 984 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 985 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 986 | proof cases | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 987 | assume "finite S" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 988 | thus ?thesis using assms by induct auto | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 989 | qed simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 990 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 991 | lemma%important borel_measurable_suminf_order[measurable (raw)]: | 
| 63332 | 992 |   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}"
 | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 993 | assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 994 | shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 995 | unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 996 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 997 | subsection%important \<open>Borel spaces on Euclidean spaces\<close> | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 998 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 999 | lemma%important borel_measurable_inner[measurable (raw)]: | 
| 50881 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 hoelzl parents: 
50526diff
changeset | 1000 |   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 | 1001 | 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 | 1002 | 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 | 1003 | 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 | 1004 | using assms | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1005 | by%unimportant (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 | 1006 | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1007 | notation | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1008 | eucl_less (infix "<e" 50) | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1009 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1010 | lemma%important box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
 | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1011 |   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 | 1012 | by auto | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1013 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1014 | lemma%important eucl_ivals[measurable]: | 
| 61076 | 1015 | fixes a b :: "'a::ordered_euclidean_space" | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1016 |   shows "{x. x <e a} \<in> sets borel"
 | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1017 |     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 | 1018 |     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 | 1019 |     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 | 1020 |     and "{a..b} \<in> sets borel"
 | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1021 |     and  "{x. a <e x \<and> x \<le> b} \<in> sets borel"
 | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1022 |     and "{x. a \<le> x \<and>  x <e b} \<in> sets borel"
 | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1023 | unfolding box_oc box_co | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1024 | 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 | 1025 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1026 | lemma%unimportant (*FIX ME this has no name *) | 
| 51683 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 hoelzl parents: 
51478diff
changeset | 1027 |   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 | 1028 |   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 | 1029 |     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 | 1030 |     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 | 1031 |     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 | 1032 | 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 | 1033 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1034 | lemma%unimportant borel_eq_box: | 
| 61076 | 1035 | "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1036 | (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 | 1037 | 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 | 1038 |   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 | 1039 | 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 | 1040 | show "M \<in> ?SIGMA" | 
| 61808 | 1041 | apply (subst open_UNION_box[OF \<open>open M\<close>]) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1042 | 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 | 1043 | 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 | 1044 | done | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1045 | 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 | 1046 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1047 | lemma%unimportant halfspace_gt_in_halfspace: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1048 | assumes i: "i \<in> A" | 
| 62372 | 1049 |   shows "{x::'a. a < x \<bullet> i} \<in>
 | 
| 61076 | 1050 |     sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> 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 | 1051 | (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 | 1052 | proof - | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1053 | 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 | 1054 | by (intro sigma_algebra_sigma_sets) simp_all | 
| 61076 | 1055 |   have *: "?set = (\<Union>n. UNIV - {x::'a. x \<bullet> i < a + 1 / real (Suc n)})"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1056 | proof (safe, simp_all add: not_less del: of_nat_Suc) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1057 | 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 | 1058 | 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 | 1059 | 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 | 1060 | 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 | 1061 | 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 | 1062 | 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 | 1063 | next | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1064 | 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 | 1065 | 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 | 1066 | 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 | 1067 | 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 | 1068 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1069 | show "?set \<in> ?SIGMA" unfolding * | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61284diff
changeset | 1070 | by (auto intro!: Diff sigma_sets_Inter 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 | 1071 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1072 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1073 | lemma%unimportant borel_eq_halfspace_less: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1074 |   "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 | 1075 | (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 | 1076 | 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 | 1077 | 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 | 1078 |   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 | 1079 | 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 | 1080 | 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 | 1081 | 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 | 1082 | (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 | 1083 | 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 | 1084 | 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 | 1085 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1086 | lemma%unimportant borel_eq_halfspace_le: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1087 |   "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 | 1088 | (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 | 1089 | 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 | 1090 | 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 | 1091 | 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 | 1092 |   have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1093 | proof (safe, simp_all del: of_nat_Suc) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1094 | 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 | 1095 | 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 | 1096 | 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 | 1097 | 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 | 1098 | 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 | 1099 | 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 | 1100 | next | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1101 | 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 | 1102 | 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 | 1103 | 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 | 1104 | 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 | 1105 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1106 |   show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1107 | 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 | 1108 | 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 | 1109 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1110 | lemma%unimportant borel_eq_halfspace_ge: | 
| 61076 | 1111 |   "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1112 | (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 | 1113 | 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 | 1114 | 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 | 1115 |   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 | 1116 |   show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1117 | 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 | 1118 | 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 | 1119 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1120 | lemma%important borel_eq_halfspace_greater: | 
| 61076 | 1121 |   "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1122 | (is "_ = ?SIGMA") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1123 | proof%unimportant (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1124 | 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 | 1125 | 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 | 1126 |   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 | 1127 |   show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1128 | 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 | 1129 | 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 | 1130 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1131 | lemma%unimportant borel_eq_atMost: | 
| 61076 | 1132 |   "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1133 | (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 | 1134 | 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 | 1135 | 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 | 1136 | 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 | 1137 |   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)})"
 | 
| 62390 | 1138 | proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_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 | 1139 | 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 | 1140 | 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 | 1141 | 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 | 1142 | 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 | 1143 | 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 | 1144 | 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 | 1145 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1146 |   show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1147 | 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 | 1148 | 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 | 1149 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1150 | lemma%unimportant borel_eq_greaterThan: | 
| 61076 | 1151 |   "borel = sigma UNIV (range (\<lambda>a::'a::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 | 1152 | (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 | 1153 | 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 | 1154 | 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 | 1155 | 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 | 1156 |   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 | 1157 |   also have *: "{x::'a. a < x\<bullet>i} =
 | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1158 |       (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n) <e x})" using i
 | 
| 62390 | 1159 | proof (safe, simp_all add: eucl_less_def split: if_split_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 | 1160 | 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 | 1161 | 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 | 1162 | 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 | 1163 |     { 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 | 1164 | 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 | 1165 | 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 | 1166 | 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 | 1167 | 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 | 1168 | 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 | 1169 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1170 |   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 | 1171 | apply (simp only:) | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1172 | 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 | 1173 | 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 | 1174 | done | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1175 | 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 | 1176 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1177 | lemma%unimportant borel_eq_lessThan: | 
| 61076 | 1178 |   "borel = sigma UNIV (range (\<lambda>a::'a::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 | 1179 | (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 | 1180 | 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 | 1181 | 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 | 1182 | 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 | 1183 |   have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
 | 
| 61808 | 1184 |   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 \<open>i\<in> Basis\<close>
 | 
| 62390 | 1185 | proof (safe, simp_all add: eucl_less_def split: if_split_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 | 1186 | 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 | 1187 | 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 | 1188 | 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 | 1189 |     { 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 | 1190 | 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 | 1191 | 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 | 1192 | 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 | 1193 | 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 | 1194 | 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 | 1195 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1196 |   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 | 1197 | apply (simp only:) | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1198 | apply (intro sets.countable_UN sets.Diff) | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1199 | 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 | 1200 | done | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1201 | 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 | 1202 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1203 | lemma%important borel_eq_atLeastAtMost: | 
| 61076 | 1204 |   "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1205 | (is "_ = ?SIGMA") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1206 | proof%unimportant (rule borel_eq_sigmaI5[OF borel_eq_atMost]) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1207 | 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 | 1208 |   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 | 1209 | 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 | 1210 | 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 | 1211 | 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 | 1212 | 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 | 1213 |     { 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 | 1214 | 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 | 1215 | 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 | 1216 | 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 | 1217 | 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 | 1218 | 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 | 1219 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1220 |   show "{..a} \<in> ?SIGMA" unfolding *
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1221 | 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 | 1222 | (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 | 1223 | 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 | 1224 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1225 | lemma%important borel_set_induct[consumes 1, case_names empty interval compl union]: | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1226 | assumes "A \<in> sets borel" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1227 |   assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1228 | un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow> (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1229 | shows "P (A::real set)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1230 | proof%unimportant - | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1231 |   let ?G = "range (\<lambda>(a,b). {a..b::real})"
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1232 | have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G" | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1233 | using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1234 | thus ?thesis | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1235 | proof (induction rule: sigma_sets_induct_disjoint) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1236 | case (union f) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1237 | from union.hyps(2) have "\<And>i. f i \<in> sets borel" by (auto simp: borel_eq_atLeastAtMost) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1238 | with union show ?case by (auto intro: un) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1239 | next | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1240 | case (basic A) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1241 |     then obtain a b where "A = {a .. b}" by auto
 | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1242 | then show ?case | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1243 | by (cases "a \<le> b") (auto intro: int empty) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1244 | qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost) | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1245 | qed | 
| 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1246 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1247 | lemma%unimportant borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 1248 | 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 | 1249 | 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 | 1250 |   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 | 1251 | 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 | 1252 |   also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))"
 | 
| 62372 | 1253 | by (intro sets.countable_nat_UN) auto | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 1254 |   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 | 1255 | qed simp | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 1256 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1257 | lemma%unimportant eucl_lessThan: "{x::real. x <e a} = lessThan a"
 | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1258 | by (simp add: eucl_less_def lessThan_def) | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1259 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1260 | lemma%unimportant borel_eq_atLeastLessThan: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1261 |   "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 | 1262 | 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 | 1263 | 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 | 1264 | 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 | 1265 |   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 | 1266 | by (auto simp: move_uminus real_arch_simple) | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1267 |   then show "{y. y <e x} \<in> ?SIGMA"
 | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1268 | 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 | 1269 | 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 | 1270 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1271 | lemma%unimportant borel_measurable_halfspacesI: | 
| 61076 | 1272 | fixes f :: "'a \<Rightarrow> 'c::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1273 | assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))" | 
| 62372 | 1274 | and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space 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 | 1275 | 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 | 1276 | 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 | 1277 | 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 | 1278 | 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 | 1279 | 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 | 1280 | next | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1281 | 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 | 1282 | 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 | 1283 | 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 | 1284 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1285 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1286 | lemma%unimportant borel_measurable_iff_halfspace_le: | 
| 61076 | 1287 | fixes f :: "'a \<Rightarrow> 'c::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1288 |   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 | 1289 | 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 | 1290 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1291 | lemma%unimportant borel_measurable_iff_halfspace_less: | 
| 61076 | 1292 | fixes f :: "'a \<Rightarrow> 'c::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1293 |   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 | 1294 | 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 | 1295 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1296 | lemma%unimportant borel_measurable_iff_halfspace_ge: | 
| 61076 | 1297 | fixes f :: "'a \<Rightarrow> 'c::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1298 |   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 | 1299 | 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 | 1300 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1301 | lemma%unimportant borel_measurable_iff_halfspace_greater: | 
| 61076 | 1302 | fixes f :: "'a \<Rightarrow> 'c::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1303 |   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 | 1304 | 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 | 1305 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1306 | lemma%unimportant borel_measurable_iff_le: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1307 |   "(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 | 1308 | 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 | 1309 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1310 | lemma%unimportant borel_measurable_iff_less: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1311 |   "(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 | 1312 | 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 | 1313 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1314 | lemma%unimportant borel_measurable_iff_ge: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1315 |   "(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 | 1316 | 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 | 1317 | 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 | 1318 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1319 | lemma%unimportant borel_measurable_iff_greater: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1320 |   "(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 | 1321 | 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 | 1322 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1323 | lemma%important borel_measurable_euclidean_space: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1324 | 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 | 1325 | shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1326 | proof%unimportant safe | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1327 | 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 | 1328 | 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 | 1329 | 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 | 1330 | 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 | 1331 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1332 | subsection%important "Borel measurable operators" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50419diff
changeset | 1333 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1334 | lemma%important borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel" | 
| 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1335 | by%unimportant (intro borel_measurable_continuous_on1 continuous_intros) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1336 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1337 | lemma%unimportant borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel" | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57259diff
changeset | 1338 |   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 | 1339 | (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 | 1340 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1341 | lemma%important 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 | 1342 |   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 | 1343 | 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 | 1344 | shows "(\<lambda>x. - g x) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1345 | by%unimportant (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 | 1346 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1347 | lemma%important 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 | 1348 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
 | 
| 49774 | 1349 | assumes f: "f \<in> borel_measurable M" | 
| 1350 | assumes g: "g \<in> borel_measurable M" | |
| 1351 | shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1352 | using%unimportant borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) | 
| 49774 | 1353 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1354 | lemma%unimportant 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 | 1355 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
 | 
| 49774 | 1356 | assumes f: "f \<in> borel_measurable M" | 
| 1357 | assumes g: "g \<in> borel_measurable M" | |
| 1358 | 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 | 1359 | 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 | 1360 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1361 | lemma%important borel_measurable_prod[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 | 1362 |   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 | 1363 | 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 | 1364 | shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1365 | proof%unimportant cases | 
| 51683 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 hoelzl parents: 
51478diff
changeset | 1366 | 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 | 1367 | 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 | 1368 | qed simp | 
| 49774 | 1369 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1370 | lemma%important 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 | 1371 |   fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
 | 
| 49774 | 1372 | assumes f: "f \<in> borel_measurable M" | 
| 1373 | assumes g: "g \<in> borel_measurable M" | |
| 1374 | shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1375 | using%unimportant f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) | 
| 62372 | 1376 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1377 | lemma%unimportant 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 | 1378 |   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 | 1379 | assumes f: "f \<in> borel_measurable M" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1380 | assumes g: "g \<in> borel_measurable M" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1381 | 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 | 1382 | 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 | 1383 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1384 | lemma%unimportant borel_measurable_uminus_eq [simp]: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1385 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1386 | shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1387 | proof | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1388 | assume ?l from borel_measurable_uminus[OF this] show ?r by simp | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1389 | qed auto | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 1390 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1391 | lemma%unimportant affine_borel_measurable_vector: | 
| 38656 | 1392 | fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" | 
| 1393 | assumes "f \<in> borel_measurable M" | |
| 1394 | shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" | |
| 1395 | proof (rule borel_measurableI) | |
| 1396 | fix S :: "'x set" assume "open S" | |
| 1397 | show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" | |
| 1398 | proof cases | |
| 1399 | assume "b \<noteq> 0" | |
| 61808 | 1400 | with \<open>open S\<close> 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 | 1401 | using open_affinity [of S "inverse b" "- a /\<^sub>R b"] | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53216diff
changeset | 1402 | by (auto simp: algebra_simps) | 
| 47694 | 1403 | hence "?S \<in> sets borel" by auto | 
| 38656 | 1404 | moreover | 
| 61808 | 1405 | from \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" | 
| 38656 | 1406 | apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) | 
| 40859 | 1407 | ultimately show ?thesis using assms unfolding in_borel_measurable_borel | 
| 38656 | 1408 | by auto | 
| 1409 | qed simp | |
| 1410 | qed | |
| 1411 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1412 | lemma%unimportant borel_measurable_const_scaleR[measurable (raw)]: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1413 | "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 | 1414 | using affine_borel_measurable_vector[of f M 0 b] by simp | 
| 38656 | 1415 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1416 | lemma%unimportant borel_measurable_const_add[measurable (raw)]: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1417 | "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 | 1418 | 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 | 1419 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1420 | lemma%unimportant 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 | 1421 | fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" | 
| 49774 | 1422 | assumes f: "f \<in> borel_measurable M" | 
| 35692 | 1423 | 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 | 1424 | 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 | 1425 |   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 | 1426 | 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 | 1427 | done | 
| 35692 | 1428 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1429 | lemma%unimportant 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 | 1430 | "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 | 1431 |     (\<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 | 1432 | by (simp add: divide_inverse) | 
| 38656 | 1433 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1434 | lemma%unimportant borel_measurable_abs[measurable (raw)]: | 
| 50003 | 1435 | "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" | 
| 1436 | unfolding abs_real_def by simp | |
| 38656 | 1437 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1438 | lemma%unimportant borel_measurable_nth[measurable (raw)]: | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1439 | "(\<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 | 1440 | by (simp add: cart_eq_inner_axis) | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41025diff
changeset | 1441 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1442 | lemma%important convex_measurable: | 
| 59415 | 1443 | fixes A :: "'a :: euclidean_space set" | 
| 62372 | 1444 | shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow> | 
| 59415 | 1445 | (\<lambda>x. q (X x)) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1446 | by%unimportant (rule measurable_compose[where f=X and N="restrict_space borel A"]) | 
| 59415 | 1447 | (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2) | 
| 41830 | 1448 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1449 | lemma%unimportant borel_measurable_ln[measurable (raw)]: | 
| 49774 | 1450 | 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 | 1451 | 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 | 1452 | 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 | 1453 |   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 | 1454 | 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 | 1455 | done | 
| 41830 | 1456 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1457 | lemma%unimportant borel_measurable_log[measurable (raw)]: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1458 | "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M" | 
| 49774 | 1459 | unfolding log_def by auto | 
| 41830 | 1460 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1461 | lemma%unimportant borel_measurable_exp[measurable]: | 
| 58656 | 1462 |   "(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 | 1463 | by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) | 
| 50419 | 1464 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1465 | lemma%unimportant measurable_real_floor[measurable]: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1466 | "(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" | 
| 47761 | 1467 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1468 | have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))" | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1469 | by (auto intro: floor_eq2) | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1470 | then show ?thesis | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1471 | by (auto simp: vimage_def measurable_count_space_eq2_countable) | 
| 47761 | 1472 | qed | 
| 1473 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1474 | lemma%unimportant measurable_real_ceiling[measurable]: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1475 | "(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1476 | unfolding ceiling_def[abs_def] by simp | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1477 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1478 | lemma%unimportant borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel" | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1479 | by simp | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1480 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1481 | lemma%unimportant 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 | 1482 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1483 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1484 | lemma%unimportant borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1485 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1486 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1487 | lemma%unimportant borel_measurable_power [measurable (raw)]: | 
| 59415 | 1488 |   fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 1489 | assumes f: "f \<in> borel_measurable M" | |
| 1490 | shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M" | |
| 1491 | 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 | 1492 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1493 | lemma%unimportant borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1494 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1495 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1496 | lemma%unimportant borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1497 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1498 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1499 | lemma%unimportant borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1500 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1501 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1502 | lemma%unimportant 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 | 1503 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1504 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1505 | lemma%unimportant 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 | 1506 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1507 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1508 | lemma%unimportant borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1509 | by (intro borel_measurable_continuous_on1 continuous_intros) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57138diff
changeset | 1510 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1511 | lemma%important borel_measurable_complex_iff: | 
| 57259 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1512 | "f \<in> borel_measurable M \<longleftrightarrow> | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1513 | (\<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 | 1514 | apply auto | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1515 | apply (subst fun_complex_eq) | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1516 | apply (intro borel_measurable_add) | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1517 | apply auto | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1518 | done | 
| 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 hoelzl parents: 
57235diff
changeset | 1519 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1520 | lemma%important powr_real_measurable [measurable]: | 
| 67278 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 eberlm <eberlm@in.tum.de> parents: 
66164diff
changeset | 1521 | assumes "f \<in> measurable M borel" "g \<in> measurable M borel" | 
| 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 eberlm <eberlm@in.tum.de> parents: 
66164diff
changeset | 1522 | shows "(\<lambda>x. f x powr g x :: real) \<in> measurable M borel" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1523 | using%unimportant assms by (simp_all add: powr_def) | 
| 67278 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 eberlm <eberlm@in.tum.de> parents: 
66164diff
changeset | 1524 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1525 | lemma%unimportant measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel" | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1526 | by simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1527 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1528 | subsection%important "Borel space on the extended reals" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1529 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1530 | lemma%unimportant borel_measurable_ereal[measurable (raw)]: | 
| 43920 | 1531 | assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" | 
| 60771 | 1532 | using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1533 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1534 | lemma%unimportant borel_measurable_real_of_ereal[measurable (raw)]: | 
| 62372 | 1535 | fixes f :: "'a \<Rightarrow> ereal" | 
| 49774 | 1536 | assumes f: "f \<in> borel_measurable M" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1537 | shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1538 | apply (rule measurable_compose[OF f]) | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1539 |   apply (rule borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"])
 | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1540 | apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1541 | done | 
| 49774 | 1542 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1543 | lemma%unimportant borel_measurable_ereal_cases: | 
| 62372 | 1544 | fixes f :: "'a \<Rightarrow> ereal" | 
| 49774 | 1545 | assumes f: "f \<in> borel_measurable M" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1546 | assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M" | 
| 49774 | 1547 | shows "(\<lambda>x. H (f x)) \<in> borel_measurable M" | 
| 1548 | proof - | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1549 | let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real_of_ereal (f x)))" | 
| 49774 | 1550 |   { 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 | 1551 | with f H show ?thesis by simp | 
| 47694 | 1552 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1553 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1554 | lemma%unimportant (*FIX ME needs a name *) | 
| 50003 | 1555 | fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M" | 
| 1556 | shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" | |
| 1557 | and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M" | |
| 1558 | and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M" | |
| 49774 | 1559 | by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) | 
| 1560 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1561 | lemma%unimportant borel_measurable_uminus_eq_ereal[simp]: | 
| 49774 | 1562 | "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") | 
| 1563 | proof | |
| 1564 | assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp | |
| 1565 | qed auto | |
| 1566 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1567 | lemma%important set_Collect_ereal2: | 
| 62372 | 1568 | fixes f g :: "'a \<Rightarrow> ereal" | 
| 49774 | 1569 | assumes f: "f \<in> borel_measurable M" | 
| 1570 | assumes g: "g \<in> borel_measurable M" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1571 |   assumes H: "{x \<in> space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \<in> sets M"
 | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1572 |     "{x \<in> space borel. H (-\<infinity>) (ereal x)} \<in> sets borel"
 | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1573 |     "{x \<in> space borel. H (\<infinity>) (ereal x)} \<in> sets borel"
 | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1574 |     "{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
 | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1575 |     "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
 | 
| 49774 | 1576 |   shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1577 | proof%unimportant - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1578 | 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_of_ereal (g x)))" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1579 | 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_of_ereal (f x))) x" | 
| 49774 | 1580 |   { 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 | 1581 | note * = this | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1582 | from assms show ?thesis | 
| 62390 | 1583 | by (subst *) (simp del: space_borel split del: if_split) | 
| 49774 | 1584 | qed | 
| 1585 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1586 | lemma%unimportant borel_measurable_ereal_iff: | 
| 43920 | 1587 | 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 | 1588 | proof | 
| 43920 | 1589 | assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" | 
| 1590 | 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 | 1591 | 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 | 1592 | qed auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1593 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1594 | lemma%unimportant borel_measurable_erealD[measurable_dest]: | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 1595 | "(\<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 | 1596 | unfolding borel_measurable_ereal_iff by simp | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 1597 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1598 | lemma%important borel_measurable_ereal_iff_real: | 
| 43923 | 1599 | fixes f :: "'a \<Rightarrow> ereal" | 
| 1600 | shows "f \<in> borel_measurable M \<longleftrightarrow> | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1601 |     ((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1602 | proof%unimportant safe | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1603 |   assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1604 |   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 | 1605 |   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
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1606 | let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real_of_ereal (f x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1607 | have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto | 
| 43920 | 1608 | 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 | 1609 | finally show "f \<in> borel_measurable M" . | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1610 | qed simp_all | 
| 41830 | 1611 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1612 | lemma%unimportant borel_measurable_ereal_iff_Iio: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1613 |   "(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 | 1614 | by (auto simp: borel_Iio measurable_iff_measure_of) | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1615 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1616 | lemma%unimportant borel_measurable_ereal_iff_Ioi: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1617 |   "(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 | 1618 | by (auto simp: borel_Ioi measurable_iff_measure_of) | 
| 35582 | 1619 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1620 | lemma%unimportant vimage_sets_compl_iff: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1621 | "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 | 1622 | proof - | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1623 |   { fix A assume "f -` A \<inter> space M \<in> sets M"
 | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1624 | 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 | 1625 | 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 | 1626 | from this[of A] this[of "-A"] show ?thesis | 
| 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1627 | by (metis double_complement) | 
| 49774 | 1628 | qed | 
| 1629 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1630 | lemma%unimportant borel_measurable_iff_Iic_ereal: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1631 |   "(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 | 1632 |   unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp
 | 
| 38656 | 1633 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1634 | lemma%unimportant borel_measurable_iff_Ici_ereal: | 
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 1635 |   "(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 | 1636 |   unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp
 | 
| 38656 | 1637 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1638 | lemma%important borel_measurable_ereal2: | 
| 62372 | 1639 | fixes f g :: "'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1640 | assumes f: "f \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1641 | assumes g: "g \<in> borel_measurable M" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1642 | assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1643 | "(\<lambda>x. H (-\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1644 | "(\<lambda>x. H (\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1645 | "(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1646 | "(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M" | 
| 49774 | 1647 | shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1648 | proof%unimportant - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1649 | 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_of_ereal (g x)))" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61424diff
changeset | 1650 | 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_of_ereal (f x))) x" | 
| 49774 | 1651 |   { 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 | 1652 | note * = this | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1653 | 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 | 1654 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41969diff
changeset | 1655 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1656 | lemma%unimportant [measurable(raw)]: | 
| 43920 | 1657 | fixes f :: "'a \<Rightarrow> ereal" | 
| 50003 | 1658 | 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 | 1659 | 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 | 1660 | and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M" | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1661 | by (simp_all add: borel_measurable_ereal2) | 
| 49774 | 1662 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1663 | lemma%unimportant [measurable(raw)]: | 
| 49774 | 1664 | fixes f g :: "'a \<Rightarrow> ereal" | 
| 1665 | assumes "f \<in> borel_measurable M" | |
| 1666 | assumes "g \<in> borel_measurable M" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1667 | 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 | 1668 | and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M" | 
| 50003 | 1669 | using assms by (simp_all add: minus_ereal_def divide_ereal_def) | 
| 38656 | 1670 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1671 | lemma%unimportant borel_measurable_ereal_sum[measurable (raw)]: | 
| 43920 | 1672 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 41096 | 1673 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 1674 | 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 | 1675 | using assms by (induction S rule: infinite_finite_induct) auto | 
| 38656 | 1676 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1677 | lemma%unimportant borel_measurable_ereal_prod[measurable (raw)]: | 
| 43920 | 1678 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 38656 | 1679 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | 
| 41096 | 1680 | 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 | 1681 | using assms by (induction S rule: infinite_finite_induct) auto | 
| 38656 | 1682 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1683 | lemma%unimportant borel_measurable_extreal_suminf[measurable (raw)]: | 
| 43920 | 1684 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 50003 | 1685 | 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 | 1686 | shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M" | 
| 50003 | 1687 | unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp | 
| 39092 | 1688 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1689 | subsection%important "Borel space on the extended non-negative reals" | 
| 62625 | 1690 | |
| 1691 | text \<open> @{type ennreal} is a topological monoid, so no rules for plus are required, also all order
 | |
| 1692 | statements are usually done on type classes. \<close> | |
| 1693 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1694 | lemma%unimportant measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel" | 
| 62625 | 1695 | by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal) | 
| 1696 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1697 | lemma%unimportant measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel" | 
| 62625 | 1698 | by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal) | 
| 1699 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1700 | lemma%unimportant borel_measurable_enn2real[measurable (raw)]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1701 | "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1702 | unfolding enn2real_def[abs_def] by measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1703 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1704 | definition%important [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M" | 
| 62625 | 1705 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1706 | lemma%unimportant is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel" | 
| 62625 | 1707 | unfolding is_borel_def[abs_def] | 
| 1708 | proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) | |
| 1709 | fix f and M :: "'a measure" | |
| 1710 | show "f \<in> borel_measurable M" if f: "enn2ereal \<circ> f \<in> borel_measurable M" | |
| 1711 | using measurable_compose[OF f measurable_e2ennreal] by simp | |
| 1712 | qed simp | |
| 1713 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1714 | context | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1715 | includes ennreal.lifting | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1716 | begin | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1717 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1718 | lemma%unimportant measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1719 | unfolding is_borel_def[symmetric] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1720 | by transfer simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1721 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1722 | lemma%important borel_measurable_ennreal_iff[simp]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1723 | assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1724 | shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1725 | proof%unimportant safe | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1726 | assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1727 | then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1728 | by measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1729 | then show "f \<in> M \<rightarrow>\<^sub>M borel" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1730 | by (rule measurable_cong[THEN iffD1, rotated]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1731 | qed measurable | 
| 62625 | 1732 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1733 | lemma%unimportant borel_measurable_times_ennreal[measurable (raw)]: | 
| 62625 | 1734 | fixes f g :: "'a \<Rightarrow> ennreal" | 
| 1735 | shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel" | |
| 1736 | unfolding is_borel_def[symmetric] by transfer simp | |
| 1737 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1738 | lemma%unimportant borel_measurable_inverse_ennreal[measurable (raw)]: | 
| 62625 | 1739 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 1740 | shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel" | |
| 1741 | unfolding is_borel_def[symmetric] by transfer simp | |
| 1742 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1743 | lemma%unimportant borel_measurable_divide_ennreal[measurable (raw)]: | 
| 62625 | 1744 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 1745 | shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel" | |
| 1746 | unfolding divide_ennreal_def by simp | |
| 1747 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1748 | lemma%unimportant borel_measurable_minus_ennreal[measurable (raw)]: | 
| 62625 | 1749 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 1750 | shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel" | |
| 1751 | unfolding is_borel_def[symmetric] by transfer simp | |
| 1752 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1753 | lemma%important borel_measurable_prod_ennreal[measurable (raw)]: | 
| 62625 | 1754 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal" | 
| 1755 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" | |
| 1756 | shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1757 | using%unimportant assms by (induction S rule: infinite_finite_induct) auto | 
| 62625 | 1758 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1759 | end | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62625diff
changeset | 1760 | |
| 62625 | 1761 | hide_const (open) is_borel | 
| 1762 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1763 | subsection%important \<open>LIMSEQ is borel measurable\<close> | 
| 39092 | 1764 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1765 | lemma%important borel_measurable_LIMSEQ_real: | 
| 39092 | 1766 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" | 
| 61969 | 1767 | assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" | 
| 39092 | 1768 | and u: "\<And>i. u i \<in> borel_measurable M" | 
| 1769 | shows "u' \<in> borel_measurable M" | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1770 | proof%unimportant - | 
| 43920 | 1771 | have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)" | 
| 46731 | 1772 | using u' by (simp add: lim_imp_Liminf) | 
| 43920 | 1773 | moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M" | 
| 39092 | 1774 | by auto | 
| 43920 | 1775 | ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) | 
| 39092 | 1776 | qed | 
| 1777 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1778 | lemma%important borel_measurable_LIMSEQ_metric: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1779 | 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 | 1780 | assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 61969 | 1781 | assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1782 | 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 | 1783 | unfolding borel_eq_closed | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1784 | proof%unimportant (safe intro!: measurable_measure_of) | 
| 62372 | 1785 | fix A :: "'b set" assume "closed A" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1786 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1787 | have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M" | 
| 62624 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 hoelzl parents: 
62390diff
changeset | 1788 | proof (rule borel_measurable_LIMSEQ_real) | 
| 61969 | 1789 | show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) \<longlonglongrightarrow> infdist (g x) A" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1790 | 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 | 1791 | 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 | 1792 | 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 | 1793 | 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 | 1794 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1795 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1796 | 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 | 1797 | proof cases | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1798 |     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 | 1799 | then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A" | 
| 61808 | 1800 | using \<open>closed A\<close> by (simp add: in_closed_iff_infdist_zero) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1801 |     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 | 1802 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1803 | 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 | 1804 | by measurable | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1805 | 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 | 1806 | qed simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1807 | qed auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56371diff
changeset | 1808 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1809 | lemma%important sets_Collect_Cauchy[measurable]: | 
| 57036 | 1810 |   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 | 1811 | assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 49774 | 1812 |   shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
 | 
| 57036 | 1813 | unfolding metric_Cauchy_iff2 using f by auto | 
| 49774 | 1814 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1815 | lemma%unimportant borel_measurable_lim_metric[measurable (raw)]: | 
| 57036 | 1816 |   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 | 1817 | assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 49774 | 1818 | shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" | 
| 1819 | proof - | |
| 63040 | 1820 | define u' where "u' x = lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" for x | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1821 | then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))" | 
| 64287 | 1822 | by (auto simp: lim_def convergent_eq_Cauchy[symmetric]) | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1823 | have "u' \<in> borel_measurable M" | 
| 57036 | 1824 | proof (rule borel_measurable_LIMSEQ_metric) | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1825 | fix x | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1826 | have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" | 
| 49774 | 1827 | by (cases "Cauchy (\<lambda>i. f i x)") | 
| 64287 | 1828 | (auto simp add: convergent_eq_Cauchy[symmetric] convergent_def) | 
| 61969 | 1829 | then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) \<longlonglongrightarrow> u' x" | 
| 62372 | 1830 | unfolding u'_def | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1831 | by (rule convergent_LIMSEQ_iff[THEN iffD1]) | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1832 | qed measurable | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1833 | then show ?thesis | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1834 | unfolding * by measurable | 
| 49774 | 1835 | qed | 
| 1836 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1837 | lemma%unimportant borel_measurable_suminf[measurable (raw)]: | 
| 57036 | 1838 |   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 | 1839 | assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 49774 | 1840 | 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 | 1841 | unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp | 
| 49774 | 1842 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1843 | lemma%unimportant Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
 | 
| 63389 | 1844 | by (simp add: pred_def) | 
| 1845 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 1846 | (* Proof by Jeremy Avigad and Luke Serafin *) | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1847 | lemma%unimportant isCont_borel_pred[measurable]: | 
| 63389 | 1848 | fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space" | 
| 1849 | shows "Measurable.pred borel (isCont f)" | |
| 1850 | proof (subst measurable_cong) | |
| 1851 | let ?I = "\<lambda>j. inverse(real (Suc j))" | |
| 1852 | show "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)" for x | |
| 1853 | unfolding continuous_at_eps_delta | |
| 1854 | proof safe | |
| 1855 | fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" | |
| 1856 | moreover have "0 < ?I i / 2" | |
| 1857 | by simp | |
| 1858 | ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2" | |
| 1859 | by (metis dist_commute) | |
| 1860 | then obtain j where j: "?I j < d" | |
| 1861 | by (metis reals_Archimedean) | |
| 1862 | ||
| 1863 | 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" | |
| 1864 | proof (safe intro!: exI[where x=j]) | |
| 1865 | fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" | |
| 1866 | have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)" | |
| 1867 | by (rule dist_triangle2) | |
| 1868 | also have "\<dots> < ?I i / 2 + ?I i / 2" | |
| 1869 | by (intro add_strict_mono d less_trans[OF _ j] *) | |
| 1870 | also have "\<dots> \<le> ?I i" | |
| 1871 | by (simp add: field_simps of_nat_Suc) | |
| 1872 | finally show "dist (f y) (f z) \<le> ?I i" | |
| 1873 | by simp | |
| 1874 | qed | |
| 1875 | next | |
| 1876 | fix e::real assume "0 < e" | |
| 1877 | then obtain n where n: "?I n < e" | |
| 1878 | by (metis reals_Archimedean) | |
| 1879 | 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" | |
| 1880 | from this[THEN spec, of "Suc n"] | |
| 1881 | 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)" | |
| 1882 | by auto | |
| 1883 | ||
| 1884 | show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" | |
| 1885 | proof (safe intro!: exI[of _ "?I j"]) | |
| 1886 | fix y assume "dist y x < ?I j" | |
| 1887 | then have "dist (f y) (f x) \<le> ?I (Suc n)" | |
| 1888 | by (intro j) (auto simp: dist_commute) | |
| 1889 | also have "?I (Suc n) < ?I n" | |
| 1890 | by simp | |
| 1891 | also note n | |
| 1892 | finally show "dist (f y) (f x) < e" . | |
| 1893 | qed simp | |
| 1894 | qed | |
| 1895 | qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less | |
| 1896 | Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros) | |
| 1897 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1898 | lemma%unimportant isCont_borel: | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 1899 | 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 | 1900 |   shows "{x. isCont f x} \<in> sets borel"
 | 
| 63389 | 1901 | by simp | 
| 62083 | 1902 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1903 | lemma%important is_real_interval: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1904 | assumes S: "is_interval S" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1905 |   shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or>
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1906 |     S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}"
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1907 | using S unfolding is_interval_1 by (blast intro: interval_cases) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1908 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1909 | lemma%important real_interval_borel_measurable: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1910 | assumes "is_interval (S::real set)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1911 | shows "S \<in> sets borel" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1912 | proof%unimportant - | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1913 |   from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or>
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1914 |     S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto
 | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1915 | then guess a .. | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1916 | then guess b .. | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1917 | thus ?thesis | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1918 | by auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1919 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1920 | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1921 | text \<open>The next lemmas hold in any second countable linorder (including ennreal or ereal for instance), | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1922 | but in the current state they are restricted to reals.\<close> | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1923 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1924 | lemma%important borel_measurable_mono_on_fnc: | 
| 62083 | 1925 | fixes f :: "real \<Rightarrow> real" and A :: "real set" | 
| 1926 | assumes "mono_on f A" | |
| 1927 | shows "f \<in> borel_measurable (restrict_space borel A)" | |
| 1928 | apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]) | |
| 1929 |   apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space)
 | |
| 1930 | apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within | |
| 62372 | 1931 | cong: measurable_cong_sets | 
| 62083 | 1932 | intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset) | 
| 1933 | done | |
| 1934 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1935 | lemma%unimportant borel_measurable_piecewise_mono: | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1936 | fixes f::"real \<Rightarrow> real" and C::"real set set" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1937 | assumes "countable C" "\<And>c. c \<in> C \<Longrightarrow> c \<in> sets borel" "\<And>c. c \<in> C \<Longrightarrow> mono_on f c" "(\<Union>C) = UNIV" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1938 | shows "f \<in> borel_measurable borel" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1939 | by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64272diff
changeset | 1940 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1941 | lemma%unimportant borel_measurable_mono: | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1942 | fixes f :: "real \<Rightarrow> real" | 
| 62083 | 1943 | shows "mono f \<Longrightarrow> f \<in> borel_measurable borel" | 
| 1944 | using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def) | |
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1945 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1946 | lemma%unimportant measurable_bdd_below_real[measurable (raw)]: | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1947 | fixes F :: "'a \<Rightarrow> 'i \<Rightarrow> real" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1948 | assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> M \<rightarrow>\<^sub>M borel" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1949 | shows "Measurable.pred M (\<lambda>x. bdd_below ((\<lambda>i. F i x)`I))" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1950 | proof (subst measurable_cong) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1951 | show "bdd_below ((\<lambda>i. F i x)`I) \<longleftrightarrow> (\<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i x)" for x | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1952 | by (auto simp: bdd_below_def intro!: bexI[of _ "of_int (floor _)"] intro: order_trans of_int_floor_le) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1953 | show "Measurable.pred M (\<lambda>w. \<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i w)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1954 | using countable_int by measurable | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1955 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1956 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1957 | lemma%important borel_measurable_cINF_real[measurable (raw)]: | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1958 | fixes F :: "_ \<Rightarrow> _ \<Rightarrow> real" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1959 | assumes [simp]: "countable I" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1960 | assumes F[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 1961 | shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1962 | proof%unimportant (rule measurable_piecewise_restrict) | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1963 |   let ?\<Omega> = "{x\<in>space M. bdd_below ((\<lambda>i. F i x)`I)}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1964 |   show "countable {?\<Omega>, - ?\<Omega>}" "space M \<subseteq> \<Union>{?\<Omega>, - ?\<Omega>}" "\<And>X. X \<in> {?\<Omega>, - ?\<Omega>} \<Longrightarrow> X \<inter> space M \<in> sets M"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1965 | by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 1966 |   fix X assume "X \<in> {?\<Omega>, - ?\<Omega>}" then show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M X)"
 | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1967 | proof safe | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 1968 | show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M ?\<Omega>)" | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1969 | by (intro borel_measurable_cINF measurable_restrict_space1 F) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1970 | (auto simp: space_restrict_space) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 1971 | show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M (-?\<Omega>))" | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1972 | proof (subst measurable_cong) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1973 | fix x assume "x \<in> space (restrict_space M (-?\<Omega>))" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1974 | then have "\<not> (\<forall>i\<in>I. - F i x \<le> y)" for y | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1975 | by (auto simp: space_restrict_space bdd_above_def bdd_above_uminus[symmetric]) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69022diff
changeset | 1976 | then show "(INF i\<in>I. F i x) = - (THE x. False)" | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1977 | by (auto simp: space_restrict_space Inf_real_def Sup_real_def Least_def simp del: Set.ball_simps(10)) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1978 | qed simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1979 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1980 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1981 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1982 | lemma%unimportant borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
 | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1983 | proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1984 | fix x :: real | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1985 |   have eq: "{..<x} = space (sigma UNIV (range atLeast)) - {x ..}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1986 | by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1987 |   show "{..<x} \<in> sets (sigma UNIV (range atLeast))"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1988 | unfolding eq by (intro sets.compl_sets) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1989 | qed auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1990 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1991 | lemma%unimportant borel_measurable_pred_less[measurable (raw)]: | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1992 |   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1993 | shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> Measurable.pred M (\<lambda>w. f w < g w)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1994 | unfolding Measurable.pred_def by (rule borel_measurable_less) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63952diff
changeset | 1995 | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1996 | no_notation | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1997 | eucl_less (infix "<e" 50) | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 1998 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 1999 | lemma%important borel_measurable_Max2[measurable (raw)]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2000 |   fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2001 | assumes "finite I" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2002 | and [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2003 |   shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2004 | by%unimportant (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image) | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2005 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2006 | lemma%unimportant measurable_compose_n [measurable (raw)]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2007 | assumes "T \<in> measurable M M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2008 | shows "(T^^n) \<in> measurable M M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2009 | by (induction n, auto simp add: measurable_compose[OF _ assms]) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2010 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2011 | lemma%unimportant measurable_real_imp_nat: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2012 | fixes f::"'a \<Rightarrow> nat" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2013 | assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2014 | shows "f \<in> measurable M (count_space UNIV)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2015 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2016 | let ?g = "(\<lambda>x. real(f x))" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2017 |   have "\<And>(n::nat). ?g-`({real n}) \<inter> space M = f-`{n} \<inter> space M" by auto
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2018 |   moreover have "\<And>(n::nat). ?g-`({real n}) \<inter> space M \<in> sets M" using assms by measurable
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2019 |   ultimately have "\<And>(n::nat). f-`{n} \<inter> space M \<in> sets M" by simp
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2020 | then show ?thesis using measurable_count_space_eq2_countable by blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2021 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2022 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2023 | lemma%unimportant measurable_equality_set [measurable]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2024 |   fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2025 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2026 |   shows "{x \<in> space M. f x = g x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2027 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2028 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2029 |   define A where "A = {x \<in> space M. f x = g x}"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2030 |   define B where "B = {y. \<exists>x::'a. y = (x,x)}"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2031 | have "A = (\<lambda>x. (f x, g x))-`B \<inter> space M" unfolding A_def B_def by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2032 | moreover have "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2033 | moreover have "B \<in> sets borel" unfolding B_def by (simp add: closed_diagonal) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2034 | ultimately have "A \<in> sets M" by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2035 | then show ?thesis unfolding A_def by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2036 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2037 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2038 | lemma%unimportant measurable_inequality_set [measurable]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2039 |   fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2040 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2041 |   shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2042 |         "{x \<in> space M. f x < g x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2043 |         "{x \<in> space M. f x \<ge> g x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2044 |         "{x \<in> space M. f x > g x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2045 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2046 | define F where "F = (\<lambda>x. (f x, g x))" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2047 | have * [measurable]: "F \<in> borel_measurable M" unfolding F_def by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2048 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2049 |   have "{x \<in> space M. f x \<le> g x} = F-`{(x, y) | x y. x \<le> y} \<inter> space M" unfolding F_def by auto
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2050 |   moreover have "{(x, y) | x y. x \<le> (y::'a)} \<in> sets borel" using closed_subdiagonal borel_closed by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2051 |   ultimately show "{x \<in> space M. f x \<le> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2052 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2053 |   have "{x \<in> space M. f x < g x} = F-`{(x, y) | x y. x < y} \<inter> space M" unfolding F_def by auto
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2054 |   moreover have "{(x, y) | x y. x < (y::'a)} \<in> sets borel" using open_subdiagonal borel_open by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2055 |   ultimately show "{x \<in> space M. f x < g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2056 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2057 |   have "{x \<in> space M. f x \<ge> g x} = F-`{(x, y) | x y. x \<ge> y} \<inter> space M" unfolding F_def by auto
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2058 |   moreover have "{(x, y) | x y. x \<ge> (y::'a)} \<in> sets borel" using closed_superdiagonal borel_closed by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2059 |   ultimately show "{x \<in> space M. f x \<ge> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2060 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2061 |   have "{x \<in> space M. f x > g x} = F-`{(x, y) | x y. x > y} \<inter> space M" unfolding F_def by auto
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2062 |   moreover have "{(x, y) | x y. x > (y::'a)} \<in> sets borel" using open_superdiagonal borel_open by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2063 |   ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2064 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2065 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2066 | lemma%unimportant measurable_limit [measurable]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2067 | fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2068 | assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2069 | shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2070 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2071 | obtain A :: "nat \<Rightarrow> 'b set" where A: | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2072 | "\<And>i. open (A i)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2073 | "\<And>i. c \<in> A i" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2074 | "\<And>S. open S \<Longrightarrow> c \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2075 | by (rule countable_basis_at_decseq) blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2076 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2077 | have [measurable]: "\<And>N i. (f N)-`(A i) \<inter> space M \<in> sets M" using A(1) by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2078 |   then have mes: "(\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M) \<in> sets M" by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2079 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2080 | have "(u \<longlonglongrightarrow> c) \<longleftrightarrow> (\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" for u::"nat \<Rightarrow> 'b" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2081 | proof | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2082 | assume "u \<longlonglongrightarrow> c" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2083 | then have "eventually (\<lambda>n. u n \<in> A i) sequentially" for i using A(1)[of i] A(2)[of i] | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2084 | by (simp add: topological_tendstoD) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2085 | then show "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2086 | next | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2087 | assume H: "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2088 | show "(u \<longlonglongrightarrow> c)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2089 | proof (rule topological_tendstoI) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2090 | fix S assume "open S" "c \<in> S" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2091 | with A(3)[OF this] obtain i where "A i \<subseteq> S" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2092 | using eventually_False_sequentially eventually_mono by blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2093 | moreover have "eventually (\<lambda>n. u n \<in> A i) sequentially" using H by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2094 | ultimately show "\<forall>\<^sub>F n in sequentially. u n \<in> S" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2095 | by (simp add: eventually_mono subset_eq) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2096 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2097 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2098 |   then have "{x. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i))"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2099 | by (auto simp add: atLeast_def eventually_at_top_linorder) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2100 |   then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M)"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2101 | by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2102 |   then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} \<in> sets M" using mes by simp
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2103 | then show ?thesis by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2104 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2105 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2106 | lemma%important measurable_limit2 [measurable]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2107 | fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2108 | assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2109 | shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2110 | proof%unimportant - | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2111 | define w where "w = (\<lambda>n x. u n x - v x)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2112 | have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2113 | have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2114 | unfolding w_def using Lim_null by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2115 | then show ?thesis using measurable_limit by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2116 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2117 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2118 | lemma%unimportant measurable_P_restriction [measurable (raw)]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2119 | assumes [measurable]: "Measurable.pred M P" "A \<in> sets M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2120 |   shows "{x \<in> A. P x} \<in> sets M"
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2121 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2122 | have "A \<subseteq> space M" using sets.sets_into_space[OF assms(2)]. | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2123 |   then have "{x \<in> A. P x} = A \<inter> {x \<in> space M. P x}" by blast
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2124 | then show ?thesis by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2125 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2126 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2127 | lemma%unimportant measurable_sum_nat [measurable (raw)]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2128 | fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2129 | assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2130 | shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2131 | proof cases | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2132 | assume "finite S" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2133 | then show ?thesis using assms by induct auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2134 | qed simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2135 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2136 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2137 | lemma%unimportant measurable_abs_powr [measurable]: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2138 | fixes p::real | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2139 | assumes [measurable]: "f \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2140 | shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2141 | unfolding powr_def by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2142 | |
| 64911 | 2143 | text \<open>The next one is a variation around \verb+measurable_restrict_space+.\<close> | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2144 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2145 | lemma%unimportant measurable_restrict_space3: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2146 | assumes "f \<in> measurable M N" and | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2147 | "f \<in> A \<rightarrow> B" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2148 | shows "f \<in> measurable (restrict_space M A) (restrict_space N B)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2149 | proof - | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2150 | have "f \<in> measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2151 | then show ?thesis by (metis Int_iff funcsetI funcset_mem | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2152 | measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2153 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2154 | |
| 64911 | 2155 | text \<open>The next one is a variation around \verb+measurable_piecewise_restrict+.\<close> | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2156 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2157 | lemma%important measurable_piecewise_restrict2: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2158 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2159 | and "space M = (\<Union>(n::nat). A n)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2160 | "\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2161 | shows "f \<in> measurable M N" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68635diff
changeset | 2162 | proof%unimportant (rule measurableI) | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2163 | fix B assume [measurable]: "B \<in> sets N" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2164 |   {
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2165 | fix n::nat | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2166 | obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2167 | then have *: "f-`B \<inter> A n = h-`B \<inter> A n" by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2168 | have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n" using assms(2) sets.sets_into_space by auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2169 | then have "h-`B \<inter> A n \<in> sets M" by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2170 | then have "f-`B \<inter> A n \<in> sets M" using * by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2171 | } | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2172 | then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M" by measurable | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2173 | moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)" using assms(2) by blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2174 | ultimately show "f-`B \<inter> space M \<in> sets M" by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2175 | next | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2176 | fix x assume "x \<in> space M" | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2177 | then obtain n where "x \<in> A n" using assms(2) by blast | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2178 | obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast | 
| 64911 | 2179 | then have "f x = h x" using \<open>x \<in> A n\<close> by blast | 
| 2180 | moreover have "h x \<in> space N" by (metis measurable_space \<open>x \<in> space M\<close> \<open>h \<in> measurable M N\<close>) | |
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2181 | ultimately show "f x \<in> space N" by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2182 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 2183 | |
| 51683 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 hoelzl parents: 
51478diff
changeset | 2184 | end |