| author | wenzelm | 
| Thu, 01 Sep 2022 10:52:30 +0200 | |
| changeset 76034 | dda3c117f13c | 
| parent 75462 | 7448423e5dba | 
| child 78475 | a5f6d2fc1b1f | 
| permissions | -rw-r--r-- | 
| 68017 | 1 | (* Title: HOL/Analysis/Change_Of_Vars.thy | 
| 2 | Authors: LC Paulson, based on material from HOL Light | |
| 3 | *) | |
| 4 | ||
| 5 | section\<open>Change of Variables Theorems\<close> | |
| 6 | ||
| 67998 | 7 | theory Change_Of_Vars | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo EmÃlio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67998diff
changeset | 8 | imports Vitali_Covering_Theorem Determinants | 
| 67998 | 9 | |
| 10 | begin | |
| 11 | ||
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69661diff
changeset | 12 | subsection \<open>Measurable Shear and Stretch\<close> | 
| 67998 | 13 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 14 | proposition | 
| 67998 | 15 | fixes a :: "real^'n" | 
| 16 |   assumes "m \<noteq> n" and ab_ne: "cbox a b \<noteq> {}" and an: "0 \<le> a$n"
 | |
| 17 | shows measurable_shear_interval: "(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` (cbox a b) \<in> lmeasurable" | |
| 18 | (is "?f ` _ \<in> _") | |
| 19 | and measure_shear_interval: "measure lebesgue ((\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` cbox a b) | |
| 20 | = measure lebesgue (cbox a b)" (is "?Q") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 21 | proof - | 
| 67998 | 22 | have lin: "linear ?f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 23 | by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) | 
| 67998 | 24 | show fab: "?f ` cbox a b \<in> lmeasurable" | 
| 25 | by (simp add: lin measurable_linear_image_interval) | |
| 26 | let ?c = "\<chi> i. if i = m then b$m + b$n else b$i" | |
| 27 | let ?mn = "axis m 1 - axis n (1::real)" | |
| 28 | have eq1: "measure lebesgue (cbox a ?c) | |
| 29 | = measure lebesgue (?f ` cbox a b) | |
| 30 |             + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m})
 | |
| 31 |             + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})"
 | |
| 32 | proof (rule measure_Un3_negligible) | |
| 33 |     show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m} \<in> lmeasurable" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
 | |
| 34 | by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) | |
| 35 |     have "negligible {x. ?mn \<bullet> x = a$m}"
 | |
| 36 | by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) | |
| 37 |     moreover have "?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) \<subseteq> {x. ?mn \<bullet> x = a$m}"
 | |
| 38 | using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') | |
| 39 |     ultimately show "negligible ((?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}))"
 | |
| 40 | by (rule negligible_subset) | |
| 41 |     have "negligible {x. ?mn \<bullet> x = b$m}"
 | |
| 42 | by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) | |
| 43 |     moreover have "(?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}) \<subseteq> {x. ?mn \<bullet> x = b$m}"
 | |
| 44 | using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') | |
| 45 |     ultimately show "negligible (?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
 | |
| 46 | by (rule negligible_subset) | |
| 47 |     have "negligible {x. ?mn \<bullet> x = b$m}"
 | |
| 48 | by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) | |
| 49 |     moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})) \<subseteq> {x. ?mn \<bullet> x = b$m}"
 | |
| 50 | using \<open>m \<noteq> n\<close> ab_ne | |
| 51 | apply (auto simp: algebra_simps mem_box_cart inner_axis') | |
| 52 | apply (drule_tac x=m in spec)+ | |
| 53 | apply simp | |
| 54 | done | |
| 55 |     ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
 | |
| 56 | by (rule negligible_subset) | |
| 57 |     show "?f ` cbox a b \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} = cbox a ?c" (is "?lhs = _")
 | |
| 58 | proof | |
| 59 | show "?lhs \<subseteq> cbox a ?c" | |
| 60 | by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) | |
| 61 | show "cbox a ?c \<subseteq> ?lhs" | |
| 62 | apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \<open>m \<noteq> n\<close>]) | |
| 63 | apply (auto simp: mem_box_cart split: if_split_asm) | |
| 64 | done | |
| 65 | qed | |
| 66 | qed (fact fab) | |
| 67 | let ?d = "\<chi> i. if i = m then a $ m - b $ m else 0" | |
| 68 |   have eq2: "measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})
 | |
| 69 | = measure lebesgue (cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i))" | |
| 70 |   proof (rule measure_translate_add[of "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}"
 | |
| 71 | "(\<chi> i. if i = m then a$m - b$m else 0)" "cbox a (\<chi> i. if i = m then a$m + b$n else b$i)"]) | |
| 72 |     show "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}) \<in> lmeasurable"
 | |
| 73 |       "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
 | |
| 74 | by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) | |
| 75 | have "\<And>x. \<lbrakk>x $ n + a $ m \<le> x $ m\<rbrakk> | |
| 76 |          \<Longrightarrow> x \<in> (+) (\<chi> i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \<le> x $ m}"
 | |
| 77 | using \<open>m \<noteq> n\<close> | |
| 78 | by (rule_tac x="x - (\<chi> i. if i = m then a$m - b$m else 0)" in image_eqI) | |
| 79 | (simp_all add: mem_box_cart) | |
| 80 |     then have imeq: "(+) ?d ` {x. b $ m \<le> ?mn \<bullet> x} = {x. a $ m \<le> ?mn \<bullet> x}"
 | |
| 81 | using \<open>m \<noteq> n\<close> by (auto simp: mem_box_cart inner_axis' algebra_simps) | |
| 82 | have "\<And>x. \<lbrakk>0 \<le> a $ n; x $ n + a $ m \<le> x $ m; | |
| 83 | \<forall>i. i \<noteq> m \<longrightarrow> a $ i \<le> x $ i \<and> x $ i \<le> b $ i\<rbrakk> | |
| 84 | \<Longrightarrow> a $ m \<le> x $ m" | |
| 85 | using \<open>m \<noteq> n\<close> by force | |
| 86 |     then have "(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})
 | |
| 87 |             = cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<inter> {x. a $ m \<le> ?mn \<bullet> x}"
 | |
| 88 | using an ab_ne | |
| 89 | apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) | |
| 90 | apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) | |
| 91 | by (metis (full_types) add_mono mult_2_right) | |
| 92 |     then show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union>
 | |
| 93 |           (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}) =
 | |
| 94 | cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") | |
| 95 | using an \<open>m \<noteq> n\<close> | |
| 96 | apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) | |
| 97 | apply (drule_tac x=n in spec)+ | |
| 98 | by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) | |
| 99 |     have "negligible{x. ?mn \<bullet> x = a$m}"
 | |
| 100 | by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) | |
| 101 |     moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
 | |
| 102 |                                  (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})) \<subseteq> {x. ?mn \<bullet> x = a$m}"
 | |
| 103 | using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') | |
| 104 |     ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
 | |
| 105 |                                  (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}))"
 | |
| 106 | by (rule negligible_subset) | |
| 107 | qed | |
| 108 |   have ac_ne: "cbox a ?c \<noteq> {}"
 | |
| 109 | using ab_ne an | |
| 110 | by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) | |
| 111 |   have ax_ne: "cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<noteq> {}"
 | |
| 112 | using ab_ne an | |
| 113 | by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) | |
| 114 | have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\<chi> i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" | |
| 115 | by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove | |
| 116 | if_distrib [of "\<lambda>u. u - z" for z] prod.remove) | |
| 117 | show ?Q | |
| 118 | using eq1 eq2 eq3 | |
| 119 | by (simp add: algebra_simps) | |
| 120 | qed | |
| 121 | ||
| 122 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 123 | proposition | 
| 67998 | 124 | fixes S :: "(real^'n) set" | 
| 125 | assumes "S \<in> lmeasurable" | |
| 126 | shows measurable_stretch: "((\<lambda>x. \<chi> k. m k * x$k) ` S) \<in> lmeasurable" (is "?f ` S \<in> _") | |
| 127 | and measure_stretch: "measure lebesgue ((\<lambda>x. \<chi> k. m k * x$k) ` S) = \<bar>prod m UNIV\<bar> * measure lebesgue S" | |
| 128 | (is "?MEQ") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 129 | proof - | 
| 67998 | 130 | have "(?f ` S) \<in> lmeasurable \<and> ?MEQ" | 
| 131 | proof (cases "\<forall>k. m k \<noteq> 0") | |
| 132 | case True | |
| 133 | have m0: "0 < \<bar>prod m UNIV\<bar>" | |
| 134 | using True by simp | |
| 135 | have "(indicat_real (?f ` S) has_integral \<bar>prod m UNIV\<bar> * measure lebesgue S) UNIV" | |
| 136 | proof (clarsimp simp add: has_integral_alt [where i=UNIV]) | |
| 137 | fix e :: "real" | |
| 138 | assume "e > 0" | |
| 139 | have "(indicat_real S has_integral (measure lebesgue S)) UNIV" | |
| 140 | using assms lmeasurable_iff_has_integral by blast | |
| 141 | then obtain B where "B>0" | |
| 142 | and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> | |
| 143 | \<exists>z. (indicat_real S has_integral z) (cbox a b) \<and> | |
| 144 | \<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>" | |
| 145 | by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \<open>e > 0\<close>) | |
| 146 | show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | |
| 147 | (\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox a b) \<and> | |
| 148 | \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)" | |
| 149 | proof (intro exI conjI allI) | |
| 150 | let ?C = "Max (range (\<lambda>k. \<bar>m k\<bar>)) * B" | |
| 151 | show "?C > 0" | |
| 152 | using True \<open>B > 0\<close> by (simp add: Max_gr_iff) | |
| 153 | show "ball 0 ?C \<subseteq> cbox u v \<longrightarrow> | |
| 154 | (\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and> | |
| 155 | \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)" for u v | |
| 156 | proof | |
| 157 | assume uv: "ball 0 ?C \<subseteq> cbox u v" | |
| 158 |           with \<open>?C > 0\<close> have cbox_ne: "cbox u v \<noteq> {}"
 | |
| 159 | using centre_in_ball by blast | |
| 160 | let ?\<alpha> = "\<lambda>k. u$k / m k" | |
| 161 | let ?\<beta> = "\<lambda>k. v$k / m k" | |
| 162 | have invm0: "\<And>k. inverse (m k) \<noteq> 0" | |
| 163 | using True by auto | |
| 164 | have "ball 0 B \<subseteq> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C" | |
| 165 | proof clarsimp | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo EmÃlio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67998diff
changeset | 166 | fix x :: "real^'n" | 
| 67998 | 167 | assume x: "norm x < B" | 
| 168 | have [simp]: "\<bar>Max (range (\<lambda>k. \<bar>m k\<bar>))\<bar> = Max (range (\<lambda>k. \<bar>m k\<bar>))" | |
| 169 | by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) | |
| 170 | have "norm (\<chi> k. m k * x $ k) \<le> norm (Max (range (\<lambda>k. \<bar>m k\<bar>)) *\<^sub>R x)" | |
| 171 | by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) | |
| 172 | also have "\<dots> < ?C" | |
| 72569 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
72445diff
changeset | 173 | using x \<open>0 < (MAX k. \<bar>m k\<bar>) * B\<close> \<open>0 < B\<close> zero_less_mult_pos2 by fastforce | 
| 67998 | 174 | finally have "norm (\<chi> k. m k * x $ k) < ?C" . | 
| 175 | then show "x \<in> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C" | |
| 176 | using stretch_Galois [of "inverse \<circ> m"] True by (auto simp: image_iff field_simps) | |
| 177 | qed | |
| 178 | then have Bsub: "ball 0 B \<subseteq> cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))" | |
| 179 | using cbox_ne uv image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric] | |
| 180 | by (force simp: field_simps) | |
| 181 | obtain z where zint: "(indicat_real S has_integral z) (cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))" | |
| 182 | and zless: "\<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>" | |
| 183 | using B [OF Bsub] by blast | |
| 184 | have ind: "indicat_real (?f ` S) = (\<lambda>x. indicator S (\<chi> k. x$k / m k))" | |
| 185 | using True stretch_Galois [of m] by (force simp: indicator_def) | |
| 186 | show "\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and> | |
| 187 | \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e" | |
| 188 | proof (simp add: ind, intro conjI exI) | |
| 189 | have "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>) | |
| 190 | ((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))" | |
| 191 | using True has_integral_stretch_cart [OF zint, of "inverse \<circ> m"] | |
| 192 | by (simp add: field_simps prod_dividef) | |
| 193 | moreover have "((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))) = cbox u v" | |
| 194 | using True image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric] | |
| 195 |                 image_stretch_interval_cart [of "\<lambda>k. 1" u v, symmetric] \<open>cbox u v \<noteq> {}\<close>
 | |
| 196 | by (simp add: field_simps image_comp o_def) | |
| 197 | ultimately show "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>) (cbox u v)" | |
| 198 | by simp | |
| 199 | have "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> | |
| 200 | = \<bar>prod m UNIV\<bar> * \<bar>z - measure lebesgue S\<bar>" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 201 | by (metis (no_types, opaque_lifting) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') | 
| 67998 | 202 | also have "\<dots> < e" | 
| 203 | using zless True by (simp add: field_simps) | |
| 204 | finally show "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e" . | |
| 205 | qed | |
| 206 | qed | |
| 207 | qed | |
| 208 | qed | |
| 209 | then show ?thesis | |
| 210 | by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) | |
| 211 | next | |
| 212 | case False | |
| 213 | then obtain k where "m k = 0" and prm: "prod m UNIV = 0" | |
| 214 | by auto | |
| 215 | have nfS: "negligible (?f ` S)" | |
| 216 | by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \<open>m k = 0\<close> in auto) | |
| 217 | then have "(?f ` S) \<in> lmeasurable" | |
| 218 | by (simp add: negligible_iff_measure) | |
| 219 | with nfS show ?thesis | |
| 220 | by (simp add: prm negligible_iff_measure0) | |
| 221 | qed | |
| 222 | then show "(?f ` S) \<in> lmeasurable" ?MEQ | |
| 223 | by metis+ | |
| 224 | qed | |
| 225 | ||
| 226 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 227 | proposition | 
| 67998 | 228 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 229 | assumes "linear f" "S \<in> lmeasurable" | |
| 230 | shows measurable_linear_image: "(f ` S) \<in> lmeasurable" | |
| 231 | and measure_linear_image: "measure lebesgue (f ` S) = \<bar>det (matrix f)\<bar> * measure lebesgue S" (is "?Q f S") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 232 | proof - | 
| 67998 | 233 | have "\<forall>S \<in> lmeasurable. (f ` S) \<in> lmeasurable \<and> ?Q f S" | 
| 234 | proof (rule induct_linear_elementary [OF \<open>linear f\<close>]; intro ballI) | |
| 235 | fix f g and S :: "(real,'n) vec set" | |
| 236 | assume "linear f" and "linear g" | |
| 237 | and f [rule_format]: "\<forall>S \<in> lmeasurable. f ` S \<in> lmeasurable \<and> ?Q f S" | |
| 238 | and g [rule_format]: "\<forall>S \<in> lmeasurable. g ` S \<in> lmeasurable \<and> ?Q g S" | |
| 239 | and S: "S \<in> lmeasurable" | |
| 240 | then have gS: "g ` S \<in> lmeasurable" | |
| 241 | by blast | |
| 242 | show "(f \<circ> g) ` S \<in> lmeasurable \<and> ?Q (f \<circ> g) S" | |
| 243 | using f [OF gS] g [OF S] matrix_compose [OF \<open>linear g\<close> \<open>linear f\<close>] | |
| 244 | by (simp add: o_def image_comp abs_mult det_mul) | |
| 245 | next | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo EmÃlio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67998diff
changeset | 246 | fix f :: "real^'n::_ \<Rightarrow> real^'n::_" and i and S :: "(real^'n::_) set" | 
| 67998 | 247 | assume "linear f" and 0: "\<And>x. f x $ i = 0" and "S \<in> lmeasurable" | 
| 248 | then have "\<not> inj f" | |
| 249 | by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) | |
| 250 | have detf: "det (matrix f) = 0" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 251 | using \<open>\<not> inj f\<close> det_nz_iff_inj[OF \<open>linear f\<close>] by blast | 
| 67998 | 252 | show "f ` S \<in> lmeasurable \<and> ?Q f S" | 
| 253 | proof | |
| 254 | show "f ` S \<in> lmeasurable" | |
| 255 | using lmeasurable_iff_indicator_has_integral \<open>linear f\<close> \<open>\<not> inj f\<close> negligible_UNIV negligible_linear_singular_image by blast | |
| 256 | have "measure lebesgue (f ` S) = 0" | |
| 257 | by (meson \<open>\<not> inj f\<close> \<open>linear f\<close> negligible_imp_measure0 negligible_linear_singular_image) | |
| 258 | also have "\<dots> = \<bar>det (matrix f)\<bar> * measure lebesgue S" | |
| 259 | by (simp add: detf) | |
| 260 | finally show "?Q f S" . | |
| 261 | qed | |
| 262 | next | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo EmÃlio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67998diff
changeset | 263 | fix c and S :: "(real^'n::_) set" | 
| 67998 | 264 | assume "S \<in> lmeasurable" | 
| 265 | show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable \<and> ?Q (\<lambda>a. \<chi> i. c i * a $ i) S" | |
| 266 | proof | |
| 267 | show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable" | |
| 268 | by (simp add: \<open>S \<in> lmeasurable\<close> measurable_stretch) | |
| 269 | show "?Q (\<lambda>a. \<chi> i. c i * a $ i) S" | |
| 270 | by (simp add: measure_stretch [OF \<open>S \<in> lmeasurable\<close>, of c] axis_def matrix_def det_diagonal) | |
| 271 | qed | |
| 272 | next | |
| 273 | fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" | |
| 274 | assume "m \<noteq> n" and "S \<in> lmeasurable" | |
| 73648 | 275 | let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. v $ Transposition.transpose m n i" | 
| 67998 | 276 | have lin: "linear ?h" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 277 | by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) | 
| 73648 | 278 | have meq: "measure lebesgue ((\<lambda>v::(real, 'n) vec. \<chi> i. v $ Transposition.transpose m n i) ` cbox a b) | 
| 67998 | 279 | = measure lebesgue (cbox a b)" for a b | 
| 280 |     proof (cases "cbox a b = {}")
 | |
| 281 | case True then show ?thesis | |
| 282 | by simp | |
| 283 | next | |
| 284 | case False | |
| 285 |       then have him: "?h ` (cbox a b) \<noteq> {}"
 | |
| 286 | by blast | |
| 287 | have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" | |
| 73648 | 288 | by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis transpose_involutory)+ | 
| 67998 | 289 | show ?thesis | 
| 290 | using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\<lambda>i. (b - a)$i", symmetric] | |
| 291 | by (simp add: eq content_cbox_cart False) | |
| 292 | qed | |
| 73648 | 293 | have "(\<chi> i j. if Transposition.transpose m n i = j then 1 else 0) = (\<chi> i j. if j = Transposition.transpose m n i then 1 else (0::real))" | 
| 67998 | 294 | by (auto intro!: Cart_lambda_cong) | 
| 73648 | 295 | then have "matrix ?h = transpose(\<chi> i j. mat 1 $ i $ Transposition.transpose m n j)" | 
| 67998 | 296 | by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) | 
| 297 | then have 1: "\<bar>det (matrix ?h)\<bar> = 1" | |
| 298 | by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) | |
| 299 | show "?h ` S \<in> lmeasurable \<and> ?Q ?h S" | |
| 300 | proof | |
| 301 | show "?h ` S \<in> lmeasurable" "?Q ?h S" | |
| 302 | using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force+ | |
| 303 | qed | |
| 304 | next | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo EmÃlio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67998diff
changeset | 305 | fix m n :: "'n" and S :: "(real, 'n) vec set" | 
| 67998 | 306 | assume "m \<noteq> n" and "S \<in> lmeasurable" | 
| 307 | let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. if i = m then v $ m + v $ n else v $ i" | |
| 308 | have lin: "linear ?h" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 309 | by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) | 
| 67998 | 310 | consider "m < n" | " n < m" | 
| 311 | using \<open>m \<noteq> n\<close> less_linear by blast | |
| 312 | then have 1: "det(matrix ?h) = 1" | |
| 313 | proof cases | |
| 314 | assume "m < n" | |
| 315 | have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n | |
| 316 | proof - | |
| 317 | have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))" | |
| 318 | using axis_def by blast | |
| 319 | then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" | |
| 320 | using \<open>j < i\<close> axis_def \<open>m < n\<close> by auto | |
| 321 | with \<open>m < n\<close> show ?thesis | |
| 322 | by (auto simp: matrix_def axis_def cong: if_cong) | |
| 323 | qed | |
| 324 | show ?thesis | |
| 325 | using \<open>m \<noteq> n\<close> by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) | |
| 326 | next | |
| 327 | assume "n < m" | |
| 328 | have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n | |
| 329 | proof - | |
| 330 | have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))" | |
| 331 | using axis_def by blast | |
| 332 | then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" | |
| 333 | using \<open>j > i\<close> axis_def \<open>m > n\<close> by auto | |
| 334 | with \<open>m > n\<close> show ?thesis | |
| 335 | by (auto simp: matrix_def axis_def cong: if_cong) | |
| 336 | qed | |
| 337 | show ?thesis | |
| 338 | using \<open>m \<noteq> n\<close> | |
| 339 | by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) | |
| 340 | qed | |
| 341 | have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b | |
| 342 |     proof (cases "cbox a b = {}")
 | |
| 343 | case True then show ?thesis by simp | |
| 344 | next | |
| 345 | case False | |
| 346 |       then have ne: "(+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b \<noteq> {}"
 | |
| 347 | by auto | |
| 348 | let ?v = "\<chi> i. if i = n then - a $ n else 0" | |
| 349 | have "?h ` cbox a b | |
| 350 | = (+) (\<chi> i. if i = m \<or> i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" | |
| 351 | using \<open>m \<noteq> n\<close> unfolding image_comp o_def by (force simp: vec_eq_iff) | |
| 352 | then have "measure lebesgue (?h ` (cbox a b)) | |
| 353 | = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) ` | |
| 354 | (+) ?v ` cbox a b)" | |
| 355 | by (rule ssubst) (rule measure_translation) | |
| 356 | also have "\<dots> = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" | |
| 357 | by (metis (no_types, lifting) cbox_translation) | |
| 358 | also have "\<dots> = measure lebesgue ((+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b)" | |
| 359 | apply (subst measure_shear_interval) | |
| 360 | using \<open>m \<noteq> n\<close> ne apply auto | |
| 361 | apply (simp add: cbox_translation) | |
| 362 | by (metis cbox_borel cbox_translation measure_completion sets_lborel) | |
| 363 | also have "\<dots> = measure lebesgue (cbox a b)" | |
| 364 | by (rule measure_translation) | |
| 365 | finally show ?thesis . | |
| 366 | qed | |
| 367 | show "?h ` S \<in> lmeasurable \<and> ?Q ?h S" | |
| 368 | using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force | |
| 369 | qed | |
| 370 | with assms show "(f ` S) \<in> lmeasurable" "?Q f S" | |
| 371 | by metis+ | |
| 372 | qed | |
| 373 | ||
| 374 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 375 | lemma | 
| 67998 | 376 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 377 | assumes f: "orthogonal_transformation f" and S: "S \<in> lmeasurable" | |
| 378 | shows measurable_orthogonal_image: "f ` S \<in> lmeasurable" | |
| 379 | and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" | |
| 380 | proof - | |
| 381 | have "linear f" | |
| 382 | by (simp add: f orthogonal_transformation_linear) | |
| 383 | then show "f ` S \<in> lmeasurable" | |
| 384 | by (metis S measurable_linear_image) | |
| 385 | show "measure lebesgue (f ` S) = measure lebesgue S" | |
| 386 | by (simp add: measure_linear_image \<open>linear f\<close> S f) | |
| 387 | qed | |
| 388 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 389 | proposition measure_semicontinuous_with_hausdist_explicit: | 
| 67998 | 390 | assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" | 
| 391 | obtains d where "d > 0" | |
| 392 | "\<And>T. \<lbrakk>T \<in> lmeasurable; \<And>y. y \<in> T \<Longrightarrow> \<exists>x. x \<in> S \<and> dist x y < d\<rbrakk> | |
| 393 | \<Longrightarrow> measure lebesgue T < measure lebesgue S + e" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 394 | proof (cases "S = {}")
 | 
| 67998 | 395 | case True | 
| 396 | with that \<open>e > 0\<close> show ?thesis by force | |
| 397 | next | |
| 398 | case False | |
| 399 |   then have frS: "frontier S \<noteq> {}"
 | |
| 400 | using \<open>bounded S\<close> frontier_eq_empty not_bounded_UNIV by blast | |
| 401 | have "S \<in> lmeasurable" | |
| 402 | by (simp add: \<open>bounded S\<close> measurable_Jordan neg) | |
| 403 | have null: "(frontier S) \<in> null_sets lebesgue" | |
| 404 | by (metis neg negligible_iff_null_sets) | |
| 405 | have "frontier S \<in> lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" | |
| 406 | using neg negligible_imp_measurable negligible_iff_measure by blast+ | |
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 407 | with \<open>e > 0\<close> sets_lebesgue_outer_open | 
| 67998 | 408 | obtain U where "open U" | 
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 409 | and U: "frontier S \<subseteq> U" "U - frontier S \<in> lmeasurable" "emeasure lebesgue (U - frontier S) < e" | 
| 67998 | 410 | by (metis fmeasurableD) | 
| 411 | with null have "U \<in> lmeasurable" | |
| 412 | by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) | |
| 413 | have "measure lebesgue (U - frontier S) = measure lebesgue U" | |
| 414 | using mS0 by (simp add: \<open>U \<in> lmeasurable\<close> fmeasurableD measure_Diff_null_set null) | |
| 415 | with U have mU: "measure lebesgue U < e" | |
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 416 | by (simp add: emeasure_eq_measure2 ennreal_less_iff) | 
| 67998 | 417 | show ?thesis | 
| 418 | proof | |
| 419 | have "U \<noteq> UNIV" | |
| 420 | using \<open>U \<in> lmeasurable\<close> by auto | |
| 421 |     then have "- U \<noteq> {}"
 | |
| 422 | by blast | |
| 423 | with \<open>open U\<close> \<open>frontier S \<subseteq> U\<close> show "setdist (frontier S) (- U) > 0" | |
| 424 | by (auto simp: \<open>bounded S\<close> open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) | |
| 425 | fix T | |
| 426 | assume "T \<in> lmeasurable" | |
| 427 | and T: "\<And>t. t \<in> T \<Longrightarrow> \<exists>y. y \<in> S \<and> dist y t < setdist (frontier S) (- U)" | |
| 428 | then have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue (T - S)" | |
| 429 | by (simp add: \<open>S \<in> lmeasurable\<close> measure_diff_le_measure_setdiff) | |
| 430 | also have "\<dots> \<le> measure lebesgue U" | |
| 431 | proof - | |
| 432 | have "T - S \<subseteq> U" | |
| 433 | proof clarify | |
| 434 | fix x | |
| 435 | assume "x \<in> T" and "x \<notin> S" | |
| 436 | then obtain y where "y \<in> S" and y: "dist y x < setdist (frontier S) (- U)" | |
| 437 | using T by blast | |
| 438 |         have "closed_segment x y \<inter> frontier S \<noteq> {}"
 | |
| 439 | using connected_Int_frontier \<open>x \<notin> S\<close> \<open>y \<in> S\<close> by blast | |
| 440 | then obtain z where z: "z \<in> closed_segment x y" "z \<in> frontier S" | |
| 441 | by auto | |
| 442 | with y have "dist z x < setdist(frontier S) (- U)" | |
| 443 | by (auto simp: dist_commute dest!: dist_in_closed_segment) | |
| 444 | with z have False if "x \<in> -U" | |
| 445 | using setdist_le_dist [OF \<open>z \<in> frontier S\<close> that] by auto | |
| 446 | then show "x \<in> U" | |
| 447 | by blast | |
| 448 | qed | |
| 449 | then show ?thesis | |
| 450 | by (simp add: \<open>S \<in> lmeasurable\<close> \<open>T \<in> lmeasurable\<close> \<open>U \<in> lmeasurable\<close> fmeasurableD measure_mono_fmeasurable sets.Diff) | |
| 451 | qed | |
| 452 | finally have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue U" . | |
| 453 | with mU show "measure lebesgue T < measure lebesgue S + e" | |
| 454 | by linarith | |
| 455 | qed | |
| 456 | qed | |
| 457 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 458 | proposition | 
| 67998 | 459 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 460 | assumes S: "S \<in> lmeasurable" | |
| 461 | and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 462 | and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S" | |
| 463 | and bounded: "\<And>x. x \<in> S \<Longrightarrow> \<bar>det (matrix (f' x))\<bar> \<le> B" | |
| 464 | shows measurable_bounded_differentiable_image: | |
| 465 | "f ` S \<in> lmeasurable" | |
| 466 | and measure_bounded_differentiable_image: | |
| 467 | "measure lebesgue (f ` S) \<le> B * measure lebesgue S" (is "?M") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 468 | proof - | 
| 67998 | 469 | have "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> B * measure lebesgue S" | 
| 470 | proof (cases "B < 0") | |
| 471 | case True | |
| 472 |     then have "S = {}"
 | |
| 473 | by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) | |
| 474 | then show ?thesis | |
| 475 | by auto | |
| 476 | next | |
| 477 | case False | |
| 478 | then have "B \<ge> 0" | |
| 479 | by arith | |
| 480 | let ?\<mu> = "measure lebesgue" | |
| 481 | have f_diff: "f differentiable_on S" | |
| 482 | using deriv by (auto simp: differentiable_on_def differentiable_def) | |
| 483 | have eps: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> (B+e) * ?\<mu> S" (is "?ME") | |
| 484 | if "e > 0" for e | |
| 485 | proof - | |
| 486 | have eps_d: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> (B+e) * (?\<mu> S + d)" (is "?MD") | |
| 487 | if "d > 0" for d | |
| 488 | proof - | |
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 489 | obtain T where T: "open T" "S \<subseteq> T" and TS: "(T-S) \<in> lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 490 | using S \<open>d > 0\<close> sets_lebesgue_outer_open by blast | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 491 | then have "?\<mu> (T-S) < d" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 492 | by (metis emeasure_eq_measure2 ennreal_leI not_less) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 493 | with S T TS have "T \<in> lmeasurable" and Tless: "?\<mu> T < ?\<mu> S + d" | 
| 67998 | 494 | by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) | 
| 495 | have "\<exists>r. 0 < r \<and> r < d \<and> ball x r \<subseteq> T \<and> f ` (S \<inter> ball x r) \<in> lmeasurable \<and> | |
| 496 | ?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)" | |
| 497 | if "x \<in> S" "d > 0" for x d | |
| 498 | proof - | |
| 499 | have lin: "linear (f' x)" | |
| 500 | and lim0: "((\<lambda>y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \<longlongrightarrow> 0) (at x within S)" | |
| 501 | using deriv \<open>x \<in> S\<close> by (auto simp: has_derivative_within bounded_linear.linear field_simps) | |
| 502 | have bo: "bounded (f' x ` ball 0 1)" | |
| 503 | by (simp add: bounded_linear_image linear_linear lin) | |
| 504 | have neg: "negligible (frontier (f' x ` ball 0 1))" | |
| 505 | using deriv has_derivative_linear \<open>x \<in> S\<close> | |
| 506 | by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) | |
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 507 |           let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 508 | have 0: "0 < e * ?unit_vol" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 509 | using \<open>e > 0\<close> by (simp add: content_ball_pos) | 
| 67998 | 510 | obtain k where "k > 0" and k: | 
| 511 | "\<And>U. \<lbrakk>U \<in> lmeasurable; \<And>y. y \<in> U \<Longrightarrow> \<exists>z. z \<in> f' x ` ball 0 1 \<and> dist z y < k\<rbrakk> | |
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 512 | \<Longrightarrow> ?\<mu> U < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol" | 
| 67998 | 513 | using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast | 
| 514 | obtain l where "l > 0" and l: "ball x l \<subseteq> T" | |
| 515 | using \<open>x \<in> S\<close> \<open>open T\<close> \<open>S \<subseteq> T\<close> openE by blast | |
| 516 | obtain \<zeta> where "0 < \<zeta>" | |
| 517 | and \<zeta>: "\<And>y. \<lbrakk>y \<in> S; y \<noteq> x; dist y x < \<zeta>\<rbrakk> | |
| 518 | \<Longrightarrow> norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70724diff
changeset | 519 | using lim0 \<open>k > 0\<close> by (simp add: Lim_within) (auto simp add: field_simps) | 
| 67998 | 520 | define r where "r \<equiv> min (min l (\<zeta>/2)) (min 1 (d/2))" | 
| 521 | show ?thesis | |
| 522 | proof (intro exI conjI) | |
| 523 | show "r > 0" "r < d" | |
| 524 | using \<open>l > 0\<close> \<open>\<zeta> > 0\<close> \<open>d > 0\<close> by (auto simp: r_def) | |
| 525 | have "r \<le> l" | |
| 526 | by (auto simp: r_def) | |
| 527 | with l show "ball x r \<subseteq> T" | |
| 528 | by auto | |
| 529 | have ex_lessK: "\<exists>x' \<in> ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" | |
| 530 | if "y \<in> S" and "dist x y < r" for y | |
| 531 | proof (cases "y = x") | |
| 532 | case True | |
| 533 | with lin linear_0 \<open>k > 0\<close> that show ?thesis | |
| 534 | by (rule_tac x=0 in bexI) (auto simp: linear_0) | |
| 535 | next | |
| 536 | case False | |
| 537 | then show ?thesis | |
| 538 | proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) | |
| 539 | have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 540 | by (simp add: lin linear_scale) | 
| 67998 | 541 | then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" | 
| 542 | by (simp add: dist_norm) | |
| 543 | also have "\<dots> = norm (f' x (y - x) - (f y - f x)) / r" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 544 | using \<open>r > 0\<close> by (simp add: divide_simps scale_right_diff_distrib [symmetric]) | 
| 67998 | 545 | also have "\<dots> \<le> norm (f y - (f x + f' x (y - x))) / norm (y - x)" | 
| 71633 | 546 | using that \<open>r > 0\<close> False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono) | 
| 67998 | 547 | also have "\<dots> < k" | 
| 548 | using that \<open>0 < \<zeta>\<close> by (simp add: dist_commute r_def \<zeta> [OF \<open>y \<in> S\<close> False]) | |
| 549 | finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . | |
| 550 | show "(y - x) /\<^sub>R r \<in> ball 0 1" | |
| 551 | using that \<open>r > 0\<close> by (simp add: dist_norm divide_simps norm_minus_commute) | |
| 552 | qed | |
| 553 | qed | |
| 554 | let ?rfs = "(\<lambda>x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \<inter> ball x r)" | |
| 555 | have rfs_mble: "?rfs \<in> lmeasurable" | |
| 556 | proof (rule bounded_set_imp_lmeasurable) | |
| 557 | have "f differentiable_on S \<inter> ball x r" | |
| 558 | using f_diff by (auto simp: fmeasurableD differentiable_on_subset) | |
| 559 | with S show "?rfs \<in> sets lebesgue" | |
| 560 | by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) | |
| 561 | let ?B = "(\<lambda>(x, y). x + y) ` (f' x ` ball 0 1 \<times> ball 0 k)" | |
| 562 | have "bounded ?B" | |
| 563 | by (simp add: bounded_plus [OF bo]) | |
| 564 | moreover have "?rfs \<subseteq> ?B" | |
| 565 | apply (auto simp: dist_norm image_iff dest!: ex_lessK) | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 566 | by (metis (no_types, opaque_lifting) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) | 
| 67998 | 567 | ultimately show "bounded (?rfs)" | 
| 568 | by (rule bounded_subset) | |
| 569 | qed | |
| 570 | then have "(\<lambda>x. r *\<^sub>R x) ` ?rfs \<in> lmeasurable" | |
| 571 | by (simp add: measurable_linear_image) | |
| 572 | with \<open>r > 0\<close> have "(+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable" | |
| 573 | by (simp add: image_comp o_def) | |
| 574 | then have "(+) (f x) ` (+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable" | |
| 575 | using measurable_translation by blast | |
| 576 | then show fsb: "f ` (S \<inter> ball x r) \<in> lmeasurable" | |
| 577 | by (simp add: image_comp o_def) | |
| 578 |             have "?\<mu> (f ` (S \<inter> ball x r)) = ?\<mu> (?rfs) * r ^ CARD('n)"
 | |
| 69661 | 579 | using \<open>r > 0\<close> fsb | 
| 580 | by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) | |
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 581 |             also have "\<dots> \<le> (\<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)"
 | 
| 67998 | 582 | proof - | 
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 583 | have "?\<mu> (?rfs) < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol" | 
| 67998 | 584 | using rfs_mble by (force intro: k dest!: ex_lessK) | 
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 585 | then have "?\<mu> (?rfs) < \<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 586 | by (simp add: lin measure_linear_image [of "f' x"]) | 
| 67998 | 587 | with \<open>r > 0\<close> show ?thesis | 
| 588 | by auto | |
| 589 | qed | |
| 590 | also have "\<dots> \<le> (B + e) * ?\<mu> (ball x r)" | |
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 591 | using bounded [OF \<open>x \<in> S\<close>] \<open>r > 0\<close> | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 592 | by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos) | 
| 67998 | 593 | finally show "?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)" . | 
| 594 | qed | |
| 595 | qed | |
| 596 | then obtain r where | |
| 597 | r0d: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> 0 < r x d \<and> r x d < d" | |
| 598 | and rT: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> ball x (r x d) \<subseteq> T" | |
| 599 | and r: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> | |
| 600 | (f ` (S \<inter> ball x (r x d))) \<in> lmeasurable \<and> | |
| 601 | ?\<mu> (f ` (S \<inter> ball x (r x d))) \<le> (B + e) * ?\<mu> (ball x (r x d))" | |
| 602 | by metis | |
| 603 |         obtain C where "countable C" and Csub: "C \<subseteq> {(x,r x t) |x t. x \<in> S \<and> 0 < t}"
 | |
| 604 | and pwC: "pairwise (\<lambda>i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" | |
| 605 | and negC: "negligible(S - (\<Union>i \<in> C. ball (fst i) (snd i)))" | |
| 606 |           apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \<in> S \<and> 0 < t}" fst snd])
 | |
| 607 | apply auto | |
| 608 | by (metis dist_eq_0_iff r0d) | |
| 609 | let ?UB = "(\<Union>(x,s) \<in> C. ball x s)" | |
| 610 | have eq: "f ` (S \<inter> ?UB) = (\<Union>(x,s) \<in> C. f ` (S \<inter> ball x s))" | |
| 611 | by auto | |
| 612 | have mle: "?\<mu> (\<Union>(x,s) \<in> K. f ` (S \<inter> ball x s)) \<le> (B + e) * (?\<mu> S + d)" (is "?l \<le> ?r") | |
| 613 | if "K \<subseteq> C" and "finite K" for K | |
| 614 | proof - | |
| 615 | have gt0: "b > 0" if "(a, b) \<in> K" for a b | |
| 616 | using Csub that \<open>K \<subseteq> C\<close> r0d by auto | |
| 617 | have inj: "inj_on (\<lambda>(x, y). ball x y) K" | |
| 618 | by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) | |
| 619 | have disjnt: "disjoint ((\<lambda>(x, y). ball x y) ` K)" | |
| 620 | using pwC that | |
| 621 | apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) | |
| 622 | by (metis subsetD fst_conv snd_conv) | |
| 623 | have "?l \<le> (\<Sum>i\<in>K. ?\<mu> (case i of (x, s) \<Rightarrow> f ` (S \<inter> ball x s)))" | |
| 624 | proof (rule measure_UNION_le [OF \<open>finite K\<close>], clarify) | |
| 625 | fix x r | |
| 626 | assume "(x,r) \<in> K" | |
| 627 | then have "x \<in> S" | |
| 628 | using Csub \<open>K \<subseteq> C\<close> by auto | |
| 629 | show "f ` (S \<inter> ball x r) \<in> sets lebesgue" | |
| 630 | by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) | |
| 631 | qed | |
| 632 | also have "\<dots> \<le> (\<Sum>(x,s) \<in> K. (B + e) * ?\<mu> (ball x s))" | |
| 633 | apply (rule sum_mono) | |
| 634 | using Csub r \<open>K \<subseteq> C\<close> by auto | |
| 635 | also have "\<dots> = (B + e) * (\<Sum>(x,s) \<in> K. ?\<mu> (ball x s))" | |
| 636 | by (simp add: prod.case_distrib sum_distrib_left) | |
| 637 | also have "\<dots> = (B + e) * sum ?\<mu> ((\<lambda>(x, y). ball x y) ` K)" | |
| 638 | using \<open>B \<ge> 0\<close> \<open>e > 0\<close> by (simp add: inj sum.reindex prod.case_distrib) | |
| 639 | also have "\<dots> = (B + e) * ?\<mu> (\<Union>(x,s) \<in> K. ball x s)" | |
| 640 | using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that | |
| 641 | by (subst measure_Union') (auto simp: disjnt measure_Union') | |
| 642 | also have "\<dots> \<le> (B + e) * ?\<mu> T" | |
| 643 | using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that apply simp | |
| 644 | apply (rule measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>]) | |
| 645 | using Csub rT by force+ | |
| 646 | also have "\<dots> \<le> (B + e) * (?\<mu> S + d)" | |
| 647 | using \<open>B \<ge> 0\<close> \<open>e > 0\<close> Tless by simp | |
| 648 | finally show ?thesis . | |
| 649 | qed | |
| 650 | have fSUB_mble: "(f ` (S \<inter> ?UB)) \<in> lmeasurable" | |
| 651 | unfolding eq using Csub r False \<open>e > 0\<close> that | |
| 652 | by (auto simp: intro!: fmeasurable_UN_bound [OF \<open>countable C\<close> _ mle]) | |
| 653 | have fSUB_meas: "?\<mu> (f ` (S \<inter> ?UB)) \<le> (B + e) * (?\<mu> S + d)" (is "?MUB") | |
| 654 | unfolding eq using Csub r False \<open>e > 0\<close> that | |
| 655 | by (auto simp: intro!: measure_UN_bound [OF \<open>countable C\<close> _ mle]) | |
| 656 | have neg: "negligible ((f ` (S \<inter> ?UB) - f ` S) \<union> (f ` S - f ` (S \<inter> ?UB)))" | |
| 657 | proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) | |
| 658 | show "f differentiable_on S - (\<Union>i\<in>C. ball (fst i) (snd i))" | |
| 659 | by (meson DiffE differentiable_on_subset subsetI f_diff) | |
| 660 | qed force | |
| 661 | show "f ` S \<in> lmeasurable" | |
| 662 | by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) | |
| 663 | show ?MD | |
| 664 | using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp | |
| 665 | qed | |
| 666 | show "f ` S \<in> lmeasurable" | |
| 667 | using eps_d [of 1] by simp | |
| 668 | show ?ME | |
| 669 | proof (rule field_le_epsilon) | |
| 670 | fix \<delta> :: real | |
| 671 | assume "0 < \<delta>" | |
| 672 | then show "?\<mu> (f ` S) \<le> (B + e) * ?\<mu> S + \<delta>" | |
| 673 | using eps_d [of "\<delta> / (B+e)"] \<open>e > 0\<close> \<open>B \<ge> 0\<close> by (auto simp: divide_simps mult_ac) | |
| 674 | qed | |
| 675 | qed | |
| 676 | show ?thesis | |
| 677 | proof (cases "?\<mu> S = 0") | |
| 678 | case True | |
| 679 | with eps have "?\<mu> (f ` S) = 0" | |
| 680 | by (metis mult_zero_right not_le zero_less_measure_iff) | |
| 681 | then show ?thesis | |
| 682 | using eps [of 1] by (simp add: True) | |
| 683 | next | |
| 684 | case False | |
| 685 | have "?\<mu> (f ` S) \<le> B * ?\<mu> S" | |
| 686 | proof (rule field_le_epsilon) | |
| 687 | fix e :: real | |
| 688 | assume "e > 0" | |
| 689 | then show "?\<mu> (f ` S) \<le> B * ?\<mu> S + e" | |
| 690 | using eps [of "e / ?\<mu> S"] False by (auto simp: algebra_simps zero_less_measure_iff) | |
| 691 | qed | |
| 692 | with eps [of 1] show ?thesis by auto | |
| 693 | qed | |
| 694 | qed | |
| 695 | then show "f ` S \<in> lmeasurable" ?M by blast+ | |
| 696 | qed | |
| 697 | ||
| 69739 | 698 | lemma m_diff_image_weak: | 
| 67998 | 699 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 700 | assumes S: "S \<in> lmeasurable" | |
| 701 | and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 702 | and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S" | |
| 69739 | 703 | shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 704 | proof - | 
| 67998 | 705 | let ?\<mu> = "measure lebesgue" | 
| 706 | have aint_S: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S" | |
| 707 | using int unfolding absolutely_integrable_on_def by auto | |
| 708 | define m where "m \<equiv> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 709 | have *: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S" | |
| 710 | if "e > 0" for e | |
| 711 | proof - | |
| 712 |     define T where "T \<equiv> \<lambda>n. {x \<in> S. n * e \<le> \<bar>det (matrix (f' x))\<bar> \<and>
 | |
| 713 | \<bar>det (matrix (f' x))\<bar> < (Suc n) * e}" | |
| 714 | have meas_t: "T n \<in> lmeasurable" for n | |
| 715 | proof - | |
| 716 | have *: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<in> borel_measurable (lebesgue_on S)" | |
| 717 | using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) | |
| 718 | have [intro]: "x \<in> sets (lebesgue_on S) \<Longrightarrow> x \<in> sets lebesgue" for x | |
| 719 | using S sets_restrict_space_subset by blast | |
| 720 |       have "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> sets lebesgue"
 | |
| 721 | using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) | |
| 722 |       then have 1: "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> lmeasurable"
 | |
| 723 | using S by (simp add: fmeasurableI2) | |
| 724 |       have "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> sets lebesgue"
 | |
| 725 | using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) | |
| 726 |       then have 2: "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> lmeasurable"
 | |
| 727 | using S by (simp add: fmeasurableI2) | |
| 728 | show ?thesis | |
| 729 | using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) | |
| 730 | qed | |
| 731 | have aint_T: "\<And>k. (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on T k" | |
| 732 | using set_integrable_subset [OF aint_S] meas_t T_def by blast | |
| 733 | have Seq: "S = (\<Union>n. T n)" | |
| 734 | apply (auto simp: T_def) | |
| 735 | apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) | |
| 736 | using that apply auto | |
| 737 | using of_int_floor_le pos_le_divide_eq apply blast | |
| 738 | by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) | |
| 739 | have meas_ft: "f ` T n \<in> lmeasurable" for n | |
| 740 | proof (rule measurable_bounded_differentiable_image) | |
| 741 | show "T n \<in> lmeasurable" | |
| 742 | by (simp add: meas_t) | |
| 743 | next | |
| 744 | fix x :: "(real,'n) vec" | |
| 745 | assume "x \<in> T n" | |
| 746 | show "(f has_derivative f' x) (at x within T n)" | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 747 | by (metis (no_types, lifting) \<open>x \<in> T n\<close> deriv has_derivative_subset mem_Collect_eq subsetI T_def) | 
| 67998 | 748 | show "\<bar>det (matrix (f' x))\<bar> \<le> (Suc n) * e" | 
| 749 | using \<open>x \<in> T n\<close> T_def by auto | |
| 750 | next | |
| 751 | show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T n" | |
| 752 | using aint_T absolutely_integrable_on_def by blast | |
| 753 | qed | |
| 754 | have disT: "disjoint (range T)" | |
| 755 | unfolding disjoint_def | |
| 756 | proof clarsimp | |
| 757 |       show "T m \<inter> T n = {}" if "T m \<noteq> T n" for m n
 | |
| 758 | using that | |
| 759 | proof (induction m n rule: linorder_less_wlog) | |
| 760 | case (less m n) | |
| 761 | with \<open>e > 0\<close> show ?case | |
| 762 | unfolding T_def | |
| 763 | proof (clarsimp simp add: Collect_conj_eq [symmetric]) | |
| 764 | fix x | |
| 765 | assume "e > 0" "m < n" "n * e \<le> \<bar>det (matrix (f' x))\<bar>" "\<bar>det (matrix (f' x))\<bar> < (1 + real m) * e" | |
| 766 | then have "n < 1 + real m" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 767 | by (metis (no_types, opaque_lifting) less_le_trans mult.commute not_le mult_le_cancel_iff2) | 
| 67998 | 768 | then show "False" | 
| 769 | using less.hyps by linarith | |
| 770 | qed | |
| 771 | qed auto | |
| 772 | qed | |
| 773 |     have injT: "inj_on T ({n. T n \<noteq> {}})"
 | |
| 774 | unfolding inj_on_def | |
| 775 | proof clarsimp | |
| 776 |       show "m = n" if "T m = T n" "T n \<noteq> {}" for m n
 | |
| 777 | using that | |
| 778 | proof (induction m n rule: linorder_less_wlog) | |
| 779 | case (less m n) | |
| 780 | have False if "T n \<subseteq> T m" "x \<in> T n" for x | |
| 781 | using \<open>e > 0\<close> \<open>m < n\<close> that | |
| 782 | apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) | |
| 72569 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
72445diff
changeset | 783 | by (metis add.commute less_le_trans nat_less_real_le not_le mult_le_cancel_iff2) | 
| 67998 | 784 | then show ?case | 
| 785 | using less.prems by blast | |
| 786 | qed auto | |
| 787 | qed | |
| 788 |     have sum_eq_Tim: "(\<Sum>k\<le>n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \<Rightarrow> real" and n
 | |
| 789 | proof (subst sum.reindex_nontrivial) | |
| 790 |       fix i j  assume "i \<in> {..n}" "j \<in> {..n}" "i \<noteq> j" "T i = T j"
 | |
| 791 | with that injT [unfolded inj_on_def] show "f (T i) = 0" | |
| 792 | by simp metis | |
| 793 | qed (use atMost_atLeast0 in auto) | |
| 794 | let ?B = "m + e * ?\<mu> S" | |
| 795 | have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" for n | |
| 796 | proof - | |
| 797 | have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> (\<Sum>k\<le>n. ((k+1) * e) * ?\<mu>(T k))" | |
| 798 | proof (rule sum_mono [OF measure_bounded_differentiable_image]) | |
| 799 | show "(f has_derivative f' x) (at x within T k)" if "x \<in> T k" for k x | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 800 | using that unfolding T_def by (blast intro: deriv has_derivative_subset) | 
| 67998 | 801 | show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k" for k | 
| 802 | using absolutely_integrable_on_def aint_T by blast | |
| 803 | show "\<bar>det (matrix (f' x))\<bar> \<le> real (k + 1) * e" if "x \<in> T k" for k x | |
| 804 | using T_def that by auto | |
| 805 | qed (use meas_t in auto) | |
| 806 | also have "\<dots> \<le> (\<Sum>k\<le>n. (k * e) * ?\<mu>(T k)) + (\<Sum>k\<le>n. e * ?\<mu>(T k))" | |
| 807 | by (simp add: algebra_simps sum.distrib) | |
| 808 | also have "\<dots> \<le> ?B" | |
| 809 | proof (rule add_mono) | |
| 810 | have "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) = (\<Sum>k\<le>n. integral (T k) (\<lambda>x. k * e))" | |
| 811 | by (simp add: lmeasure_integral [OF meas_t] | |
| 68403 | 812 | flip: integral_mult_right integral_mult_left) | 
| 67998 | 813 | also have "\<dots> \<le> (\<Sum>k\<le>n. integral (T k) (\<lambda>x. (abs (det (matrix (f' x))))))" | 
| 814 | proof (rule sum_mono) | |
| 815 | fix k | |
| 816 |           assume "k \<in> {..n}"
 | |
| 817 | show "integral (T k) (\<lambda>x. k * e) \<le> integral (T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 818 | proof (rule integral_le [OF integrable_on_const [OF meas_t]]) | |
| 819 | show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k" | |
| 820 | using absolutely_integrable_on_def aint_T by blast | |
| 821 | next | |
| 822 | fix x assume "x \<in> T k" | |
| 823 | show "k * e \<le> \<bar>det (matrix (f' x))\<bar>" | |
| 824 | using \<open>x \<in> T k\<close> T_def by blast | |
| 825 | qed | |
| 826 | qed | |
| 827 |         also have "\<dots> = sum (\<lambda>T. integral T (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) (T ` {..n})"
 | |
| 828 | by (auto intro: sum_eq_Tim) | |
| 829 | also have "\<dots> = integral (\<Union>k\<le>n. T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 830 | proof (rule integral_unique [OF has_integral_Union, symmetric]) | |
| 831 |           fix S  assume "S \<in> T ` {..n}"
 | |
| 832 | then show "((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) S" | |
| 833 | using absolutely_integrable_on_def aint_T by blast | |
| 834 | next | |
| 835 |           show "pairwise (\<lambda>S S'. negligible (S \<inter> S')) (T ` {..n})"
 | |
| 836 | using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) | |
| 837 | qed auto | |
| 838 | also have "\<dots> \<le> m" | |
| 839 | unfolding m_def | |
| 840 | proof (rule integral_subset_le) | |
| 841 | have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on (\<Union>k\<le>n. T k)" | |
| 842 | apply (rule set_integrable_subset [OF aint_S]) | |
| 843 | apply (intro measurable meas_t fmeasurableD) | |
| 844 | apply (force simp: Seq) | |
| 845 | done | |
| 846 | then show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on (\<Union>k\<le>n. T k)" | |
| 847 | using absolutely_integrable_on_def by blast | |
| 848 | qed (use Seq int in auto) | |
| 849 | finally show "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) \<le> m" . | |
| 850 | next | |
| 851 |         have "(\<Sum>k\<le>n. ?\<mu> (T k)) = sum ?\<mu> (T ` {..n})"
 | |
| 852 | by (auto intro: sum_eq_Tim) | |
| 853 | also have "\<dots> = ?\<mu> (\<Union>k\<le>n. T k)" | |
| 854 | using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) | |
| 855 | also have "\<dots> \<le> ?\<mu> S" | |
| 856 | using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) | |
| 857 | finally have "(\<Sum>k\<le>n. ?\<mu> (T k)) \<le> ?\<mu> S" . | |
| 858 | then show "(\<Sum>k\<le>n. e * ?\<mu> (T k)) \<le> e * ?\<mu> S" | |
| 859 | by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) | |
| 860 | qed | |
| 861 | finally show "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" . | |
| 862 | qed | |
| 863 | moreover have "measure lebesgue (\<Union>k\<le>n. f ` T k) \<le> (\<Sum>k\<le>n. ?\<mu> (f ` T k))" for n | |
| 864 | by (simp add: fmeasurableD meas_ft measure_UNION_le) | |
| 865 | ultimately have B_ge_m: "?\<mu> (\<Union>k\<le>n. (f ` T k)) \<le> ?B" for n | |
| 866 | by (meson order_trans) | |
| 867 | have "(\<Union>n. f ` T n) \<in> lmeasurable" | |
| 868 | by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) | |
| 869 | moreover have "?\<mu> (\<Union>n. f ` T n) \<le> m + e * ?\<mu> S" | |
| 870 | by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) | |
| 871 | ultimately show "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S" | |
| 872 | by (auto simp: Seq image_Union) | |
| 873 | qed | |
| 874 | show ?thesis | |
| 875 | proof | |
| 876 | show "f ` S \<in> lmeasurable" | |
| 877 | using * linordered_field_no_ub by blast | |
| 878 | let ?x = "m - ?\<mu> (f ` S)" | |
| 879 | have False if "?\<mu> (f ` S) > integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 880 | proof - | |
| 881 | have ml: "m < ?\<mu> (f ` S)" | |
| 882 | using m_def that by blast | |
| 883 | then have "?\<mu> S \<noteq> 0" | |
| 884 | using "*"(2) bgauge_existence_lemma by fastforce | |
| 885 | with ml have 0: "0 < - (m - ?\<mu> (f ` S))/2 / ?\<mu> S" | |
| 886 | using that zero_less_measure_iff by force | |
| 887 | then show ?thesis | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 888 | using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) | 
| 67998 | 889 | qed | 
| 890 | then show "?\<mu> (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 891 | by fastforce | |
| 892 | qed | |
| 893 | qed | |
| 894 | ||
| 895 | ||
| 69739 | 896 | theorem | 
| 67998 | 897 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 898 | assumes S: "S \<in> sets lebesgue" | |
| 899 | and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 900 | and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S" | |
| 901 | shows measurable_differentiable_image: "f ` S \<in> lmeasurable" | |
| 902 | and measure_differentiable_image: | |
| 903 | "measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is "?M") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 904 | proof - | 
| 67998 | 905 | let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S" | 
| 906 | let ?\<mu> = "measure lebesgue" | |
| 907 | have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_" | |
| 908 | apply (auto simp: mem_box_cart) | |
| 909 | apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) | |
| 910 | by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) | |
| 911 | then have Seq: "S = (\<Union>n. ?I n)" | |
| 912 | by auto | |
| 913 | have fIn: "f ` ?I n \<in> lmeasurable" | |
| 914 | and mfIn: "?\<mu> (f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is ?MN) for n | |
| 915 | proof - | |
| 916 | have In: "?I n \<in> lmeasurable" | |
| 917 | by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) | |
| 918 | moreover have "\<And>x. x \<in> ?I n \<Longrightarrow> (f has_derivative f' x) (at x within ?I n)" | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 919 | by (meson Int_iff deriv has_derivative_subset subsetI) | 
| 67998 | 920 | moreover have int_In: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on ?I n" | 
| 921 | proof - | |
| 922 | have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S" | |
| 923 | using int absolutely_integrable_integrable_bound by force | |
| 924 | then have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on ?I n" | |
| 925 | by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) | |
| 926 | then show ?thesis | |
| 927 | using absolutely_integrable_on_def by blast | |
| 928 | qed | |
| 929 | ultimately have "f ` ?I n \<in> lmeasurable" "?\<mu> (f ` ?I n) \<le> integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 930 | using m_diff_image_weak by metis+ | |
| 931 | moreover have "integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 932 | by (simp add: int_In int integral_subset_le) | |
| 933 | ultimately show "f ` ?I n \<in> lmeasurable" ?MN | |
| 934 | by auto | |
| 935 | qed | |
| 936 | have "?I k \<subseteq> ?I n" if "k \<le> n" for k n | |
| 937 | by (rule Int_mono) (use that in \<open>auto simp: subset_interval_imp_cart\<close>) | |
| 938 | then have "(\<Union>k\<le>n. f ` ?I k) = f ` ?I n" for n | |
| 939 | by (fastforce simp add:) | |
| 940 | with mfIn have "?\<mu> (\<Union>k\<le>n. f ` ?I k) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" for n | |
| 941 | by simp | |
| 942 | then have "(\<Union>n. f ` ?I n) \<in> lmeasurable" "?\<mu> (\<Union>n. f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 943 | by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ | |
| 944 | then show "f ` S \<in> lmeasurable" ?M | |
| 945 | by (metis Seq image_UN)+ | |
| 946 | qed | |
| 947 | ||
| 948 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 949 | lemma borel_measurable_simple_function_limit_increasing: | 
| 67998 | 950 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 951 | shows "(f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> f x)) \<longleftrightarrow> | |
| 952 | (\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> f x) \<and> (\<forall>n x. g n x \<le> (g(Suc n) x)) \<and> | |
| 953 | (\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and> | |
| 954 | (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x))" | |
| 955 | (is "?lhs = ?rhs") | |
| 956 | proof | |
| 957 | assume f: ?lhs | |
| 958 |   have leb_f: "{x. a \<le> f x \<and> f x < b} \<in> sets lebesgue" for a b
 | |
| 959 | proof - | |
| 960 |     have "{x. a \<le> f x \<and> f x < b} = {x. f x < b} - {x. f x < a}"
 | |
| 961 | by auto | |
| 962 | also have "\<dots> \<in> sets lebesgue" | |
| 963 | using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto | |
| 964 | finally show ?thesis . | |
| 965 | qed | |
| 966 | have "g n x \<le> f x" | |
| 967 | if inc_g: "\<And>n x. 0 \<le> g n x \<and> g n x \<le> g (Suc n) x" | |
| 968 | and meas_g: "\<And>n. g n \<in> borel_measurable lebesgue" | |
| 969 | and fin: "\<And>n. finite(range (g n))" and lim: "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" for g n x | |
| 970 | proof - | |
| 971 | have "\<exists>r>0. \<forall>N. \<exists>n\<ge>N. dist (g n x) (f x) \<ge> r" if "g n x > f x" | |
| 972 | proof - | |
| 973 | have g: "g n x \<le> g (N + n) x" for N | |
| 974 | by (rule transitive_stepwise_le) (use inc_g in auto) | |
| 975 | have "\<exists>na\<ge>N. g n x - f x \<le> dist (g na x) (f x)" for N | |
| 976 | apply (rule_tac x="N+n" in exI) | |
| 977 | using g [of N] by (auto simp: dist_norm) | |
| 978 | with that show ?thesis | |
| 979 | using diff_gt_0_iff_gt by blast | |
| 980 | qed | |
| 981 | with lim show ?thesis | |
| 982 | apply (auto simp: lim_sequentially) | |
| 983 | by (meson less_le_not_le not_le_imp_less) | |
| 984 | qed | |
| 985 | moreover | |
| 986 |   let ?\<Omega> = "\<lambda>n k. indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n}"
 | |
| 987 | let ?g = "\<lambda>n x. (\<Sum>k::real | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x)" | |
| 988 | have "\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> (g(Suc n) x)) \<and> | |
| 989 | (\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and>(\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x)" | |
| 990 | proof (intro exI allI conjI) | |
| 991 | show "0 \<le> ?g n x" for n x | |
| 992 | proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) | |
| 993 | fix k::real | |
| 994 | assume "k \<in> \<int>" and k: "\<bar>k\<bar> \<le> 2 ^ (2*n)" | |
| 995 | show "0 \<le> k/2^n * ?\<Omega> n k x" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 996 | using f \<open>k \<in> \<int>\<close> apply (auto simp: indicator_def field_split_simps Ints_def) | 
| 67998 | 997 | apply (drule spec [where x=x]) | 
| 998 | using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] | |
| 999 | by linarith | |
| 1000 | qed | |
| 1001 | show "?g n x \<le> ?g (Suc n) x" for n x | |
| 1002 | proof - | |
| 1003 | have "?g n x = | |
| 1004 | (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). | |
| 1005 |               k/2^n * (indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x +
 | |
| 1006 |               indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x))"
 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1007 | by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) | 
| 67998 | 1008 |       also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x) +
 | 
| 1009 |                        (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x)"
 | |
| 1010 | by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) | |
| 1011 |       also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1)/2 ^ Suc n} x) +
 | |
| 1012 |                        (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2 * k+1) + 1)/2 ^ Suc n} x)"
 | |
| 1013 | by (force simp: field_simps indicator_def intro: sum.cong) | |
| 1014 |       also have "\<dots> \<le> (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x))"
 | |
| 1015 | (is "?a + _ \<le> ?b") | |
| 1016 | proof - | |
| 1017 | have *: "\<lbrakk>sum f I \<le> sum h I; a + sum h I \<le> b\<rbrakk> \<Longrightarrow> a + sum f I \<le> b" for I a b f and h :: "real\<Rightarrow>real" | |
| 1018 | by linarith | |
| 1019 | let ?h = "\<lambda>k. (2*k+1)/2 ^ Suc n * | |
| 1020 |                       (indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2*k+1) + 1)/2 ^ Suc n} x)"
 | |
| 1021 | show ?thesis | |
| 1022 | proof (rule *) | |
| 1023 | show "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). | |
| 1024 |                   2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1 + 1)/2 ^ Suc n} x)
 | |
| 1025 |                 \<le> sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1026 | by (rule sum_mono) (simp add: indicator_def field_split_simps) | 
| 67998 | 1027 | next | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1028 |           have \<alpha>: "?a = (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
 | 
| 67998 | 1029 |                          k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
 | 
| 1030 | by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) | |
| 1031 |           have \<beta>: "sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
 | |
| 1032 |                    = (\<Sum>k \<in> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
 | |
| 1033 |                       k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
 | |
| 1034 | by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1035 |           have 0: "(*) 2 ` {k \<in> \<int>. P k} \<inter> (\<lambda>x. 2 * x + 1) ` {k \<in> \<int>. P k} = {}" for P :: "real \<Rightarrow> bool"
 | 
| 67998 | 1036 | proof - | 
| 1037 | have "2 * i \<noteq> 2 * j + 1" for i j :: int by arith | |
| 1038 | thus ?thesis | |
| 1039 | unfolding Ints_def by auto (use of_int_eq_iff in fastforce) | |
| 1040 | qed | |
| 1041 |           have "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1042 |                 = (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
 | 
| 67998 | 1043 |                   k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
 | 
| 1044 | unfolding \<alpha> \<beta> | |
| 1045 | using finite_abs_int_segment [of "2 ^ (2*n)"] | |
| 1046 | by (subst sum_Un) (auto simp: 0) | |
| 1047 | also have "\<dots> \<le> ?b" | |
| 1048 | proof (rule sum_mono2) | |
| 1049 |             show "finite {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
 | |
| 1050 | by (rule finite_abs_int_segment) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1051 |             show "(*) 2 ` {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2^(2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2^(2*n)} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
 | 
| 67998 | 1052 | apply auto | 
| 1053 | using one_le_power [of "2::real" "2*n"] by linarith | |
| 1054 | have *: "\<lbrakk>x \<in> (S \<union> T) - U; \<And>x. x \<in> S \<Longrightarrow> x \<in> U; \<And>x. x \<in> T \<Longrightarrow> x \<in> U\<rbrakk> \<Longrightarrow> P x" for S T U P | |
| 1055 | by blast | |
| 1056 | have "0 \<le> b" if "b \<in> \<int>" "f x * (2 * 2^n) < b + 1" for b | |
| 1057 | proof - | |
| 1058 | have "0 \<le> f x * (2 * 2^n)" | |
| 1059 | by (simp add: f) | |
| 1060 | also have "\<dots> < b+1" | |
| 1061 | by (simp add: that) | |
| 1062 | finally show "0 \<le> b" | |
| 1063 | using \<open>b \<in> \<int>\<close> by (auto simp: elim!: Ints_cases) | |
| 1064 | qed | |
| 1065 |             then show "0 \<le> b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \<le> f y \<and> f y < (b + 1)/2 ^ Suc n} x"
 | |
| 1066 |                   if "b \<in> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)} -
 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1067 |                           ((*) 2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})" for b
 | 
| 67998 | 1068 | using that by (simp add: indicator_def divide_simps) | 
| 1069 | qed | |
| 1070 |           finally show "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<le> ?b" .
 | |
| 1071 | qed | |
| 1072 | qed | |
| 1073 | finally show ?thesis . | |
| 1074 | qed | |
| 1075 | show "?g n \<in> borel_measurable lebesgue" for n | |
| 1076 | apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) | |
| 1077 | using leb_f sets_restrict_UNIV by auto | |
| 1078 | show "finite (range (?g n))" for n | |
| 1079 | proof - | |
| 1080 | have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) | |
| 1081 |               \<in> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}" for x
 | |
| 1082 | proof (cases "\<exists>k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n) \<and> k/2^n \<le> f x \<and> f x < (k+1)/2^n") | |
| 1083 | case True | |
| 1084 | then show ?thesis | |
| 1085 | by (blast intro: indicator_sum_eq) | |
| 1086 | next | |
| 1087 | case False | |
| 1088 | then have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) = 0" | |
| 1089 | by auto | |
| 1090 | then show ?thesis by force | |
| 1091 | qed | |
| 1092 |       then have "range (?g n) \<subseteq> ((\<lambda>k. (k/2^n)) ` {k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n)})"
 | |
| 1093 | by auto | |
| 1094 |       moreover have "finite ((\<lambda>k::real. (k/2^n)) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})"
 | |
| 1095 | by (intro finite_imageI finite_abs_int_segment) | |
| 1096 | ultimately show ?thesis | |
| 1097 | by (rule finite_subset) | |
| 1098 | qed | |
| 1099 | show "(\<lambda>n. ?g n x) \<longlonglongrightarrow> f x" for x | |
| 1100 | proof (clarsimp simp add: lim_sequentially) | |
| 1101 | fix e::real | |
| 1102 | assume "e > 0" | |
| 1103 | obtain N1 where N1: "2 ^ N1 > abs(f x)" | |
| 1104 | using real_arch_pow by fastforce | |
| 1105 | obtain N2 where N2: "(1/2) ^ N2 < e" | |
| 1106 | using real_arch_pow_inv \<open>e > 0\<close> by fastforce | |
| 1107 | have "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x) < e" if "N1 + N2 \<le> n" for n | |
| 1108 | proof - | |
| 1109 | let ?m = "real_of_int \<lfloor>2^n * f x\<rfloor>" | |
| 1110 | have "\<bar>?m\<bar> \<le> 2^n * 2^N1" | |
| 1111 | using N1 apply (simp add: f) | |
| 1112 | by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) | |
| 1113 | also have "\<dots> \<le> 2 ^ (2*n)" | |
| 1114 | by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff | |
| 1115 | power_add power_increasing_iff semiring_norm(76)) | |
| 1116 | finally have m_le: "\<bar>?m\<bar> \<le> 2 ^ (2*n)" . | |
| 1117 | have "?m/2^n \<le> f x" "f x < (?m + 1)/2^n" | |
| 1118 | by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) | |
| 1119 | then have eq: "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x) | |
| 1120 | = dist (?m/2^n) (f x)" | |
| 1121 | by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) | |
| 1122 | have "\<bar>2^n\<bar> * \<bar>?m/2^n - f x\<bar> = \<bar>2^n * (?m/2^n - f x)\<bar>" | |
| 1123 | by (simp add: abs_mult) | |
| 1124 | also have "\<dots> < 2 ^ N2 * e" | |
| 1125 | using N2 by (simp add: divide_simps mult.commute) linarith | |
| 1126 | also have "\<dots> \<le> \<bar>2^n\<bar> * e" | |
| 1127 | using that \<open>e > 0\<close> by auto | |
| 1128 | finally have "dist (?m/2^n) (f x) < e" | |
| 1129 | by (simp add: dist_norm) | |
| 1130 | then show ?thesis | |
| 1131 | using eq by linarith | |
| 1132 | qed | |
| 1133 | then show "\<exists>no. \<forall>n\<ge>no. dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k * ?\<Omega> n k x/2^n) (f x) < e" | |
| 1134 | by force | |
| 1135 | qed | |
| 1136 | qed | |
| 1137 | ultimately show ?rhs | |
| 1138 | by metis | |
| 1139 | next | |
| 1140 | assume RHS: ?rhs | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 1141 | with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] | 
| 67998 | 1142 | show ?lhs | 
| 1143 | by (blast intro: order_trans) | |
| 1144 | qed | |
| 1145 | ||
| 69683 | 1146 | subsection\<open>Borel measurable Jacobian determinant\<close> | 
| 67998 | 1147 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1148 | lemma lemma_partial_derivatives0: | 
| 67998 | 1149 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1150 | assumes "linear f" and lim0: "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within S)" | |
| 1151 |     and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>)"
 | |
| 1152 | shows "f x = 0" | |
| 1153 | proof - | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1154 | interpret linear f by fact | 
| 67998 | 1155 |   have "dim {x. f x = 0} \<le> DIM('a)"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1156 | by (rule dim_subset_UNIV) | 
| 67998 | 1157 |   moreover have False if less: "dim {x. f x = 0} < DIM('a)"
 | 
| 1158 | proof - | |
| 1159 | obtain d where "d \<noteq> 0" and d: "\<And>y. f y = 0 \<Longrightarrow> d \<bullet> y = 0" | |
| 1160 | using orthogonal_to_subspace_exists [OF less] orthogonal_def | |
| 68074 | 1161 | by (metis (mono_tags, lifting) mem_Collect_eq span_base) | 
| 67998 | 1162 | then obtain k where "k > 0" | 
| 1163 |       and k: "\<And>e. e > 0 \<Longrightarrow> \<exists>y. y \<in> S - {0} \<and> norm y < e \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>"
 | |
| 1164 | using lb by blast | |
| 1165 | have "\<exists>h. \<forall>n. ((h n \<in> S \<and> h n \<noteq> 0 \<and> k * norm (h n) \<le> \<bar>d \<bullet> h n\<bar>) \<and> norm (h n) < 1 / real (Suc n)) \<and> | |
| 1166 | norm (h (Suc n)) < norm (h n)" | |
| 1167 | proof (rule dependent_nat_choice) | |
| 1168 | show "\<exists>y. (y \<in> S \<and> y \<noteq> 0 \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>) \<and> norm y < 1 / real (Suc 0)" | |
| 1169 | by simp (metis DiffE insertCI k not_less not_one_le_zero) | |
| 1170 | qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) | |
| 1171 |     then obtain \<alpha> where \<alpha>: "\<And>n. \<alpha> n \<in> S - {0}" and kd: "\<And>n. k * norm(\<alpha> n) \<le> \<bar>d \<bullet> \<alpha> n\<bar>"
 | |
| 1172 | and norm_lt: "\<And>n. norm(\<alpha> n) < 1/(Suc n)" | |
| 1173 | by force | |
| 1174 | let ?\<beta> = "\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)" | |
| 1175 | have com: "\<And>g. (\<forall>n. g n \<in> sphere (0::'a) 1) | |
| 1176 | \<Longrightarrow> \<exists>l \<in> sphere 0 1. \<exists>\<rho>::nat\<Rightarrow>nat. strict_mono \<rho> \<and> (g \<circ> \<rho>) \<longlonglongrightarrow> l" | |
| 1177 | using compact_sphere compact_def by metis | |
| 1178 | moreover have "\<forall>n. ?\<beta> n \<in> sphere 0 1" | |
| 1179 | using \<alpha> by auto | |
| 1180 | ultimately obtain l::'a and \<rho>::"nat\<Rightarrow>nat" | |
| 1181 | where l: "l \<in> sphere 0 1" and "strict_mono \<rho>" and to_l: "(?\<beta> \<circ> \<rho>) \<longlonglongrightarrow> l" | |
| 1182 | by meson | |
| 1183 | moreover have "continuous (at l) (\<lambda>x. (\<bar>d \<bullet> x\<bar> - k))" | |
| 1184 | by (intro continuous_intros) | |
| 1185 | ultimately have lim_dl: "((\<lambda>x. (\<bar>d \<bullet> x\<bar> - k)) \<circ> (?\<beta> \<circ> \<rho>)) \<longlonglongrightarrow> (\<bar>d \<bullet> l\<bar> - k)" | |
| 1186 | by (meson continuous_imp_tendsto) | |
| 1187 | have "\<forall>\<^sub>F i in sequentially. 0 \<le> ((\<lambda>x. \<bar>d \<bullet> x\<bar> - k) \<circ> ((\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>)) i" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1188 | using \<alpha> kd by (auto simp: field_split_simps) | 
| 67998 | 1189 | then have "k \<le> \<bar>d \<bullet> l\<bar>" | 
| 1190 | using tendsto_lowerbound [OF lim_dl, of 0] by auto | |
| 1191 | moreover have "d \<bullet> l = 0" | |
| 1192 | proof (rule d) | |
| 1193 | show "f l = 0" | |
| 1194 | proof (rule LIMSEQ_unique [of "f \<circ> ?\<beta> \<circ> \<rho>"]) | |
| 1195 | have "isCont f l" | |
| 1196 | using \<open>linear f\<close> linear_continuous_at linear_conv_bounded_linear by blast | |
| 1197 | then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> f l" | |
| 1198 | unfolding comp_assoc | |
| 1199 | using to_l continuous_imp_tendsto by blast | |
| 1200 | have "\<alpha> \<longlonglongrightarrow> 0" | |
| 1201 | using norm_lt LIMSEQ_norm_0 by metis | |
| 1202 | with \<open>strict_mono \<rho>\<close> have "(\<alpha> \<circ> \<rho>) \<longlonglongrightarrow> 0" | |
| 1203 | by (metis LIMSEQ_subseq_LIMSEQ) | |
| 1204 | with lim0 \<alpha> have "((\<lambda>x. f x /\<^sub>R norm x) \<circ> (\<alpha> \<circ> \<rho>)) \<longlonglongrightarrow> 0" | |
| 1205 | by (force simp: tendsto_at_iff_sequentially) | |
| 1206 | then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> 0" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1207 | by (simp add: o_def scale) | 
| 67998 | 1208 | qed | 
| 1209 | qed | |
| 1210 | ultimately show False | |
| 1211 | using \<open>k > 0\<close> by auto | |
| 1212 | qed | |
| 1213 |   ultimately have dim: "dim {x. f x = 0} = DIM('a)"
 | |
| 1214 | by force | |
| 1215 | then show ?thesis | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1216 | using dim_eq_full | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1217 | by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1218 | mem_Collect_eq module_hom_zero span_base span_raw_def) | 
| 67998 | 1219 | qed | 
| 1220 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1221 | lemma lemma_partial_derivatives: | 
| 67998 | 1222 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1223 | assumes "linear f" and lim: "((\<lambda>x. f (x - a) /\<^sub>R norm (x - a)) \<longlongrightarrow> 0) (at a within S)" | |
| 1224 |     and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0.  \<forall>e>0. \<exists>x \<in> S - {a}. norm(a - x) < e \<and> k * norm(a - x) \<le> \<bar>v \<bullet> (x - a)\<bar>)"
 | |
| 1225 | shows "f x = 0" | |
| 1226 | proof - | |
| 1227 | have "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within (\<lambda>x. x-a) ` S)" | |
| 1228 | using lim by (simp add: Lim_within dist_norm) | |
| 1229 | then show ?thesis | |
| 1230 | proof (rule lemma_partial_derivatives0 [OF \<open>linear f\<close>]) | |
| 1231 | fix v :: "'a" | |
| 1232 | assume v: "v \<noteq> 0" | |
| 1233 |     show "\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> (\<lambda>x. x - a) ` S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>"
 | |
| 1234 | using lb [OF v] by (force simp: norm_minus_commute) | |
| 1235 | qed | |
| 1236 | qed | |
| 1237 | ||
| 1238 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1239 | proposition borel_measurable_partial_derivatives: | 
| 67998 | 1240 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n"
 | 
| 1241 | assumes S: "S \<in> sets lebesgue" | |
| 1242 | and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 1243 | shows "(\<lambda>x. (matrix(f' x)$m$n)) \<in> borel_measurable (lebesgue_on S)" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1244 | proof - | 
| 67998 | 1245 | have contf: "continuous_on S f" | 
| 1246 | using continuous_on_eq_continuous_within f has_derivative_continuous by blast | |
| 1247 |   have "{x \<in> S.  (matrix (f' x)$m$n) \<le> b} \<in> sets lebesgue" for b
 | |
| 1248 | proof (rule sets_negligible_symdiff) | |
| 1249 |     let ?T = "{x \<in> S. \<forall>e>0. \<exists>d>0. \<exists>A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>) \<and>
 | |
| 1250 | (\<forall>y \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x))}" | |
| 1251 | let ?U = "S \<inter> | |
| 1252 |               (\<Inter>e \<in> {e \<in> \<rat>. e > 0}.
 | |
| 1253 |                 \<Union>A \<in> {A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>)}.
 | |
| 1254 |                   \<Union>d \<in> {d \<in> \<rat>. 0 < d}.
 | |
| 1255 |                      S \<inter> (\<Inter>y \<in> S. {x \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x)}))"
 | |
| 1256 | have "?T = ?U" | |
| 1257 | proof (intro set_eqI iffI) | |
| 1258 | fix x | |
| 1259 | assume xT: "x \<in> ?T" | |
| 1260 | then show "x \<in> ?U" | |
| 1261 | proof (clarsimp simp add:) | |
| 1262 | fix q :: real | |
| 1263 | assume "q \<in> \<rat>" "q > 0" | |
| 1264 | then obtain d A where "d > 0" and A: "A $ m $ n < b" "\<And>i j. A $ i $ j \<in> \<rat>" | |
| 1265 | "\<And>y. \<lbrakk>y\<in>S; norm (y - x) < d\<rbrakk> \<Longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)" | |
| 1266 | using xT by auto | |
| 1267 | then obtain \<delta> where "d > \<delta>" "\<delta> > 0" "\<delta> \<in> \<rat>" | |
| 1268 | using Rats_dense_in_real by blast | |
| 1269 | with A show "\<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> | |
| 1270 | (\<exists>s. s \<in> \<rat> \<and> 0 < s \<and> (\<forall>y\<in>S. norm (y - x) < s \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)))" | |
| 1271 | by force | |
| 1272 | qed | |
| 1273 | next | |
| 1274 | fix x | |
| 1275 | assume xU: "x \<in> ?U" | |
| 1276 | then show "x \<in> ?T" | |
| 1277 | proof clarsimp | |
| 1278 | fix e :: "real" | |
| 1279 | assume "e > 0" | |
| 1280 | then obtain \<epsilon> where \<epsilon>: "e > \<epsilon>" "\<epsilon> > 0" "\<epsilon> \<in> \<rat>" | |
| 1281 | using Rats_dense_in_real by blast | |
| 1282 | with xU obtain A r where "x \<in> S" and Ar: "A $ m $ n < b" "\<forall>i j. A $ i $ j \<in> \<rat>" "r \<in> \<rat>" "r > 0" | |
| 1283 | and "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> \<epsilon> * norm (y - x)" | |
| 1284 | by (auto simp: split: if_split_asm) | |
| 1285 | then have "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)" | |
| 1286 | by (meson \<open>e > \<epsilon>\<close> less_eq_real_def mult_right_mono norm_ge_zero order_trans) | |
| 1287 | then show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))" | |
| 1288 | using \<open>x \<in> S\<close> Ar by blast | |
| 1289 | qed | |
| 1290 | qed | |
| 1291 | moreover have "?U \<in> sets lebesgue" | |
| 1292 | proof - | |
| 1293 |       have coQ: "countable {e \<in> \<rat>. 0 < e}"
 | |
| 1294 | using countable_Collect countable_rat by blast | |
| 1295 |       have ne: "{e \<in> \<rat>. (0::real) < e} \<noteq> {}"
 | |
| 1296 | using zero_less_one Rats_1 by blast | |
| 1297 |       have coA: "countable {A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>)}"
 | |
| 1298 | proof (rule countable_subset) | |
| 1299 |         show "countable {A. \<forall>i j. A $ i $ j \<in> \<rat>}"
 | |
| 1300 | using countable_vector [OF countable_vector, of "\<lambda>i j. \<rat>"] by (simp add: countable_rat) | |
| 1301 | qed blast | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1302 |       have *: "\<lbrakk>U \<noteq> {} \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)\<rbrakk>
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1303 | \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)" for U | 
| 67998 | 1304 | by fastforce | 
| 1305 |       have eq: "{x::(real,'m)vec. P x \<and> (Q x \<longrightarrow> R x)} = {x. P x \<and> \<not> Q x} \<union> {x. P x \<and> R x}" for P Q R
 | |
| 1306 | by auto | |
| 1307 |       have sets: "S \<inter> (\<Inter>y\<in>S. {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)})
 | |
| 1308 | \<in> sets lebesgue" for e A d | |
| 1309 | proof - | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1310 | have clo: "closedin (top_of_set S) | 
| 67998 | 1311 |                      {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)}"
 | 
| 1312 | for y | |
| 1313 | proof - | |
| 1314 | have cont1: "continuous_on S (\<lambda>x. norm (y - x))" | |
| 1315 | and cont2: "continuous_on S (\<lambda>x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" | |
| 1316 | by (force intro: contf continuous_intros)+ | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1317 |           have clo1: "closedin (top_of_set S) {x \<in> S. d \<le> norm (y - x)}"
 | 
| 67998 | 1318 |             using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def)
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1319 | have clo2: "closedin (top_of_set S) | 
| 67998 | 1320 |                        {x \<in> S. norm (f y - f x - (A *v y - A *v x)) \<le> e * norm (y - x)}"
 | 
| 1321 |             using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def)
 | |
| 1322 | show ?thesis | |
| 1323 | by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) | |
| 1324 | qed | |
| 1325 | show ?thesis | |
| 1326 | by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ | |
| 1327 | qed | |
| 1328 | show ?thesis | |
| 1329 | by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto | |
| 1330 | qed | |
| 1331 | ultimately show "?T \<in> sets lebesgue" | |
| 1332 | by simp | |
| 1333 |     let ?M = "(?T - {x \<in> S. matrix (f' x) $ m $ n \<le> b} \<union> ({x \<in> S. matrix (f' x) $ m $ n \<le> b} - ?T))"
 | |
| 1334 |     let ?\<Theta> = "\<lambda>x v. \<forall>\<xi>>0. \<exists>e>0. \<forall>y \<in> S-{x}. norm (x - y) < e \<longrightarrow> \<bar>v \<bullet> (y - x)\<bar> < \<xi> * norm (x - y)"
 | |
| 1335 |     have nN: "negligible {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
 | |
| 1336 | unfolding negligible_eq_zero_density | |
| 1337 | proof clarsimp | |
| 1338 | fix x v and r e :: "real" | |
| 1339 | assume "x \<in> S" "v \<noteq> 0" "r > 0" "e > 0" | |
| 1340 | and Theta [rule_format]: "?\<Theta> x v" | |
| 1341 |       moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0"
 | |
| 1342 | by (simp add: \<open>v \<noteq> 0\<close> \<open>e > 0\<close>) | |
| 1343 | ultimately obtain d where "d > 0" | |
| 1344 |          and dless: "\<And>y. \<lbrakk>y \<in> S - {x}; norm (x - y) < d\<rbrakk> \<Longrightarrow>
 | |
| 1345 |                         \<bar>v \<bullet> (y - x)\<bar> < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)"
 | |
| 1346 | by metis | |
| 1347 |       let ?W = "ball x (min d r) \<inter> {y. \<bar>v \<bullet> (y - x)\<bar> < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}"
 | |
| 1348 |       have "open {x. \<bar>v \<bullet> (x - a)\<bar> < b}" for a b
 | |
| 1349 | by (intro open_Collect_less continuous_intros) | |
| 1350 | show "\<exists>d>0. d \<le> r \<and> | |
| 1351 |             (\<exists>U. {x' \<in> S. \<exists>v\<noteq>0. ?\<Theta> x' v} \<inter> ball x d \<subseteq> U \<and>
 | |
| 1352 | U \<in> lmeasurable \<and> measure lebesgue U < e * content (ball x d))" | |
| 1353 | proof (intro exI conjI) | |
| 1354 | show "0 < min d r" "min d r \<le> r" | |
| 1355 | using \<open>r > 0\<close> \<open>d > 0\<close> by auto | |
| 1356 |         show "{x' \<in> S. \<exists>v. v \<noteq> 0 \<and> (\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {x'}. norm (x' - z) < e \<longrightarrow> \<bar>v \<bullet> (z - x')\<bar> < \<xi> * norm (x' - z))} \<inter> ball x (min d r) \<subseteq> ?W"
 | |
| 1357 | proof (clarsimp simp: dist_norm norm_minus_commute) | |
| 68001 
0a2a1b6507c1
correction of TeX errors and other oversights
 paulson <lp15@cam.ac.uk> parents: 
67999diff
changeset | 1358 | fix y w | 
| 67998 | 1359 | assume "y \<in> S" "w \<noteq> 0" | 
| 1360 | and less [rule_format]: | |
| 1361 |                     "\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {y}. norm (y - z) < e \<longrightarrow> \<bar>w \<bullet> (z - y)\<bar> < \<xi> * norm (y - z)"
 | |
| 1362 | and d: "norm (y - x) < d" and r: "norm (y - x) < r" | |
| 1363 |             show "\<bar>v \<bullet> (y - x)\<bar> < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
 | |
| 1364 | proof (cases "y = x") | |
| 1365 | case True | |
| 1366 | with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> \<open>v \<noteq> 0\<close> show ?thesis | |
| 1367 | by simp | |
| 1368 | next | |
| 1369 | case False | |
| 1370 |               have "\<bar>v \<bullet> (y - x)\<bar> < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)"
 | |
| 1371 | apply (rule dless) | |
| 1372 | using False \<open>y \<in> S\<close> d by (auto simp: norm_minus_commute) | |
| 1373 |               also have "\<dots> \<le> norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
 | |
| 1374 | using d r \<open>e > 0\<close> by (simp add: field_simps norm_minus_commute mult_left_mono) | |
| 1375 | finally show ?thesis . | |
| 1376 | qed | |
| 1377 | qed | |
| 1378 | show "?W \<in> lmeasurable" | |
| 1379 | by (simp add: fmeasurable_Int_fmeasurable borel_open) | |
| 1380 | obtain k::'m where True | |
| 1381 | by metis | |
| 1382 | obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" | |
| 1383 | using rotation_rightward_line by metis | |
| 1384 | define b where "b \<equiv> norm v" | |
| 1385 | have "b > 0" | |
| 1386 | using \<open>v \<noteq> 0\<close> by (auto simp: b_def) | |
| 1387 | obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" | |
| 1388 | by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) | |
| 1389 |           let ?v = "\<chi> i. min d r / CARD('m)"
 | |
| 1390 |           let ?v' = "\<chi> i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r"
 | |
| 1391 | let ?x' = "inv T x" | |
| 1392 |           let ?W' = "(ball ?x' (min d r) \<inter> {y. \<bar>(y - ?x')$k\<bar> < e * min d r / (2 * CARD('m) ^ CARD('m))})"
 | |
| 1393 | have abs: "x - e \<le> y \<and> y \<le> x + e \<longleftrightarrow> abs(y - x) \<le> e" for x y e::real | |
| 1394 | by auto | |
| 1395 | have "?W = T ` ?W'" | |
| 1396 | proof - | |
| 1397 | have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" | |
| 1398 | by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) | |
| 1399 |             have 2: "{y. \<bar>v \<bullet> (y - x)\<bar> < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} =
 | |
| 1400 |                       T ` {y. \<bar>y $ k - ?x' $ k\<bar> < e * min d r / (2 * real CARD('m) ^ CARD('m))}"
 | |
| 1401 | proof - | |
| 1402 | have *: "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = b * \<bar>inv T y $ k - ?x' $ k\<bar>" for y | |
| 1403 | proof - | |
| 1404 | have "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = \<bar>(b *\<^sub>R axis k 1) \<bullet> inv T (y - x)\<bar>" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 1405 | by (metis (no_types, opaque_lifting) b_def eqb invT orthogonal_transformation_def v) | 
| 67998 | 1406 | also have "\<dots> = b * \<bar>(axis k 1) \<bullet> inv T (y - x)\<bar>" | 
| 1407 | using \<open>b > 0\<close> by (simp add: abs_mult) | |
| 1408 | also have "\<dots> = b * \<bar>inv T y $ k - ?x' $ k\<bar>" | |
| 1409 | using orthogonal_transformation_linear [OF invT] | |
| 1410 | by (simp add: inner_axis' linear_diff) | |
| 1411 | finally show ?thesis | |
| 1412 | by simp | |
| 1413 | qed | |
| 1414 | show ?thesis | |
| 1415 | using v b_def [symmetric] | |
| 1416 | using \<open>b > 0\<close> by (simp add: * bij_image_Collect_eq [OF \<open>bij T\<close>] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) | |
| 1417 | qed | |
| 1418 | show ?thesis | |
| 1419 | using \<open>b > 0\<close> by (simp add: image_Int \<open>inj T\<close> 1 2 b_def [symmetric]) | |
| 1420 | qed | |
| 1421 | moreover have "?W' \<in> lmeasurable" | |
| 1422 | by (auto intro: fmeasurable_Int_fmeasurable) | |
| 1423 | ultimately have "measure lebesgue ?W = measure lebesgue ?W'" | |
| 1424 | by (metis measure_orthogonal_image T) | |
| 1425 | also have "\<dots> \<le> measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" | |
| 1426 | proof (rule measure_mono_fmeasurable) | |
| 1427 | show "?W' \<subseteq> cbox (?x' - ?v') (?x' + ?v')" | |
| 1428 | apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) | |
| 1429 | by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) | |
| 1430 | qed auto | |
| 1431 | also have "\<dots> \<le> e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" | |
| 1432 | proof - | |
| 1433 |             have "cbox (?x' - ?v) (?x' + ?v) \<noteq> {}"
 | |
| 1434 | using \<open>r > 0\<close> \<open>d > 0\<close> by (auto simp: interval_eq_empty_cart divide_less_0_iff) | |
| 1435 | with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> show ?thesis | |
| 1436 | apply (simp add: content_cbox_if_cart mem_box_cart) | |
| 1437 | apply (auto simp: prod_nonneg) | |
| 71172 | 1438 | apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) | 
| 67998 | 1439 | done | 
| 1440 | qed | |
| 1441 | also have "\<dots> \<le> e/2 * measure lebesgue (cball ?x' (min d r))" | |
| 1442 | proof (rule mult_left_mono [OF measure_mono_fmeasurable]) | |
| 1443 | have *: "norm (?x' - y) \<le> min d r" | |
| 1444 |               if y: "\<And>i. \<bar>?x' $ i - y $ i\<bar> \<le> min d r / real CARD('m)" for y
 | |
| 1445 | proof - | |
| 1446 | have "norm (?x' - y) \<le> (\<Sum>i\<in>UNIV. \<bar>(?x' - y) $ i\<bar>)" | |
| 1447 | by (rule norm_le_l1_cart) | |
| 1448 |               also have "\<dots> \<le> real CARD('m) * (min d r / real CARD('m))"
 | |
| 1449 | by (rule sum_bounded_above) (use y in auto) | |
| 1450 | finally show ?thesis | |
| 1451 | by simp | |
| 1452 | qed | |
| 1453 | show "cbox (?x' - ?v) (?x' + ?v) \<subseteq> cball ?x' (min d r)" | |
| 1454 | apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) | |
| 1455 | by (simp add: abs_diff_le_iff abs_minus_commute) | |
| 1456 | qed (use \<open>e > 0\<close> in auto) | |
| 1457 | also have "\<dots> < e * content (cball ?x' (min d r))" | |
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 1458 | using \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> by (auto intro: content_cball_pos) | 
| 67998 | 1459 | also have "\<dots> = e * content (ball x (min d r))" | 
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 1460 | using \<open>r > 0\<close> \<open>d > 0\<close> content_ball_conv_unit_ball[of "min d r" "inv T x"] | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 1461 | content_ball_conv_unit_ball[of "min d r" x] | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71189diff
changeset | 1462 | by (simp add: content_cball_conv_ball) | 
| 67998 | 1463 | finally show "measure lebesgue ?W < e * content (ball x (min d r))" . | 
| 1464 | qed | |
| 1465 | qed | |
| 1466 | have *: "(\<And>x. (x \<notin> S) \<Longrightarrow> (x \<in> T \<longleftrightarrow> x \<in> U)) \<Longrightarrow> (T - U) \<union> (U - T) \<subseteq> S" for S T U :: "(real,'m) vec set" | |
| 1467 | by blast | |
| 1468 |     have MN: "?M \<subseteq> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
 | |
| 1469 | proof (rule *) | |
| 1470 | fix x | |
| 1471 |       assume x: "x \<notin> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
 | |
| 1472 |       show "(x \<in> ?T) \<longleftrightarrow> (x \<in> {x \<in> S. matrix (f' x) $ m $ n \<le> b})"
 | |
| 1473 | proof (cases "x \<in> S") | |
| 1474 | case True | |
| 1475 | then have x: "\<not> ?\<Theta> x v" if "v \<noteq> 0" for v | |
| 1476 | using x that by force | |
| 1477 | show ?thesis | |
| 1478 | proof (rule iffI; clarsimp) | |
| 1479 | assume b: "\<forall>e>0. \<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> | |
| 1480 | (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))" | |
| 1481 | (is "\<forall>e>0. \<exists>d>0. \<exists>A. ?\<Phi> e d A") | |
| 1482 | then have "\<forall>k. \<exists>d>0. \<exists>A. ?\<Phi> (1 / Suc k) d A" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 1483 | by (metis (no_types, opaque_lifting) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) | 
| 67998 | 1484 | then obtain \<delta> A where \<delta>: "\<And>k. \<delta> k > 0" | 
| 1485 | and Ab: "\<And>k. A k $ m $ n < b" | |
| 1486 | and A: "\<And>k y. \<lbrakk>y \<in> S; norm (y - x) < \<delta> k\<rbrakk> \<Longrightarrow> | |
| 1487 | norm (f y - f x - A k *v (y - x)) \<le> 1/(Suc k) * norm (y - x)" | |
| 1488 | by metis | |
| 1489 | have "\<forall>i j. \<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a" | |
| 1490 | proof (intro allI) | |
| 1491 | fix i j | |
| 1492 | have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n | |
| 1493 | by (metis cart_eq_inner_axis matrix_vector_mul_component) | |
| 1494 |             let ?CA = "{x. Cauchy (\<lambda>n. (A n) *v x)}"
 | |
| 1495 | have "subspace ?CA" | |
| 1496 | unfolding subspace_def convergent_eq_Cauchy [symmetric] | |
| 1497 | by (force simp: algebra_simps intro: tendsto_intros) | |
| 1498 | then have CA_eq: "?CA = span ?CA" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1499 | by (metis span_eq_iff) | 
| 67998 | 1500 | also have "\<dots> = UNIV" | 
| 1501 | proof - | |
| 1502 |               have "dim ?CA \<le> CARD('m)"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1503 | using dim_subset_UNIV[of ?CA] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1504 | by auto | 
| 67998 | 1505 |               moreover have "False" if less: "dim ?CA < CARD('m)"
 | 
| 1506 | proof - | |
| 1507 | obtain d where "d \<noteq> 0" and d: "\<And>y. y \<in> span ?CA \<Longrightarrow> orthogonal d y" | |
| 1508 | using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) | |
| 1509 | with x [OF \<open>d \<noteq> 0\<close>] obtain \<xi> where "\<xi> > 0" | |
| 1510 |                   and \<xi>: "\<And>e. e > 0 \<Longrightarrow> \<exists>y \<in> S - {x}. norm (x - y) < e \<and> \<xi> * norm (x - y) \<le> \<bar>d \<bullet> (y - x)\<bar>"
 | |
| 1511 | by (fastforce simp: not_le Bex_def) | |
| 1512 |                 obtain \<gamma> z where \<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
 | |
| 1513 | and \<gamma>le: "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>" | |
| 1514 | and \<gamma>x: "\<gamma> \<longlonglongrightarrow> x" | |
| 1515 | and z: "(\<lambda>n. (\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x)) \<longlonglongrightarrow> z" | |
| 1516 | proof - | |
| 1517 |                   have "\<exists>\<gamma>. (\<forall>i. (\<gamma> i \<in> S - {x} \<and>
 | |
| 1518 | \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar> \<and> norm(\<gamma> i - x) < 1/Suc i) \<and> | |
| 1519 | norm(\<gamma>(Suc i) - x) < norm(\<gamma> i - x))" | |
| 1520 | proof (rule dependent_nat_choice) | |
| 1521 |                     show "\<exists>y. y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1 / Suc 0"
 | |
| 1522 | using \<xi> [of 1] by (auto simp: dist_norm norm_minus_commute) | |
| 1523 | next | |
| 1524 | fix y i | |
| 1525 |                     assume "y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1/Suc i"
 | |
| 1526 | then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" | |
| 1527 | by auto | |
| 1528 |                     then obtain y' where "y' \<in> S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))"
 | |
| 1529 | "\<xi> * norm (x - y') \<le> \<bar>d \<bullet> (y' - x)\<bar>" | |
| 1530 | using \<xi> by metis | |
| 1531 |                     with \<xi> show "\<exists>y'. (y' \<in> S - {x} \<and> \<xi> * norm (y' - x) \<le> \<bar>d \<bullet> (y' - x)\<bar> \<and>
 | |
| 1532 | norm (y' - x) < 1/(Suc (Suc i))) \<and> norm (y' - x) < norm (y - x)" | |
| 1533 | by (auto simp: dist_norm norm_minus_commute) | |
| 1534 | qed | |
| 1535 | then obtain \<gamma> where | |
| 1536 |                         \<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
 | |
| 1537 | and \<gamma>le: "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>" | |
| 1538 | and \<gamma>conv: "\<And>i. norm(\<gamma> i - x) < 1/(Suc i)" | |
| 1539 | by blast | |
| 1540 | let ?f = "\<lambda>i. (\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x)" | |
| 1541 | have "?f i \<in> sphere 0 1" for i | |
| 1542 | using \<gamma>Sx by auto | |
| 1543 | then obtain l \<rho> where "l \<in> sphere 0 1" "strict_mono \<rho>" and l: "(?f \<circ> \<rho>) \<longlonglongrightarrow> l" | |
| 1544 | using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson | |
| 1545 | show thesis | |
| 1546 | proof | |
| 1547 |                     show "(\<gamma> \<circ> \<rho>) i \<in> S - {x}" "\<xi> * norm ((\<gamma> \<circ> \<rho>) i - x) \<le> \<bar>d \<bullet> ((\<gamma> \<circ> \<rho>) i - x)\<bar>" for i
 | |
| 1548 | using \<gamma>Sx \<gamma>le by auto | |
| 1549 | have "\<gamma> \<longlonglongrightarrow> x" | |
| 1550 | proof (clarsimp simp add: LIMSEQ_def dist_norm) | |
| 1551 | fix r :: "real" | |
| 1552 | assume "r > 0" | |
| 1553 | with real_arch_invD obtain no where "no \<noteq> 0" "real no > 1/r" | |
| 1554 | by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) | |
| 1555 | with \<gamma>conv show "\<exists>no. \<forall>n\<ge>no. norm (\<gamma> n - x) < r" | |
| 1556 | by (metis \<open>r > 0\<close> add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) | |
| 1557 | qed | |
| 1558 | with \<open>strict_mono \<rho>\<close> show "(\<gamma> \<circ> \<rho>) \<longlonglongrightarrow> x" | |
| 1559 | by (metis LIMSEQ_subseq_LIMSEQ) | |
| 1560 | show "(\<lambda>n. ((\<gamma> \<circ> \<rho>) n - x) /\<^sub>R norm ((\<gamma> \<circ> \<rho>) n - x)) \<longlonglongrightarrow> l" | |
| 1561 | using l by (auto simp: o_def) | |
| 1562 | qed | |
| 1563 | qed | |
| 1564 | have "isCont (\<lambda>x. (\<bar>d \<bullet> x\<bar> - \<xi>)) z" | |
| 1565 | by (intro continuous_intros) | |
| 1566 | from isCont_tendsto_compose [OF this z] | |
| 1567 | have lim: "(\<lambda>y. \<bar>d \<bullet> ((\<gamma> y - x) /\<^sub>R norm (\<gamma> y - x))\<bar> - \<xi>) \<longlonglongrightarrow> \<bar>d \<bullet> z\<bar> - \<xi>" | |
| 1568 | by auto | |
| 1569 | moreover have "\<forall>\<^sub>F i in sequentially. 0 \<le> \<bar>d \<bullet> ((\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x))\<bar> - \<xi>" | |
| 1570 | proof (rule eventuallyI) | |
| 1571 | fix n | |
| 1572 | show "0 \<le> \<bar>d \<bullet> ((\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x))\<bar> - \<xi>" | |
| 1573 | using \<gamma>le [of n] \<gamma>Sx by (auto simp: abs_mult divide_simps) | |
| 1574 | qed | |
| 1575 | ultimately have "\<xi> \<le> \<bar>d \<bullet> z\<bar>" | |
| 1576 | using tendsto_lowerbound [where a=0] by fastforce | |
| 1577 | have "Cauchy (\<lambda>n. (A n) *v z)" | |
| 1578 | proof (clarsimp simp add: Cauchy_def) | |
| 1579 | fix \<epsilon> :: "real" | |
| 1580 | assume "0 < \<epsilon>" | |
| 1581 | then obtain N::nat where "N > 0" and N: "\<epsilon>/2 > 1/N" | |
| 1582 | by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) | |
| 1583 | show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (A m *v z) (A n *v z) < \<epsilon>" | |
| 1584 | proof (intro exI allI impI) | |
| 1585 | fix i j | |
| 1586 | assume ij: "N \<le> i" "N \<le> j" | |
| 1587 | let ?V = "\<lambda>i k. A i *v ((\<gamma> k - x) /\<^sub>R norm (\<gamma> k - x))" | |
| 1588 | have "\<forall>\<^sub>F k in sequentially. dist (\<gamma> k) x < min (\<delta> i) (\<delta> j)" | |
| 1589 | using \<gamma>x [unfolded tendsto_iff] by (meson min_less_iff_conj \<delta>) | |
| 1590 | then have even: "\<forall>\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \<le> 0" | |
| 1591 | proof (rule eventually_mono, clarsimp) | |
| 1592 | fix p | |
| 1593 | assume p: "dist (\<gamma> p) x < \<delta> i" "dist (\<gamma> p) x < \<delta> j" | |
| 1594 | let ?C = "\<lambda>k. f (\<gamma> p) - f x - A k *v (\<gamma> p - x)" | |
| 1595 | have "norm ((A i - A j) *v (\<gamma> p - x)) = norm (?C j - ?C i)" | |
| 1596 | by (simp add: algebra_simps) | |
| 1597 | also have "\<dots> \<le> norm (?C j) + norm (?C i)" | |
| 1598 | using norm_triangle_ineq4 by blast | |
| 1599 | also have "\<dots> \<le> 1/(Suc j) * norm (\<gamma> p - x) + 1/(Suc i) * norm (\<gamma> p - x)" | |
| 1600 | by (metis A Diff_iff \<gamma>Sx dist_norm p add_mono) | |
| 1601 | also have "\<dots> \<le> 1/N * norm (\<gamma> p - x) + 1/N * norm (\<gamma> p - x)" | |
| 1602 | apply (intro add_mono mult_right_mono) | |
| 1603 | using ij \<open>N > 0\<close> by (auto simp: field_simps) | |
| 1604 | also have "\<dots> = 2 / N * norm (\<gamma> p - x)" | |
| 1605 | by simp | |
| 1606 | finally have no_le: "norm ((A i - A j) *v (\<gamma> p - x)) \<le> 2 / N * norm (\<gamma> p - x)" . | |
| 1607 | have "norm (?V i p - ?V j p) = | |
| 1608 | norm ((A i - A j) *v ((\<gamma> p - x) /\<^sub>R norm (\<gamma> p - x)))" | |
| 1609 | by (simp add: algebra_simps) | |
| 1610 | also have "\<dots> = norm ((A i - A j) *v (\<gamma> p - x)) / norm (\<gamma> p - x)" | |
| 1611 | by (simp add: divide_inverse matrix_vector_mult_scaleR) | |
| 1612 | also have "\<dots> \<le> 2 / N" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1613 | using no_le by (auto simp: field_split_simps) | 
| 67998 | 1614 | finally show "norm (?V i p - ?V j p) \<le> 2 / N" . | 
| 1615 | qed | |
| 1616 | have "isCont (\<lambda>w. (norm(A i *v w - A j *v w) - 2 / N)) z" | |
| 1617 | by (intro continuous_intros) | |
| 1618 | from isCont_tendsto_compose [OF this z] | |
| 1619 | have lim: "(\<lambda>w. norm (A i *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x)) - | |
| 1620 | A j *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x))) - 2 / N) | |
| 1621 | \<longlonglongrightarrow> norm (A i *v z - A j *v z) - 2 / N" | |
| 1622 | by auto | |
| 1623 | have "dist (A i *v z) (A j *v z) \<le> 2 / N" | |
| 1624 | using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) | |
| 1625 | with N show "dist (A i *v z) (A j *v z) < \<epsilon>" | |
| 1626 | by linarith | |
| 1627 | qed | |
| 1628 | qed | |
| 1629 | then have "d \<bullet> z = 0" | |
| 1630 | using CA_eq d orthogonal_def by auto | |
| 1631 | then show False | |
| 1632 | using \<open>0 < \<xi>\<close> \<open>\<xi> \<le> \<bar>d \<bullet> z\<bar>\<close> by auto | |
| 1633 | qed | |
| 1634 | ultimately show ?thesis | |
| 1635 | using dim_eq_full by fastforce | |
| 1636 | qed | |
| 1637 | finally have "?CA = UNIV" . | |
| 1638 | then have "Cauchy (\<lambda>n. (A n) *v axis j 1)" | |
| 1639 | by auto | |
| 1640 | then obtain L where "(\<lambda>n. A n *v axis j 1) \<longlonglongrightarrow> L" | |
| 1641 | by (auto simp: Cauchy_convergent_iff convergent_def) | |
| 1642 | then have "(\<lambda>x. (A x *v axis j 1) $ i) \<longlonglongrightarrow> L $ i" | |
| 1643 | by (rule tendsto_vec_nth) | |
| 1644 | then show "\<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a" | |
| 1645 | by (force simp: vax) | |
| 1646 | qed | |
| 1647 | then obtain B where B: "\<And>i j. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> B $ i $ j" | |
| 1648 | by (auto simp: lambda_skolem) | |
| 1649 | have lin_df: "linear (f' x)" | |
| 1650 | and lim_df: "((\<lambda>y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \<longlongrightarrow> 0) (at x within S)" | |
| 1651 | using \<open>x \<in> S\<close> assms by (auto simp: has_derivative_within linear_linear) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1652 | moreover | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1653 | interpret linear "f' x" by fact | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1654 | have "(matrix (f' x) - B) *v w = 0" for w | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1655 | proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1656 | show "linear ((*v) (matrix (f' x) - B))" | 
| 67998 | 1657 | by (rule matrix_vector_mul_linear) | 
| 1658 | have "((\<lambda>y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)" | |
| 71633 | 1659 | using tendsto_minus [OF lim_df] by (simp add: field_split_simps) | 
| 67998 | 1660 | then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)" | 
| 1661 | proof (rule Lim_transform) | |
| 1662 | have "((\<lambda>y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \<longlongrightarrow> 0) (at x within S)" | |
| 1663 | proof (clarsimp simp add: Lim_within dist_norm) | |
| 1664 | fix e :: "real" | |
| 1665 | assume "e > 0" | |
| 1666 | then obtain q::nat where "q \<noteq> 0" and qe2: "1/q < e/2" | |
| 1667 | by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) | |
| 1668 | let ?g = "\<lambda>p. sum (\<lambda>i. sum (\<lambda>j. abs((A p - B)$i$j)) UNIV) UNIV" | |
| 1669 | have "(\<lambda>k. onorm (\<lambda>y. (A k - B) *v y)) \<longlonglongrightarrow> 0" | |
| 1670 | proof (rule Lim_null_comparison) | |
| 1671 | show "\<forall>\<^sub>F k in sequentially. norm (onorm (\<lambda>y. (A k - B) *v y)) \<le> ?g k" | |
| 1672 | proof (rule eventually_sequentiallyI) | |
| 1673 | fix k :: "nat" | |
| 1674 | assume "0 \<le> k" | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1675 | have "0 \<le> onorm ((*v) (A k - B))" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1676 | using matrix_vector_mul_bounded_linear | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1677 | by (rule onorm_pos_le) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1678 | then show "norm (onorm ((*v) (A k - B))) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>(A k - B) $ i $ j\<bar>)" | 
| 67998 | 1679 | by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) | 
| 1680 | qed | |
| 1681 | next | |
| 1682 | show "?g \<longlonglongrightarrow> 0" | |
| 1683 | using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) | |
| 1684 | qed | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1685 | with \<open>e > 0\<close> obtain p where "\<And>n. n \<ge> p \<Longrightarrow> \<bar>onorm ((*v) (A n - B))\<bar> < e/2" | 
| 67998 | 1686 | unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1687 | then have pqe2: "\<bar>onorm ((*v) (A (p + q) - B))\<bar> < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) | 
| 67998 | 1688 | using le_add1 by blast | 
| 1689 | show "\<exists>d>0. \<forall>y\<in>S. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> | |
| 1690 | inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" | |
| 1691 | proof (intro exI, safe) | |
| 1692 | show "0 < \<delta>(p + q)" | |
| 1693 | by (simp add: \<delta>) | |
| 1694 | next | |
| 1695 | fix y | |
| 1696 | assume y: "y \<in> S" "norm (y - x) < \<delta>(p + q)" and "y \<noteq> x" | |
| 1697 | have *: "\<lbrakk>norm(b - c) < e - d; norm(y - x - b) \<le> d\<rbrakk> \<Longrightarrow> norm(y - x - c) < e" | |
| 1698 | for b c d e x and y:: "real^'n" | |
| 1699 | using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp | |
| 1700 | have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" | |
| 1701 | proof (rule *) | |
| 1702 | show "norm (f y - f x - A (p + q) *v (y - x)) \<le> norm (y - x) / (Suc (p + q))" | |
| 1703 | using A [OF y] by simp | |
| 1704 | have "norm (A (p + q) *v (y - x) - B *v (y - x)) \<le> onorm(\<lambda>x. (A(p + q) - B) *v x) * norm(y - x)" | |
| 1705 | by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) | |
| 1706 | also have "\<dots> < (e/2) * norm (y - x)" | |
| 1707 | using \<open>y \<noteq> x\<close> pqe2 by auto | |
| 1708 | also have "\<dots> \<le> (e - 1 / (Suc (p + q))) * norm (y - x)" | |
| 1709 | proof (rule mult_right_mono) | |
| 1710 | have "1 / Suc (p + q) \<le> 1 / q" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1711 | using \<open>q \<noteq> 0\<close> by (auto simp: field_split_simps) | 
| 67998 | 1712 | also have "\<dots> < e/2" | 
| 1713 | using qe2 by auto | |
| 1714 | finally show "e / 2 \<le> e - 1 / real (Suc (p + q))" | |
| 1715 | by linarith | |
| 1716 | qed auto | |
| 1717 | finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" | |
| 1718 | by (simp add: algebra_simps) | |
| 1719 | qed | |
| 1720 | then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1721 | using \<open>y \<noteq> x\<close> by (simp add: field_split_simps algebra_simps) | 
| 67998 | 1722 | qed | 
| 1723 | qed | |
| 1724 | then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R | |
| 1725 | norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) | |
| 1726 | (at x within S)" | |
| 71633 | 1727 | by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR) | 
| 67998 | 1728 | qed | 
| 1729 | qed (use x in \<open>simp; auto simp: not_less\<close>) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1730 | ultimately have "f' x = (*v) B" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1731 | by (force simp: algebra_simps scalar_mult_eq_scaleR) | 
| 67998 | 1732 | show "matrix (f' x) $ m $ n \<le> b" | 
| 1733 | proof (rule tendsto_upperbound [of "\<lambda>i. (A i $ m $ n)" _ sequentially]) | |
| 1734 | show "(\<lambda>i. A i $ m $ n) \<longlonglongrightarrow> matrix (f' x) $ m $ n" | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1735 | by (simp add: B \<open>f' x = (*v) B\<close>) | 
| 67998 | 1736 | show "\<forall>\<^sub>F i in sequentially. A i $ m $ n \<le> b" | 
| 1737 | by (simp add: Ab less_eq_real_def) | |
| 1738 | qed auto | |
| 1739 | next | |
| 1740 | fix e :: "real" | |
| 1741 | assume "x \<in> S" and b: "matrix (f' x) $ m $ n \<le> b" and "e > 0" | |
| 1742 | then obtain d where "d>0" | |
| 1743 | and d: "\<And>y. y\<in>S \<Longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> norm (f y - f x - f' x (y - x)) / (norm (y - x)) | |
| 1744 | < e/2" | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70724diff
changeset | 1745 | using f [OF \<open>x \<in> S\<close>] | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70724diff
changeset | 1746 | by (simp add: Deriv.has_derivative_at_within Lim_within) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70724diff
changeset | 1747 | (auto simp add: field_simps dest: spec [of _ "e/2"]) | 
| 67998 | 1748 | let ?A = "matrix(f' x) - (\<chi> i j. if i = m \<and> j = n then e / 4 else 0)" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1749 | obtain B where BRats: "\<And>i j. B$i$j \<in> \<rat>" and Bo_e6: "onorm((*v) (?A - B)) < e/6" | 
| 67998 | 1750 | using matrix_rational_approximation \<open>e > 0\<close> | 
| 1751 | by (metis zero_less_divide_iff zero_less_numeral) | |
| 1752 | show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> | |
| 1753 | (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))" | |
| 1754 | proof (intro exI conjI ballI allI impI) | |
| 1755 | show "d>0" | |
| 1756 | by (rule \<open>d>0\<close>) | |
| 1757 | show "B $ m $ n < b" | |
| 1758 | proof - | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1759 | have "\<bar>matrix ((*v) (?A - B)) $ m $ n\<bar> \<le> onorm ((*v) (?A - B))" | 
| 67998 | 1760 | using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis | 
| 1761 | then show ?thesis | |
| 1762 | using b Bo_e6 by simp | |
| 1763 | qed | |
| 1764 | show "B $ i $ j \<in> \<rat>" for i j | |
| 1765 | using BRats by auto | |
| 1766 | show "norm (f y - f x - B *v (y - x)) \<le> e * norm (y - x)" | |
| 1767 | if "y \<in> S" and y: "norm (y - x) < d" for y | |
| 1768 | proof (cases "y = x") | |
| 1769 | case True then show ?thesis | |
| 1770 | by simp | |
| 1771 | next | |
| 1772 | case False | |
| 1773 | have *: "norm(d' - d) \<le> e/2 \<Longrightarrow> norm(y - (x + d')) < e/2 \<Longrightarrow> norm(y - x - d) \<le> e" for d d' e and x y::"real^'n" | |
| 1774 | using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp | |
| 1775 | show ?thesis | |
| 1776 | proof (rule *) | |
| 1777 | have split246: "\<lbrakk>norm y \<le> e / 6; norm(x - y) \<le> e / 4\<rbrakk> \<Longrightarrow> norm x \<le> e/2" if "e > 0" for e and x y :: "real^'n" | |
| 1778 | using norm_triangle_le [of y "x-y" "e/2"] \<open>e > 0\<close> by simp | |
| 1779 | have "linear (f' x)" | |
| 1780 | using True f has_derivative_linear by blast | |
| 1781 | then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1782 | by (simp add: matrix_vector_mult_diff_rdistrib) | 
| 67998 | 1783 | also have "\<dots> \<le> (e * norm (y - x)) / 2" | 
| 1784 | proof (rule split246) | |
| 1785 | have "norm ((?A - B) *v (y - x)) / norm (y - x) \<le> onorm(\<lambda>x. (?A - B) *v x)" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1786 | by (rule le_onorm) auto | 
| 67998 | 1787 | also have "\<dots> < e/6" | 
| 1788 | by (rule Bo_e6) | |
| 1789 | finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . | |
| 1790 | then show "norm ((?A - B) *v (y - x)) \<le> e * norm (y - x) / 6" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1791 | by (simp add: field_split_simps False) | 
| 67998 | 1792 | have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x))" | 
| 1793 | by (simp add: algebra_simps) | |
| 1794 | also have "\<dots> = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" | |
| 1795 | proof - | |
| 1796 | have "(\<Sum>j\<in>UNIV. (if i = m \<and> j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i | |
| 1797 | proof (cases "i=m") | |
| 1798 | case True then show ?thesis | |
| 1799 | by (auto simp: if_distrib [of "\<lambda>z. z * _"] cong: if_cong) | |
| 1800 | next | |
| 1801 | case False then show ?thesis | |
| 1802 | by (simp add: axis_def) | |
| 1803 | qed | |
| 1804 | then have "(\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" | |
| 1805 | by (auto simp: vec_eq_iff matrix_vector_mult_def) | |
| 1806 | then show ?thesis | |
| 1807 | by metis | |
| 1808 | qed | |
| 1809 | also have "\<dots> \<le> e * norm (y - x) / 4" | |
| 1810 | using \<open>e > 0\<close> apply (simp add: norm_mult abs_mult) | |
| 1811 | by (metis component_le_norm_cart vector_minus_component) | |
| 1812 | finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \<le> e * norm (y - x) / 4" . | |
| 1813 | show "0 < e * norm (y - x)" | |
| 1814 | by (simp add: False \<open>e > 0\<close>) | |
| 1815 | qed | |
| 1816 | finally show "norm (f' x (y - x) - B *v (y - x)) \<le> (e * norm (y - x)) / 2" . | |
| 1817 | show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" | |
| 1818 | using False d [OF \<open>y \<in> S\<close>] y by (simp add: dist_norm field_simps) | |
| 1819 | qed | |
| 1820 | qed | |
| 1821 | qed | |
| 1822 | qed | |
| 1823 | qed auto | |
| 1824 | qed | |
| 1825 | show "negligible ?M" | |
| 1826 | using negligible_subset [OF nN MN] . | |
| 1827 | qed | |
| 1828 | then show ?thesis | |
| 1829 | by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) | |
| 1830 | qed | |
| 1831 | ||
| 1832 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1833 | theorem borel_measurable_det_Jacobian: | 
| 67998 | 1834 |  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 1835 | assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 1836 | shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)" | |
| 1837 | unfolding det_def | |
| 70136 | 1838 | by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) | 
| 67998 | 1839 | |
| 68001 
0a2a1b6507c1
correction of TeX errors and other oversights
 paulson <lp15@cam.ac.uk> parents: 
67999diff
changeset | 1840 | text\<open>The localisation wrt S uses the same argument for many similar results.\<close> | 
| 
0a2a1b6507c1
correction of TeX errors and other oversights
 paulson <lp15@cam.ac.uk> parents: 
67999diff
changeset | 1841 | (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1842 | theorem borel_measurable_lebesgue_on_preimage_borel: | 
| 67998 | 1843 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1844 | assumes "S \<in> sets lebesgue" | |
| 1845 | shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow> | |
| 1846 |          (\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1847 | proof - | 
| 67998 | 1848 |   have "{x. (if x \<in> S then f x else 0) \<in> T} \<in> sets lebesgue \<longleftrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue"
 | 
| 1849 | if "T \<in> sets borel" for T | |
| 1850 | proof (cases "0 \<in> T") | |
| 1851 | case True | |
| 1852 |       then have "{x \<in> S. f x \<in> T} = {x. (if x \<in> S then f x else 0) \<in> T} \<inter> S"
 | |
| 1853 |                 "{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T} \<union> -S"
 | |
| 1854 | by auto | |
| 1855 | then show ?thesis | |
| 1856 | by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) | |
| 1857 | next | |
| 1858 | case False | |
| 1859 |       then have "{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T}"
 | |
| 1860 | by auto | |
| 1861 | then show ?thesis | |
| 1862 | by auto | |
| 1863 | qed | |
| 1864 | then show ?thesis | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 1865 | unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] | 
| 67998 | 1866 | by blast | 
| 1867 | qed | |
| 1868 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1869 | lemma sets_lebesgue_almost_borel: | 
| 67998 | 1870 | assumes "S \<in> sets lebesgue" | 
| 1871 | obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S" | |
| 1872 | proof - | |
| 1873 | obtain T N N' where "S = T \<union> N" "N \<subseteq> N'" "N' \<in> null_sets lborel" "T \<in> sets borel" | |
| 1874 | using sets_completionE [OF assms] by auto | |
| 1875 | then show thesis | |
| 1876 | by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) | |
| 1877 | qed | |
| 1878 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1879 | lemma double_lebesgue_sets: | 
| 67998 | 1880 | assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T" | 
| 1881 |  shows "(\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue) \<longleftrightarrow>
 | |
| 1882 | f \<in> borel_measurable (lebesgue_on S) \<and> | |
| 1883 |           (\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue)"
 | |
| 1884 | (is "?lhs \<longleftrightarrow> _ \<and> ?rhs") | |
| 1885 | unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] | |
| 1886 | proof (intro iffI allI conjI impI, safe) | |
| 1887 | fix V :: "'b set" | |
| 1888 |   assume *: "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
 | |
| 1889 | and "V \<in> sets borel" | |
| 1890 | then have V: "V \<in> sets lebesgue" | |
| 1891 | by simp | |
| 1892 |   have "{x \<in> S. f x \<in> V} = {x \<in> S. f x \<in> T \<inter> V}"
 | |
| 1893 | using fim by blast | |
| 1894 |   also have "{x \<in> S. f x \<in> T \<inter> V} \<in> sets lebesgue"
 | |
| 1895 | using T V * le_inf_iff by blast | |
| 1896 |   finally show "{x \<in> S. f x \<in> V} \<in> sets lebesgue" .
 | |
| 1897 | next | |
| 1898 | fix U :: "'b set" | |
| 1899 |   assume "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
 | |
| 1900 | "negligible U" "U \<subseteq> T" | |
| 1901 |   then show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
 | |
| 1902 | using negligible_imp_sets by blast | |
| 1903 | next | |
| 1904 | fix U :: "'b set" | |
| 1905 |   assume 1 [rule_format]: "(\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
 | |
| 1906 |      and 2 [rule_format]: "\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
 | |
| 1907 | and "U \<in> sets lebesgue" "U \<subseteq> T" | |
| 1908 | then obtain C N where C: "C \<in> sets borel \<and> negligible N \<and> C \<union> N = U" | |
| 1909 | using sets_lebesgue_almost_borel | |
| 1910 | by metis | |
| 1911 |   then have "{x \<in> S. f x \<in> C} \<in> sets lebesgue"
 | |
| 1912 | by (blast intro: 1) | |
| 1913 |   moreover have "{x \<in> S. f x \<in> N} \<in> sets lebesgue"
 | |
| 1914 | using C \<open>U \<subseteq> T\<close> by (blast intro: 2) | |
| 1915 |   moreover have "{x \<in> S. f x \<in> C \<union> N} = {x \<in> S. f x \<in> C} \<union> {x \<in> S. f x \<in> N}"
 | |
| 1916 | by auto | |
| 1917 |   ultimately show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
 | |
| 1918 | using C by auto | |
| 1919 | qed | |
| 1920 | ||
| 1921 | ||
| 69683 | 1922 | subsection\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close> | 
| 67998 | 1923 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1924 | lemma Sard_lemma00: | 
| 67998 | 1925 | fixes P :: "'b::euclidean_space set" | 
| 1926 | assumes "a \<ge> 0" and a: "a *\<^sub>R i \<noteq> 0" and i: "i \<in> Basis" | |
| 1927 |     and P: "P \<subseteq> {x. a *\<^sub>R i \<bullet> x = 0}"
 | |
| 1928 | and "0 \<le> m" "0 \<le> e" | |
| 1929 | obtains S where "S \<in> lmeasurable" | |
| 1930 |             and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
 | |
| 1931 |             and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (DIM('b) - 1)"
 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1932 | proof - | 
| 67998 | 1933 | have "a > 0" | 
| 1934 | using assms by simp | |
| 1935 | let ?v = "(\<Sum>j\<in>Basis. (if j = i then e else m) *\<^sub>R j)" | |
| 1936 | show thesis | |
| 1937 | proof | |
| 1938 | have "- e \<le> x \<bullet> i" "x \<bullet> i \<le> e" | |
| 1939 | if "t \<in> P" "norm (x - t) \<le> e" for x t | |
| 1940 | using \<open>a > 0\<close> that Basis_le_norm [of i "x-t"] P i | |
| 1941 | by (auto simp: inner_commute algebra_simps) | |
| 1942 | moreover have "- m \<le> x \<bullet> j" "x \<bullet> j \<le> m" | |
| 1943 | if "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e" "j \<in> Basis" and "j \<noteq> i" | |
| 1944 | for x t j | |
| 1945 | using that Basis_le_norm [of j x] by auto | |
| 1946 | ultimately | |
| 1947 |     show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> cbox (-?v) ?v"
 | |
| 1948 | by (auto simp: mem_box) | |
| 1949 | have *: "\<forall>k\<in>Basis. - ?v \<bullet> k \<le> ?v \<bullet> k" | |
| 1950 | using \<open>0 \<le> m\<close> \<open>0 \<le> e\<close> by (auto simp: inner_Basis) | |
| 1951 |     have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)"
 | |
| 1952 | by (metis DIM_positive Suc_pred power_Suc) | |
| 1953 |     show "measure lebesgue (cbox (-?v) ?v) \<le> 2 * e * (2 * m) ^ (DIM('b) - 1)"
 | |
| 1954 | using \<open>i \<in> Basis\<close> | |
| 1955 | by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) | |
| 1956 | qed blast | |
| 1957 | qed | |
| 1958 | ||
| 68001 
0a2a1b6507c1
correction of TeX errors and other oversights
 paulson <lp15@cam.ac.uk> parents: 
67999diff
changeset | 1959 | text\<open>As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\<close>
 | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1960 | lemma Sard_lemma0: | 
| 67998 | 1961 |   fixes P :: "(real^'n::{finite,wellorder}) set"
 | 
| 1962 | assumes "a \<noteq> 0" | |
| 1963 |     and P: "P \<subseteq> {x. a \<bullet> x = 0}" and "0 \<le> m" "0 \<le> e"
 | |
| 1964 | obtains S where "S \<in> lmeasurable" | |
| 1965 |     and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
 | |
| 1966 |     and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
 | |
| 1967 | proof - | |
| 1968 | obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" | |
| 1969 | using rotation_rightward_line by metis | |
| 1970 | have Tinv [simp]: "T (inv T x) = x" for x | |
| 1971 | by (simp add: T orthogonal_transformation_surj surj_f_inv_f) | |
| 1972 | obtain S where S: "S \<in> lmeasurable" | |
| 1973 |     and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> T-`P. norm(z - t) \<le> e)} \<subseteq> S"
 | |
| 1974 |     and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
 | |
| 1975 | proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) | |
| 1976 | have "norm a *\<^sub>R axis k 1 \<bullet> x = 0" if "T x \<in> P" for x | |
| 1977 | proof - | |
| 1978 | have "a \<bullet> T x = 0" | |
| 1979 | using P that by blast | |
| 1980 | then show ?thesis | |
| 1981 | by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) | |
| 1982 | qed | |
| 1983 |     then show "T -` P \<subseteq> {x. norm a *\<^sub>R axis k 1 \<bullet> x = 0}"
 | |
| 1984 | by auto | |
| 1985 | qed (use assms T in auto) | |
| 1986 | show thesis | |
| 1987 | proof | |
| 1988 | show "T ` S \<in> lmeasurable" | |
| 1989 | using S measurable_orthogonal_image T by blast | |
| 1990 |     have "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
 | |
| 1991 | proof clarsimp | |
| 1992 | fix x t | |
| 1993 | assume "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e" | |
| 1994 | then have "norm (inv T x) \<le> m" | |
| 1995 | using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) | |
| 1996 | moreover have "\<exists>t\<in>T -` P. norm (inv T x - t) \<le> e" | |
| 1997 | proof | |
| 1998 | have "T (inv T x - inv T t) = x - t" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 1999 | using T linear_diff orthogonal_transformation_def | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 2000 | by (metis (no_types, opaque_lifting) Tinv) | 
| 67998 | 2001 | then have "norm (inv T x - inv T t) = norm (x - t)" | 
| 2002 | by (metis T orthogonal_transformation_norm) | |
| 2003 | then show "norm (inv T x - inv T t) \<le> e" | |
| 2004 | using \<open>norm (x - t) \<le> e\<close> by linarith | |
| 2005 | next | |
| 2006 | show "inv T t \<in> T -` P" | |
| 2007 | using \<open>t \<in> P\<close> by force | |
| 2008 | qed | |
| 2009 |       ultimately show "x \<in> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
 | |
| 2010 | by force | |
| 2011 | qed | |
| 2012 |     then show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` S"
 | |
| 2013 | using image_mono [OF subS] by (rule order_trans) | |
| 2014 |     show "measure lebesgue (T ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
 | |
| 2015 | using mS T by (simp add: S measure_orthogonal_image) | |
| 2016 | qed | |
| 2017 | qed | |
| 2018 | ||
| 68001 
0a2a1b6507c1
correction of TeX errors and other oversights
 paulson <lp15@cam.ac.uk> parents: 
67999diff
changeset | 2019 | text\<open>As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\<close>
 | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2020 | lemma Sard_lemma1: | 
| 67998 | 2021 |   fixes P :: "(real^'n::{finite,wellorder}) set"
 | 
| 2022 |    assumes P: "dim P < CARD('n)" and "0 \<le> m" "0 \<le> e"
 | |
| 2023 | obtains S where "S \<in> lmeasurable" | |
| 2024 |             and "{z. norm(z - w) \<le> m \<and> (\<exists>t \<in> P. norm(z - w - t) \<le> e)} \<subseteq> S"
 | |
| 2025 |             and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2026 | proof - | 
| 67998 | 2027 |   obtain a where "a \<noteq> 0" "P \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2028 | using lowdim_subset_hyperplane [of P] P span_base by auto | 
| 67998 | 2029 | then obtain S where S: "S \<in> lmeasurable" | 
| 2030 |     and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
 | |
| 2031 |     and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
 | |
| 2032 | by (rule Sard_lemma0 [OF _ _ \<open>0 \<le> m\<close> \<open>0 \<le> e\<close>]) | |
| 2033 | show thesis | |
| 2034 | proof | |
| 2035 | show "(+)w ` S \<in> lmeasurable" | |
| 2036 | by (metis measurable_translation S) | |
| 2037 |     show "{z. norm (z - w) \<le> m \<and> (\<exists>t\<in>P. norm (z - w - t) \<le> e)} \<subseteq> (+)w ` S"
 | |
| 2038 | using subS by force | |
| 2039 |     show "measure lebesgue ((+)w ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
 | |
| 2040 | by (metis measure_translation mS) | |
| 2041 | qed | |
| 2042 | qed | |
| 2043 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2044 | lemma Sard_lemma2: | 
| 67998 | 2045 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
 | 
| 2046 |   assumes mlen: "CARD('m) \<le> CARD('n)" (is "?m \<le> ?n")
 | |
| 2047 | and "B > 0" "bounded S" | |
| 2048 | and derS: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 2049 |     and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
 | |
| 2050 | and B: "\<And>x. x \<in> S \<Longrightarrow> onorm(f' x) \<le> B" | |
| 2051 | shows "negligible(f ` S)" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2052 | proof - | 
| 67998 | 2053 | have lin_f': "\<And>x. x \<in> S \<Longrightarrow> linear(f' x)" | 
| 2054 | using derS has_derivative_linear by blast | |
| 2055 | show ?thesis | |
| 2056 | proof (clarsimp simp add: negligible_outer_le) | |
| 2057 | fix e :: "real" | |
| 2058 | assume "e > 0" | |
| 2059 | obtain c where csub: "S \<subseteq> cbox (- (vec c)) (vec c)" and "c > 0" | |
| 2060 | proof - | |
| 2061 | obtain b where b: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> b" | |
| 2062 | using \<open>bounded S\<close> by (auto simp: bounded_iff) | |
| 2063 | show thesis | |
| 2064 | proof | |
| 2065 | have "- \<bar>b\<bar> - 1 \<le> x $ i \<and> x $ i \<le> \<bar>b\<bar> + 1" if "x \<in> S" for x i | |
| 2066 | using component_le_norm_cart [of x i] b [OF that] by auto | |
| 2067 | then show "S \<subseteq> cbox (- vec (\<bar>b\<bar> + 1)) (vec (\<bar>b\<bar> + 1))" | |
| 2068 | by (auto simp: mem_box_cart) | |
| 2069 | qed auto | |
| 2070 | qed | |
| 2071 |     then have box_cc: "box (- (vec c)) (vec c) \<noteq> {}" and cbox_cc: "cbox (- (vec c)) (vec c) \<noteq> {}"
 | |
| 2072 | by (auto simp: interval_eq_empty_cart) | |
| 2073 | obtain d where "d > 0" "d \<le> B" | |
| 2074 | and d: "(d * 2) * (4 * B) ^ (?n - 1) \<le> e / (2*c) ^ ?m / ?m ^ ?m" | |
| 2075 | apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) | |
| 2076 | using \<open>B > 0\<close> \<open>c > 0\<close> \<open>e > 0\<close> | |
| 2077 | by (simp_all add: divide_simps min_mult_distrib_right) | |
| 2078 | have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and> | |
| 2079 | (x \<in> S | |
| 2080 | \<longrightarrow> (\<forall>y. y \<in> S \<and> norm(y - x) < r | |
| 2081 | \<longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)))" for x | |
| 2082 | proof (cases "x \<in> S") | |
| 2083 | case True | |
| 2084 | then obtain r where "r > 0" | |
| 2085 | and "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < r\<rbrakk> | |
| 2086 | \<Longrightarrow> norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)" | |
| 2087 | using derS \<open>d > 0\<close> by (force simp: has_derivative_within_alt) | |
| 2088 | then show ?thesis | |
| 2089 | by (rule_tac x="min r (1/2)" in exI) simp | |
| 2090 | next | |
| 2091 | case False | |
| 2092 | then show ?thesis | |
| 2093 | by (rule_tac x="1/2" in exI) simp | |
| 2094 | qed | |
| 2095 | then obtain r where r12: "\<And>x. 0 < r x \<and> r x \<le> 1/2" | |
| 2096 | and r: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < r x\<rbrakk> | |
| 2097 | \<Longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)" | |
| 2098 | by metis | |
| 2099 | then have ga: "gauge (\<lambda>x. ball x (r x))" | |
| 2100 | by (auto simp: gauge_def) | |
| 2101 | obtain \<D> where \<D>: "countable \<D>" and sub_cc: "\<Union>\<D> \<subseteq> cbox (- vec c) (vec c)" | |
| 2102 |       and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>u v. K = cbox u v)"
 | |
| 2103 |       and djointish: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 2104 | and covered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> ball x (r x)" | |
| 2105 | and close: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i::'m. v $ i - u $ i = 2*c / 2^n" | |
| 2106 | and covers: "S \<subseteq> \<Union>\<D>" | |
| 2107 | apply (rule covering_lemma [OF csub box_cc ga]) | |
| 2108 | apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) | |
| 2109 | done | |
| 2110 | let ?\<mu> = "measure lebesgue" | |
| 2111 | have "\<exists>T. T \<in> lmeasurable \<and> f ` (K \<inter> S) \<subseteq> T \<and> ?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K" | |
| 2112 | if "K \<in> \<D>" for K | |
| 2113 | proof - | |
| 2114 | obtain u v where uv: "K = cbox u v" | |
| 2115 | using cbox \<open>K \<in> \<D>\<close> by blast | |
| 2116 |       then have uv_ne: "cbox u v \<noteq> {}"
 | |
| 2117 | using cbox that by fastforce | |
| 2118 | obtain x where x: "x \<in> S \<inter> cbox u v" "cbox u v \<subseteq> ball x (r x)" | |
| 2119 | using \<open>K \<in> \<D>\<close> covered uv by blast | |
| 2120 | then have "dim (range (f' x)) < ?n" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2121 | using rank_dim_range [of "matrix (f' x)"] x rank[of x] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2122 | by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') | 
| 67998 | 2123 | then obtain T where T: "T \<in> lmeasurable" | 
| 2124 |             and subT: "{z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))} \<subseteq> T"
 | |
| 2125 | and measT: "?\<mu> T \<le> (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" | |
| 2126 | (is "_ \<le> ?DVU") | |
| 2127 | apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) | |
| 2128 | using \<open>B > 0\<close> \<open>d > 0\<close> by simp_all | |
| 2129 | show ?thesis | |
| 2130 | proof (intro exI conjI) | |
| 2131 |         have "f ` (K \<inter> S) \<subseteq> {z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))}"
 | |
| 2132 | unfolding uv | |
| 2133 | proof (clarsimp simp: mult.assoc, intro conjI) | |
| 2134 | fix y | |
| 2135 | assume y: "y \<in> cbox u v" and "y \<in> S" | |
| 2136 | then have "norm (y - x) < r x" | |
| 2137 | by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) | |
| 2138 | then have le_dyx: "norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)" | |
| 2139 | using r [of x y] x \<open>y \<in> S\<close> by blast | |
| 2140 | have yx_le: "norm (y - x) \<le> norm (v - u)" | |
| 2141 | proof (rule norm_le_componentwise_cart) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2142 | show "norm ((y - x) $ i) \<le> norm ((v - u) $ i)" for i | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2143 | using x y by (force simp: mem_box_cart dest!: spec [where x=i]) | 
| 67998 | 2144 | qed | 
| 2145 | have *: "\<lbrakk>norm(y - x - z) \<le> d; norm z \<le> B; d \<le> B\<rbrakk> \<Longrightarrow> norm(y - x) \<le> 2 * B" | |
| 2146 | for x y z :: "real^'n::_" and d B | |
| 2147 | using norm_triangle_ineq2 [of "y - x" z] by auto | |
| 2148 | show "norm (f y - f x) \<le> 2 * (B * norm (v - u))" | |
| 2149 | proof (rule * [OF le_dyx]) | |
| 2150 | have "norm (f' x (y - x)) \<le> onorm (f' x) * norm (y - x)" | |
| 2151 | using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) | |
| 2152 | also have "\<dots> \<le> B * norm (v - u)" | |
| 2153 | proof (rule mult_mono) | |
| 2154 | show "onorm (f' x) \<le> B" | |
| 2155 | using B x by blast | |
| 2156 | qed (use \<open>B > 0\<close> yx_le in auto) | |
| 2157 | finally show "norm (f' x (y - x)) \<le> B * norm (v - u)" . | |
| 2158 | show "d * norm (y - x) \<le> B * norm (v - u)" | |
| 2159 | using \<open>B > 0\<close> by (auto intro: mult_mono [OF \<open>d \<le> B\<close> yx_le]) | |
| 2160 | qed | |
| 2161 | show "\<exists>t. norm (f y - f x - f' x t) \<le> d * norm (v - u)" | |
| 2162 | apply (rule_tac x="y-x" in exI) | |
| 2163 | using \<open>d > 0\<close> yx_le le_dyx mult_left_mono [where c=d] | |
| 72569 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
72445diff
changeset | 2164 | by (meson order_trans mult_le_cancel_iff2) | 
| 67998 | 2165 | qed | 
| 2166 | with subT show "f ` (K \<inter> S) \<subseteq> T" by blast | |
| 2167 | show "?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K" | |
| 2168 | proof (rule order_trans [OF measT]) | |
| 2169 | have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" | |
| 2170 | using \<open>c > 0\<close> | |
| 71633 | 2171 | apply (simp add: algebra_simps) | 
| 67998 | 2172 | by (metis Suc_pred power_Suc zero_less_card_finite) | 
| 2173 | also have "\<dots> \<le> (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" | |
| 2174 | by (rule mult_right_mono [OF d]) auto | |
| 2175 | also have "\<dots> \<le> e / (2*c) ^ ?m * ?\<mu> K" | |
| 2176 | proof - | |
| 2177 | have "u \<in> ball (x) (r x)" "v \<in> ball x (r x)" | |
| 2178 | using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ | |
| 2179 | moreover have "r x \<le> 1/2" | |
| 2180 | using r12 by auto | |
| 2181 | ultimately have "norm (v - u) \<le> 1" | |
| 2182 | using norm_triangle_half_r [of x u 1 v] | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 2183 | by (metis (no_types, opaque_lifting) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) | 
| 67998 | 2184 | then have "norm (v - u) ^ ?n \<le> norm (v - u) ^ ?m" | 
| 2185 | by (simp add: power_decreasing [OF mlen]) | |
| 2186 | also have "\<dots> \<le> ?\<mu> K * real (?m ^ ?m)" | |
| 2187 | proof - | |
| 2188 | obtain n where n: "\<And>i. v$i - u$i = 2 * c / 2^n" | |
| 2189 | using close [of u v] \<open>K \<in> \<D>\<close> uv by blast | |
| 2190 | have "norm (v - u) ^ ?m \<le> (\<Sum>i\<in>UNIV. \<bar>(v - u) $ i\<bar>) ^ ?m" | |
| 2191 | by (intro norm_le_l1_cart power_mono) auto | |
| 2192 |               also have "\<dots> \<le> (\<Prod>i\<in>UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)"
 | |
| 2193 | by (simp add: n field_simps \<open>c > 0\<close> less_eq_real_def) | |
| 2194 | also have "\<dots> = ?\<mu> K * real (?m ^ ?m)" | |
| 2195 | by (simp add: uv uv_ne content_cbox_cart) | |
| 2196 | finally show ?thesis . | |
| 2197 | qed | |
| 2198 | finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \<le> ?\<mu> K" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2199 | by (simp add: field_split_simps) | 
| 67998 | 2200 | show ?thesis | 
| 2201 | using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \<open>c > 0\<close> \<open>e > 0\<close> by auto | |
| 2202 | qed | |
| 2203 | finally show "?DVU \<le> e / (2*c) ^ ?m * ?\<mu> K" . | |
| 2204 | qed | |
| 2205 | qed (use T in auto) | |
| 2206 | qed | |
| 2207 | then obtain g where meas_g: "\<And>K. K \<in> \<D> \<Longrightarrow> g K \<in> lmeasurable" | |
| 2208 | and sub_g: "\<And>K. K \<in> \<D> \<Longrightarrow> f ` (K \<inter> S) \<subseteq> g K" | |
| 2209 | and le_g: "\<And>K. K \<in> \<D> \<Longrightarrow> ?\<mu> (g K) \<le> e / (2*c)^?m * ?\<mu> K" | |
| 2210 | by metis | |
| 2211 | have le_e: "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> e" | |
| 2212 | if "\<F> \<subseteq> \<D>" "finite \<F>" for \<F> | |
| 2213 | proof - | |
| 2214 | have "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> (\<Sum>i\<in>\<F>. ?\<mu> (g i))" | |
| 2215 | using meas_g \<open>\<F> \<subseteq> \<D>\<close> by (auto intro: measure_UNION_le [OF \<open>finite \<F>\<close>]) | |
| 2216 | also have "\<dots> \<le> (\<Sum>K\<in>\<F>. e / (2*c) ^ ?m * ?\<mu> K)" | |
| 2217 | using \<open>\<F> \<subseteq> \<D>\<close> sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) | |
| 2218 | also have "\<dots> = e / (2*c) ^ ?m * (\<Sum>K\<in>\<F>. ?\<mu> K)" | |
| 2219 | by (simp add: sum_distrib_left) | |
| 2220 | also have "\<dots> \<le> e" | |
| 2221 | proof - | |
| 2222 | have "\<F> division_of \<Union>\<F>" | |
| 2223 | proof (rule division_ofI) | |
| 2224 |           show "K \<subseteq> \<Union>\<F>"  "K \<noteq> {}" "\<exists>a b. K = cbox a b" if "K \<in> \<F>" for K
 | |
| 2225 | using \<open>K \<in> \<F>\<close> covered cbox \<open>\<F> \<subseteq> \<D>\<close> by (auto simp: Union_upper) | |
| 2226 |           show "interior K \<inter> interior L = {}" if "K \<in> \<F>" and "L \<in> \<F>" and "K \<noteq> L" for K L
 | |
| 2227 | by (metis (mono_tags, lifting) \<open>\<F> \<subseteq> \<D>\<close> pairwiseD djointish pairwise_subset that) | |
| 2228 | qed (use that in auto) | |
| 2229 | then have "sum ?\<mu> \<F> \<le> ?\<mu> (\<Union>\<F>)" | |
| 2230 | by (simp add: content_division) | |
| 2231 | also have "\<dots> \<le> ?\<mu> (cbox (- vec c) (vec c) :: (real, 'm) vec set)" | |
| 2232 | proof (rule measure_mono_fmeasurable) | |
| 2233 | show "\<Union>\<F> \<subseteq> cbox (- vec c) (vec c)" | |
| 2234 | by (meson Sup_subset_mono sub_cc order_trans \<open>\<F> \<subseteq> \<D>\<close>) | |
| 2235 | qed (use \<open>\<F> division_of \<Union>\<F>\<close> lmeasurable_division in auto) | |
| 2236 | also have "\<dots> = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" | |
| 2237 | by simp | |
| 2238 | also have "\<dots> \<le> (2 ^ ?m * c ^ ?m)" | |
| 2239 | using \<open>c > 0\<close> by (simp add: content_cbox_if_cart) | |
| 2240 | finally have "sum ?\<mu> \<F> \<le> (2 ^ ?m * c ^ ?m)" . | |
| 2241 | then show ?thesis | |
| 71633 | 2242 | using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps) | 
| 67998 | 2243 | qed | 
| 2244 | finally show ?thesis . | |
| 2245 | qed | |
| 2246 | show "\<exists>T. f ` S \<subseteq> T \<and> T \<in> lmeasurable \<and> ?\<mu> T \<le> e" | |
| 2247 | proof (intro exI conjI) | |
| 69325 | 2248 | show "f ` S \<subseteq> \<Union> (g ` \<D>)" | 
| 67998 | 2249 | using covers sub_g by force | 
| 69325 | 2250 | show "\<Union> (g ` \<D>) \<in> lmeasurable" | 
| 67998 | 2251 | by (rule fmeasurable_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e]) | 
| 69325 | 2252 | show "?\<mu> (\<Union> (g ` \<D>)) \<le> e" | 
| 67998 | 2253 | by (rule measure_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e]) | 
| 2254 | qed | |
| 2255 | qed | |
| 2256 | qed | |
| 2257 | ||
| 2258 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2259 | theorem baby_Sard: | 
| 67998 | 2260 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
 | 
| 2261 |   assumes mlen: "CARD('m) \<le> CARD('n)"
 | |
| 2262 | and der: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 2263 |     and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
 | |
| 2264 | shows "negligible(f ` S)" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2265 | proof - | 
| 67998 | 2266 |   let ?U = "\<lambda>n. {x \<in> S. norm(x) \<le> n \<and> onorm(f' x) \<le> real n}"
 | 
| 2267 | have "\<And>x. x \<in> S \<Longrightarrow> \<exists>n. norm x \<le> real n \<and> onorm (f' x) \<le> real n" | |
| 2268 | by (meson linear order_trans real_arch_simple) | |
| 2269 | then have eq: "S = (\<Union>n. ?U n)" | |
| 2270 | by auto | |
| 2271 | have "negligible (f ` ?U n)" for n | |
| 2272 | proof (rule Sard_lemma2 [OF mlen]) | |
| 2273 | show "0 < real n + 1" | |
| 2274 | by auto | |
| 2275 | show "bounded (?U n)" | |
| 2276 | using bounded_iff by blast | |
| 2277 | show "(f has_derivative f' x) (at x within ?U n)" if "x \<in> ?U n" for x | |
| 2278 | using der that by (force intro: has_derivative_subset) | |
| 2279 | qed (use rank in auto) | |
| 2280 | then show ?thesis | |
| 2281 | by (subst eq) (simp add: image_Union negligible_Union_nat) | |
| 2282 | qed | |
| 2283 | ||
| 2284 | ||
| 69683 | 2285 | subsection\<open>A one-way version of change-of-variables not assuming injectivity. \<close> | 
| 67998 | 2286 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2287 | lemma integral_on_image_ubound_weak: | 
| 67998 | 2288 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
 | 
| 2289 | assumes S: "S \<in> sets lebesgue" | |
| 2290 | and f: "f \<in> borel_measurable (lebesgue_on (g ` S))" | |
| 2291 | and nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)" | |
| 2292 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2293 | and det_int_fg: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S" | |
| 2294 |       and meas_gim: "\<And>T. \<lbrakk>T \<subseteq> g ` S; T \<in> sets lebesgue\<rbrakk> \<Longrightarrow> {x \<in> S. g x \<in> T} \<in> sets lebesgue"
 | |
| 2295 | shows "f integrable_on (g ` S) \<and> | |
| 2296 | integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))" | |
| 2297 | (is "_ \<and> _ \<le> ?b") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2298 | proof - | 
| 67998 | 2299 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar>" | 
| 2300 | have cont_g: "continuous_on S g" | |
| 2301 | using der_g has_derivative_continuous_on by blast | |
| 2302 | have [simp]: "space (lebesgue_on S) = S" | |
| 2303 | by (simp add: S) | |
| 2304 | have gS_in_sets_leb: "g ` S \<in> sets lebesgue" | |
| 2305 | apply (rule differentiable_image_in_sets_lebesgue) | |
| 2306 | using der_g by (auto simp: S differentiable_def differentiable_on_def) | |
| 2307 | obtain h where nonneg_h: "\<And>n x. 0 \<le> h n x" | |
| 2308 | and h_le_f: "\<And>n x. x \<in> S \<Longrightarrow> h n (g x) \<le> f (g x)" | |
| 2309 | and h_inc: "\<And>n x. h n x \<le> h (Suc n) x" | |
| 2310 | and h_meas: "\<And>n. h n \<in> borel_measurable lebesgue" | |
| 2311 | and fin_R: "\<And>n. finite(range (h n))" | |
| 2312 | and lim: "\<And>x. x \<in> g ` S \<Longrightarrow> (\<lambda>n. h n x) \<longlonglongrightarrow> f x" | |
| 2313 | proof - | |
| 2314 | let ?f = "\<lambda>x. if x \<in> g ` S then f x else 0" | |
| 2315 | have "?f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> ?f x)" | |
| 2316 | by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) | |
| 2317 | then show ?thesis | |
| 2318 | apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) | |
| 2319 | apply (rename_tac h) | |
| 2320 | by (rule_tac h=h in that) (auto split: if_split_asm) | |
| 2321 | qed | |
| 2322 |   have h_lmeas: "{t. h n (g t) = y} \<inter> S \<in> sets lebesgue" for y n
 | |
| 2323 | proof - | |
| 2324 | have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" | |
| 2325 | by simp | |
| 2326 |     then have "((h n) -`{y} \<inter> g ` S) \<in> sets (lebesgue_on (g ` S))"
 | |
| 2327 | by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) | |
| 2328 |     then have "({u. h n u = y} \<inter> g ` S) \<in> sets lebesgue"
 | |
| 2329 | using gS_in_sets_leb | |
| 2330 | by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) | |
| 2331 |     then have "{x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)} \<in> sets lebesgue"
 | |
| 2332 |       using meas_gim[of "({u. h n u = y} \<inter> g ` S)"] by force
 | |
| 2333 |     moreover have "{t. h n (g t) = y} \<inter> S = {x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)}"
 | |
| 2334 | by blast | |
| 2335 | ultimately show ?thesis | |
| 2336 | by auto | |
| 2337 | qed | |
| 2338 | have hint: "h n integrable_on g ` S \<and> integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * h n (g x))" | |
| 2339 | (is "?INT \<and> ?lhs \<le> ?rhs") for n | |
| 2340 | proof - | |
| 2341 | let ?R = "range (h n)" | |
| 2342 |     have hn_eq: "h n = (\<lambda>x. \<Sum>y\<in>?R. y * indicat_real {x. h n x = y} x)"
 | |
| 2343 | by (simp add: indicator_def if_distrib fin_R cong: if_cong) | |
| 2344 |     have yind: "(\<lambda>t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \<and>
 | |
| 2345 |                 (integral (g ` S) (\<lambda>t. y * indicator {x. h n x = y} t))
 | |
| 2346 |                  \<le> integral S (\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y * indicator {x. h n x = y} (g t))"
 | |
| 2347 | if y: "y \<in> ?R" for y::real | |
| 2348 | proof (cases "y=0") | |
| 2349 | case True | |
| 2350 | then show ?thesis using gS_in_sets_leb integrable_0 by force | |
| 2351 | next | |
| 2352 | case False | |
| 2353 | with that have "y > 0" | |
| 2354 | using less_eq_real_def nonneg_h by fastforce | |
| 2355 |       have "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) integrable_on S"
 | |
| 2356 | proof (rule measurable_bounded_by_integrable_imp_integrable) | |
| 2357 |         have "(\<lambda>x. ?D x) \<in> borel_measurable (lebesgue_on ({t. h n (g t) = y} \<inter> S))"
 | |
| 2358 | apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2359 | by (meson der_g IntD2 has_derivative_subset inf_le2) | 
| 67998 | 2360 |         then have "(\<lambda>x. if x \<in> {t. h n (g t) = y} \<inter> S then ?D x else 0) \<in> borel_measurable lebesgue"
 | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 2361 | by (rule borel_measurable_if_I [OF _ h_lmeas]) | 
| 67998 | 2362 |         then show "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) \<in> borel_measurable (lebesgue_on S)"
 | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 2363 | by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) | 
| 67998 | 2364 | show "(\<lambda>x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" | 
| 2365 | by (rule integrable_cmul) (use det_int_fg in auto) | |
| 2366 |         show "norm (if x \<in> {t. h n (g t) = y} then ?D x else 0) \<le> ?D x *\<^sub>R f (g x) /\<^sub>R y"
 | |
| 2367 | if "x \<in> S" for x | |
| 2368 | using nonneg_h [of n x] \<open>y > 0\<close> nonneg_fg [of x] h_le_f [of x n] that | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2369 | by (auto simp: divide_simps mult_left_mono) | 
| 67998 | 2370 | qed (use S in auto) | 
| 2371 |       then have int_det: "(\<lambda>t. \<bar>det (matrix (g' t))\<bar>) integrable_on ({t. h n (g t) = y} \<inter> S)"
 | |
| 2372 | using integrable_restrict_Int by force | |
| 2373 |       have "(g ` ({t. h n (g t) = y} \<inter> S)) \<in> lmeasurable"
 | |
| 2374 | apply (rule measurable_differentiable_image [OF h_lmeas]) | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2375 | apply (blast intro: has_derivative_subset [OF der_g]) | 
| 67998 | 2376 | apply (rule int_det) | 
| 2377 | done | |
| 2378 |       moreover have "g ` ({t. h n (g t) = y} \<inter> S) = {x. h n x = y} \<inter> g ` S"
 | |
| 2379 | by blast | |
| 2380 |       moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \<inter> S))
 | |
| 2381 |                      \<le> integral ({t. h n (g t) = y} \<inter> S) (\<lambda>t. \<bar>det (matrix (g' t))\<bar>)"
 | |
| 2382 | apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2383 | apply (blast intro: has_derivative_subset [OF der_g]) | 
| 67998 | 2384 | done | 
| 2385 | ultimately show ?thesis | |
| 2386 |         using \<open>y > 0\<close> integral_restrict_Int [of S "{t. h n (g t) = y}" "\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y"]
 | |
| 71633 | 2387 | apply (simp add: integrable_on_indicator integral_indicator) | 
| 73536 | 2388 | apply (simp add: indicator_def of_bool_def if_distrib cong: if_cong) | 
| 67998 | 2389 | done | 
| 2390 | qed | |
| 2391 | have hn_int: "h n integrable_on g ` S" | |
| 2392 | apply (subst hn_eq) | |
| 2393 | using yind by (force intro: integrable_sum [OF fin_R]) | |
| 2394 | then show ?thesis | |
| 2395 | proof | |
| 2396 |       have "?lhs = integral (g ` S) (\<lambda>x. \<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} x)"
 | |
| 2397 | by (metis hn_eq) | |
| 2398 |       also have "\<dots> = (\<Sum>y\<in>range (h n). integral (g ` S) (\<lambda>x. y * indicat_real {x. h n x = y} x))"
 | |
| 2399 | by (rule integral_sum [OF fin_R]) (use yind in blast) | |
| 2400 |       also have "\<dots> \<le> (\<Sum>y\<in>range (h n). integral S (\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)))"
 | |
| 2401 | using yind by (force intro: sum_mono) | |
| 2402 |       also have "\<dots> = integral S (\<lambda>u. \<Sum>y\<in>range (h n). \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u))"
 | |
| 2403 | proof (rule integral_sum [OF fin_R, symmetric]) | |
| 2404 | fix y assume y: "y \<in> ?R" | |
| 2405 | with nonneg_h have "y \<ge> 0" | |
| 2406 | by auto | |
| 2407 |         show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) integrable_on S"
 | |
| 2408 | proof (rule measurable_bounded_by_integrable_imp_integrable) | |
| 2409 |           have "(\<lambda>x. indicat_real {x. h n x = y} (g x)) \<in> borel_measurable (lebesgue_on S)"
 | |
| 2410 | using h_lmeas S | |
| 2411 | by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) | |
| 2412 |           then show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) \<in> borel_measurable (lebesgue_on S)"
 | |
| 2413 | by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) | |
| 2414 | next | |
| 2415 | fix x | |
| 2416 | assume "x \<in> S" | |
| 2417 |           have "y * indicat_real {x. h n x = y} (g x) \<le> f (g x)"
 | |
| 73536 | 2418 | by (metis \<open>x \<in> S\<close> h_le_f indicator_simps(1) indicator_simps(2) mem_Collect_eq mult.commute mult.left_neutral mult_zero_left nonneg_fg) | 
| 67998 | 2419 |           with \<open>y \<ge> 0\<close> show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \<le> ?D x * f(g x)"
 | 
| 2420 | by (simp add: abs_mult mult.assoc mult_left_mono) | |
| 2421 | qed (use S det_int_fg in auto) | |
| 2422 | qed | |
| 2423 | also have "\<dots> = integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> * | |
| 2424 |                                         (\<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} (g T)))"
 | |
| 2425 | by (simp add: sum_distrib_left mult.assoc) | |
| 2426 | also have "\<dots> = ?rhs" | |
| 2427 | by (metis hn_eq) | |
| 2428 | finally show "integral (g ` S) (h n) \<le> ?rhs" . | |
| 2429 | qed | |
| 2430 | qed | |
| 2431 | have le: "integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<le> ?b" for n | |
| 2432 | proof (rule integral_le) | |
| 2433 | show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) integrable_on S" | |
| 2434 | proof (rule measurable_bounded_by_integrable_imp_integrable) | |
| 2435 | have "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> *\<^sub>R h n (g T)) \<in> borel_measurable (lebesgue_on S)" | |
| 2436 | proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \<open>S \<in> sets lebesgue\<close>) | |
| 2437 |         have eq: "{x \<in> S. f x \<le> a} = (\<Union>b \<in> (f ` S) \<inter> atMost a. {x. f x = b} \<inter> S)" for f and a::real
 | |
| 2438 | by auto | |
| 2439 |         have "finite ((\<lambda>x. h n (g x)) ` S \<inter> {..a})" for a
 | |
| 2440 | by (force intro: finite_subset [OF _ fin_R]) | |
| 2441 | with h_lmeas [of n] show "(\<lambda>x. h n (g x)) \<in> borel_measurable (lebesgue_on S)" | |
| 2442 | apply (simp add: borel_measurable_vimage_halfspace_component_le \<open>S \<in> sets lebesgue\<close> sets_restrict_space_iff eq) | |
| 2443 | by (metis (mono_tags) SUP_inf sets.finite_UN) | |
| 2444 | qed (use der_g in blast) | |
| 2445 | then show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<in> borel_measurable (lebesgue_on S)" | |
| 2446 | by simp | |
| 2447 | show "norm (?D x * h n (g x)) \<le> ?D x *\<^sub>R f (g x)" | |
| 2448 | if "x \<in> S" for x | |
| 2449 | by (simp add: h_le_f mult_left_mono nonneg_h that) | |
| 2450 | qed (use S det_int_fg in auto) | |
| 2451 | show "?D x * h n (g x) \<le> ?D x * f (g x)" if "x \<in> S" for x | |
| 2452 | by (simp add: \<open>x \<in> S\<close> h_le_f mult_left_mono) | |
| 2453 | show "(\<lambda>x. ?D x * f (g x)) integrable_on S" | |
| 2454 | using det_int_fg by blast | |
| 2455 | qed | |
| 2456 | have "f integrable_on g ` S \<and> (\<lambda>k. integral (g ` S) (h k)) \<longlonglongrightarrow> integral (g ` S) f" | |
| 2457 | proof (rule monotone_convergence_increasing) | |
| 2458 | have "\<bar>integral (g ` S) (h n)\<bar> \<le> integral S (\<lambda>x. ?D x * f (g x))" for n | |
| 2459 | proof - | |
| 2460 | have "\<bar>integral (g ` S) (h n)\<bar> = integral (g ` S) (h n)" | |
| 2461 | using hint by (simp add: integral_nonneg nonneg_h) | |
| 2462 | also have "\<dots> \<le> integral S (\<lambda>x. ?D x * f (g x))" | |
| 2463 | using hint le by (meson order_trans) | |
| 2464 | finally show ?thesis . | |
| 2465 | qed | |
| 2466 | then show "bounded (range (\<lambda>k. integral (g ` S) (h k)))" | |
| 2467 | by (force simp: bounded_iff) | |
| 2468 | qed (use h_inc lim hint in auto) | |
| 2469 | moreover have "integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * f (g x))" for n | |
| 2470 | using hint by (blast intro: le order_trans) | |
| 2471 | ultimately show ?thesis | |
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68403diff
changeset | 2472 | by (auto intro: Lim_bounded) | 
| 67998 | 2473 | qed | 
| 2474 | ||
| 2475 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2476 | lemma integral_on_image_ubound_nonneg: | 
| 67998 | 2477 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
 | 
| 2478 | assumes nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)" | |
| 2479 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2480 | and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S" | |
| 2481 | shows "f integrable_on (g ` S) \<and> integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))" | |
| 2482 | (is "_ \<and> _ \<le> ?b") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2483 | proof - | 
| 67998 | 2484 | let ?D = "\<lambda>x. det (matrix (g' x))" | 
| 2485 |   define S' where "S' \<equiv> {x \<in> S. ?D x * f(g x) \<noteq> 0}"
 | |
| 2486 | then have der_gS': "\<And>x. x \<in> S' \<Longrightarrow> (g has_derivative g' x) (at x within S')" | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2487 | by (metis (mono_tags, lifting) der_g has_derivative_subset mem_Collect_eq subset_iff) | 
| 67998 | 2488 | have "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) integrable_on UNIV" | 
| 2489 | by (simp add: integrable_restrict_UNIV intS) | |
| 2490 | then have Df_borel: "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> borel_measurable lebesgue" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2491 | using integrable_imp_measurable lebesgue_on_UNIV_eq by force | 
| 67998 | 2492 | have S': "S' \<in> sets lebesgue" | 
| 2493 | proof - | |
| 2494 | from Df_borel borel_measurable_vimage_open [of _ UNIV] | |
| 2495 |     have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> T} \<in> sets lebesgue"
 | |
| 2496 | if "open T" for T | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2497 | using that unfolding lebesgue_on_UNIV_eq | 
| 67998 | 2498 | by (fastforce simp add: dest!: spec) | 
| 2499 |     then have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> -{0}} \<in> sets lebesgue"
 | |
| 2500 | using open_Compl by blast | |
| 2501 | then show ?thesis | |
| 2502 | by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) | |
| 2503 | qed | |
| 2504 | then have gS': "g ` S' \<in> sets lebesgue" | |
| 2505 | proof (rule differentiable_image_in_sets_lebesgue) | |
| 2506 | show "g differentiable_on S'" | |
| 2507 | using der_g unfolding S'_def differentiable_def differentiable_on_def | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2508 | by (blast intro: has_derivative_subset) | 
| 67998 | 2509 | qed auto | 
| 2510 | have f: "f \<in> borel_measurable (lebesgue_on (g ` S'))" | |
| 2511 | proof (clarsimp simp add: borel_measurable_vimage_open) | |
| 2512 | fix T :: "real set" | |
| 2513 | assume "open T" | |
| 2514 |     have "{x \<in> g ` S'. f x \<in> T} = g ` {x \<in> S'. f(g x) \<in> T}"
 | |
| 2515 | by blast | |
| 2516 |     moreover have "g ` {x \<in> S'. f(g x) \<in> T} \<in> sets lebesgue"
 | |
| 2517 | proof (rule differentiable_image_in_sets_lebesgue) | |
| 2518 | let ?h = "\<lambda>x. \<bar>?D x\<bar> * f (g x) /\<^sub>R \<bar>?D x\<bar>" | |
| 2519 | have "(\<lambda>x. if x \<in> S' then \<bar>?D x\<bar> * f (g x) else 0) = (\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0)" | |
| 2520 | by (auto simp: S'_def) | |
| 2521 | also have "\<dots> \<in> borel_measurable lebesgue" | |
| 2522 | by (rule Df_borel) | |
| 2523 | finally have *: "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) \<in> borel_measurable (lebesgue_on S')" | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 2524 | by (simp add: borel_measurable_if_D) | 
| 67998 | 2525 | have "?h \<in> borel_measurable (lebesgue_on S')" | 
| 2526 | by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') | |
| 2527 | moreover have "?h x = f(g x)" if "x \<in> S'" for x | |
| 2528 | using that by (auto simp: S'_def) | |
| 2529 | ultimately have "(\<lambda>x. f(g x)) \<in> borel_measurable (lebesgue_on S')" | |
| 2530 | by (metis (no_types, lifting) measurable_lebesgue_cong) | |
| 2531 |       then show "{x \<in> S'. f (g x) \<in> T} \<in> sets lebesgue"
 | |
| 2532 | by (simp add: \<open>S' \<in> sets lebesgue\<close> \<open>open T\<close> borel_measurable_vimage_open sets_restrict_space_iff) | |
| 2533 |       show "g differentiable_on {x \<in> S'. f (g x) \<in> T}"
 | |
| 2534 | using der_g unfolding S'_def differentiable_def differentiable_on_def | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2535 | by (blast intro: has_derivative_subset) | 
| 67998 | 2536 | qed auto | 
| 2537 |     ultimately have "{x \<in> g ` S'. f x \<in> T} \<in> sets lebesgue"
 | |
| 2538 | by metis | |
| 2539 |     then show "{x \<in> g ` S'. f x \<in> T} \<in> sets (lebesgue_on (g ` S'))"
 | |
| 2540 | by (simp add: \<open>g ` S' \<in> sets lebesgue\<close> sets_restrict_space_iff) | |
| 2541 | qed | |
| 2542 | have intS': "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) integrable_on S'" | |
| 2543 | using intS | |
| 2544 | by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) | |
| 2545 |   have lebS': "{x \<in> S'. g x \<in> T} \<in> sets lebesgue" if "T \<subseteq> g ` S'" "T \<in> sets lebesgue" for T
 | |
| 2546 | proof - | |
| 2547 | have "g \<in> borel_measurable (lebesgue_on S')" | |
| 2548 | using der_gS' has_derivative_continuous_on S' | |
| 2549 | by (blast intro: continuous_imp_measurable_on_sets_lebesgue) | |
| 2550 |     moreover have "{x \<in> S'. g x \<in> U} \<in> sets lebesgue" if "negligible U" "U \<subseteq> g ` S'" for U
 | |
| 2551 | proof (intro negligible_imp_sets negligible_differentiable_vimage that) | |
| 2552 | fix x | |
| 2553 | assume x: "x \<in> S'" | |
| 2554 | then have "linear (g' x)" | |
| 2555 | using der_gS' has_derivative_linear by blast | |
| 2556 | with x show "inj (g' x)" | |
| 2557 | by (auto simp: S'_def det_nz_iff_inj) | |
| 2558 | qed (use der_gS' in auto) | |
| 2559 | ultimately show ?thesis | |
| 2560 | using double_lebesgue_sets [OF S' gS' order_refl] that by blast | |
| 2561 | qed | |
| 2562 | have int_gS': "f integrable_on g ` S' \<and> integral (g ` S') f \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))" | |
| 2563 | using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast | |
| 2564 |   have "negligible (g ` {x \<in> S. det(matrix(g' x)) = 0})"
 | |
| 2565 | proof (rule baby_Sard, simp_all) | |
| 2566 | fix x | |
| 2567 | assume x: "x \<in> S \<and> det (matrix (g' x)) = 0" | |
| 2568 |     then show "(g has_derivative g' x) (at x within {x \<in> S. det (matrix (g' x)) = 0})"
 | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2569 | by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI) | 
| 67998 | 2570 |     then show "rank (matrix (g' x)) < CARD('n)"
 | 
| 2571 | using det_nz_iff_inj matrix_vector_mul_linear x | |
| 2572 | by (fastforce simp add: less_rank_noninjective) | |
| 2573 | qed | |
| 2574 |   then have negg: "negligible (g ` S - g ` {x \<in> S. ?D x \<noteq> 0})"
 | |
| 2575 | by (rule negligible_subset) (auto simp: S'_def) | |
| 2576 |   have null: "g ` {x \<in> S. ?D x \<noteq> 0} - g ` S = {}"
 | |
| 2577 | by (auto simp: S'_def) | |
| 2578 |   let ?F = "{x \<in> S. f (g x) \<noteq> 0}"
 | |
| 2579 |   have eq: "g ` S' = g ` ?F \<inter> g ` {x \<in> S. ?D x \<noteq> 0}"
 | |
| 2580 | by (auto simp: S'_def image_iff) | |
| 2581 | show ?thesis | |
| 2582 | proof | |
| 2583 |     have "((\<lambda>x. if x \<in> g ` ?F then f x else 0) integrable_on g ` {x \<in> S. ?D x \<noteq> 0})"
 | |
| 2584 | using int_gS' eq integrable_restrict_Int [where f=f] | |
| 2585 | by simp | |
| 2586 |     then have "f integrable_on g ` {x \<in> S. ?D x \<noteq> 0}"
 | |
| 2587 | by (auto simp: image_iff elim!: integrable_eq) | |
| 2588 | then show "f integrable_on g ` S" | |
| 2589 | apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) | |
| 2590 | using negg null by auto | |
| 2591 |     have "integral (g ` S) f = integral (g ` {x \<in> S. ?D x \<noteq> 0}) f"
 | |
| 2592 | using negg by (auto intro: negligible_subset integral_spike_set) | |
| 2593 |     also have "\<dots> = integral (g ` {x \<in> S. ?D x \<noteq> 0}) (\<lambda>x. if x \<in> g ` ?F then f x else 0)"
 | |
| 2594 | by (auto simp: image_iff intro!: integral_cong) | |
| 2595 | also have "\<dots> = integral (g ` S') f" | |
| 2596 | using eq integral_restrict_Int by simp | |
| 2597 | also have "\<dots> \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))" | |
| 2598 | by (metis int_gS') | |
| 2599 | also have "\<dots> \<le> ?b" | |
| 2600 | by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) | |
| 2601 | finally show "integral (g ` S) f \<le> ?b" . | |
| 2602 | qed | |
| 2603 | qed | |
| 2604 | ||
| 2605 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2606 | lemma absolutely_integrable_on_image_real: | 
| 67998 | 2607 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
 | 
| 2608 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2609 | and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S" | |
| 2610 | shows "f absolutely_integrable_on (g ` S)" | |
| 2611 | proof - | |
| 2612 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)" | |
| 2613 |   let ?N = "{x \<in> S. f (g x) < 0}" and ?P = "{x \<in> S. f (g x) > 0}"
 | |
| 2614 |   have eq: "{x. (if x \<in> S then ?D x else 0) > 0} = {x \<in> S. ?D x > 0}"
 | |
| 2615 |            "{x. (if x \<in> S then ?D x else 0) < 0} = {x \<in> S. ?D x < 0}"
 | |
| 2616 | by auto | |
| 2617 | have "?D integrable_on S" | |
| 2618 | using intS absolutely_integrable_on_def by blast | |
| 2619 | then have "(\<lambda>x. if x \<in> S then ?D x else 0) integrable_on UNIV" | |
| 2620 | by (simp add: integrable_restrict_UNIV) | |
| 2621 | then have D_borel: "(\<lambda>x. if x \<in> S then ?D x else 0) \<in> borel_measurable (lebesgue_on UNIV)" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2622 | using integrable_imp_measurable lebesgue_on_UNIV_eq by blast | 
| 67998 | 2623 |   then have Dlt: "{x \<in> S. ?D x < 0} \<in> sets lebesgue"
 | 
| 2624 | unfolding borel_measurable_vimage_halfspace_component_lt | |
| 2625 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2626 |   from D_borel have Dgt: "{x \<in> S. ?D x > 0} \<in> sets lebesgue"
 | |
| 2627 | unfolding borel_measurable_vimage_halfspace_component_gt | |
| 2628 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2629 | ||
| 2630 | have dfgbm: "?D \<in> borel_measurable (lebesgue_on S)" | |
| 2631 | using intS absolutely_integrable_on_def integrable_imp_measurable by blast | |
| 2632 | have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \<in> ?N" for x | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2633 | using der_g has_derivative_subset that by force | 
| 67998 | 2634 | have "(\<lambda>x. - f x) integrable_on g ` ?N \<and> | 
| 2635 | integral (g ` ?N) (\<lambda>x. - f x) \<le> integral ?N (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x))" | |
| 2636 | proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) | |
| 2637 |     have 1: "?D integrable_on {x \<in> S. ?D x < 0}"
 | |
| 2638 | using Dlt | |
| 2639 | by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) | |
| 2640 | have "uminus \<circ> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N" | |
| 2641 | by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) | |
| 2642 | then show "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N" | |
| 2643 | by (simp add: integrable_neg_iff o_def) | |
| 2644 | qed auto | |
| 2645 | then have "f integrable_on g ` ?N" | |
| 2646 | by (simp add: integrable_neg_iff) | |
| 2647 |   moreover have "g ` ?N = {y \<in> g ` S. f y < 0}"
 | |
| 2648 | by auto | |
| 2649 |   ultimately have "f integrable_on {y \<in> g ` S. f y < 0}"
 | |
| 2650 | by simp | |
| 2651 |   then have N: "f absolutely_integrable_on {y \<in> g ` S. f y < 0}"
 | |
| 2652 | by (rule absolutely_integrable_absolutely_integrable_ubound) auto | |
| 2653 | ||
| 2654 | have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \<in> ?P" for x | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2655 | using der_g has_derivative_subset that by force | 
| 67998 | 2656 | have "f integrable_on g ` ?P \<and> integral (g ` ?P) f \<le> integral ?P ?D" | 
| 2657 | proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) | |
| 2658 |     have "?D integrable_on {x \<in> S. 0 < ?D x}"
 | |
| 2659 | using Dgt | |
| 2660 | by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) | |
| 2661 | then show "?D integrable_on ?P" | |
| 2662 | apply (rule integrable_spike_set) | |
| 2663 | by (auto simp: zero_less_mult_iff empty_imp_negligible) | |
| 2664 | qed auto | |
| 2665 | then have "f integrable_on g ` ?P" | |
| 2666 | by metis | |
| 2667 |   moreover have "g ` ?P = {y \<in> g ` S. f y > 0}"
 | |
| 2668 | by auto | |
| 2669 |   ultimately have "f integrable_on {y \<in> g ` S. f y > 0}"
 | |
| 2670 | by simp | |
| 2671 |   then have P: "f absolutely_integrable_on {y \<in> g ` S. f y > 0}"
 | |
| 2672 | by (rule absolutely_integrable_absolutely_integrable_lbound) auto | |
| 2673 | have "(\<lambda>x. if x \<in> g ` S \<and> f x < 0 \<or> x \<in> g ` S \<and> 0 < f x then f x else 0) = (\<lambda>x. if x \<in> g ` S then f x else 0)" | |
| 2674 | by auto | |
| 2675 | then show ?thesis | |
| 2676 | using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] | |
| 2677 | by simp | |
| 2678 | qed | |
| 2679 | ||
| 2680 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2681 | proposition absolutely_integrable_on_image: | 
| 67998 | 2682 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 2683 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2684 | and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S" | |
| 2685 | shows "f absolutely_integrable_on (g ` S)" | |
| 2686 | apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) | |
| 70136 | 2687 | using absolutely_integrable_component [OF intS] by auto | 
| 67998 | 2688 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2689 | proposition integral_on_image_ubound: | 
| 67998 | 2690 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
 | 
| 2691 | assumes"\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)" | |
| 2692 | and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2693 | and "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S" | |
| 2694 | shows "integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))" | |
| 70136 | 2695 | using integral_on_image_ubound_nonneg [OF assms] by simp | 
| 67998 | 2696 | |
| 2697 | ||
| 69683 | 2698 | subsection\<open>Change-of-variables theorem\<close> | 
| 67998 | 2699 | |
| 2700 | text\<open>The classic change-of-variables theorem. We have two versions with quite general hypotheses, | |
| 2701 | the first that the transforming function has a continuous inverse, the second that the base set is | |
| 2702 | Lebesgue measurable.\<close> | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2703 | lemma cov_invertible_nonneg_le: | 
| 67998 | 2704 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
 | 
| 2705 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2706 | and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)" | |
| 2707 | and f0: "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y" | |
| 2708 | and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x" | |
| 2709 | and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y" | |
| 2710 | and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id" | |
| 2711 | shows "f integrable_on T \<and> (integral T f) \<le> b \<longleftrightarrow> | |
| 2712 | (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S \<and> | |
| 2713 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) \<le> b" | |
| 2714 | (is "?lhs = ?rhs") | |
| 2715 | proof - | |
| 2716 | have Teq: "T = g`S" and Seq: "S = h`T" | |
| 2717 | using hg gh image_iff by fastforce+ | |
| 2718 | have gS: "g differentiable_on S" | |
| 2719 | by (meson der_g differentiable_def differentiable_on_def) | |
| 2720 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)" | |
| 2721 | show ?thesis | |
| 2722 | proof | |
| 2723 | assume ?lhs | |
| 2724 | then have fT: "f integrable_on T" and intf: "integral T f \<le> b" | |
| 2725 | by blast+ | |
| 2726 | show ?rhs | |
| 2727 | proof | |
| 2728 | let ?fgh = "\<lambda>x. \<bar>det (matrix (h' x))\<bar> * (\<bar>det (matrix (g' (h x)))\<bar> * f (g (h x)))" | |
| 2729 | have ddf: "?fgh x = f x" | |
| 2730 | if "x \<in> T" for x | |
| 2731 | proof - | |
| 2732 | have "matrix (h' x) ** matrix (g' (h x)) = mat 1" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2733 | using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
68001diff
changeset | 2734 | by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ | 
| 67998 | 2735 | then have "\<bar>det (matrix (h' x))\<bar> * \<bar>det (matrix (g' (h x)))\<bar> = 1" | 
| 2736 | by (metis abs_1 abs_mult det_I det_mul) | |
| 2737 | then show ?thesis | |
| 2738 | by (simp add: gh that) | |
| 2739 | qed | |
| 2740 | have "?D integrable_on (h ` T)" | |
| 2741 | proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) | |
| 2742 | show "(\<lambda>x. ?fgh x) absolutely_integrable_on T" | |
| 2743 | proof (subst absolutely_integrable_on_iff_nonneg) | |
| 2744 | show "(\<lambda>x. ?fgh x) integrable_on T" | |
| 2745 | using ddf fT integrable_eq by force | |
| 2746 | qed (simp add: zero_le_mult_iff f0 gh) | |
| 2747 | qed (use der_h in auto) | |
| 2748 | with Seq show "(\<lambda>x. ?D x) integrable_on S" | |
| 2749 | by simp | |
| 2750 | have "integral S (\<lambda>x. ?D x) \<le> integral T (\<lambda>x. ?fgh x)" | |
| 2751 | unfolding Seq | |
| 2752 | proof (rule integral_on_image_ubound) | |
| 2753 | show "(\<lambda>x. ?fgh x) integrable_on T" | |
| 2754 | using ddf fT integrable_eq by force | |
| 2755 | qed (use f0 gh der_h in auto) | |
| 2756 | also have "\<dots> = integral T f" | |
| 2757 | by (force simp: ddf intro: integral_cong) | |
| 2758 | also have "\<dots> \<le> b" | |
| 2759 | by (rule intf) | |
| 2760 | finally show "integral S (\<lambda>x. ?D x) \<le> b" . | |
| 2761 | qed | |
| 2762 | next | |
| 2763 | assume R: ?rhs | |
| 2764 | then have "f integrable_on g ` S" | |
| 2765 | using der_g f0 hg integral_on_image_ubound_nonneg by blast | |
| 2766 | moreover have "integral (g ` S) f \<le> integral S (\<lambda>x. ?D x)" | |
| 2767 | by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) | |
| 2768 | ultimately show ?lhs | |
| 2769 | using R by (simp add: Teq) | |
| 2770 | qed | |
| 2771 | qed | |
| 2772 | ||
| 2773 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2774 | lemma cov_invertible_nonneg_eq: | 
| 67998 | 2775 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
 | 
| 2776 | assumes "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2777 | and "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)" | |
| 2778 | and "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y" | |
| 2779 | and "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x" | |
| 2780 | and "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y" | |
| 2781 | and "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id" | |
| 2782 | shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) has_integral b) S \<longleftrightarrow> (f has_integral b) T" | |
| 2783 | using cov_invertible_nonneg_le [OF assms] | |
| 2784 | by (simp add: has_integral_iff) (meson eq_iff) | |
| 2785 | ||
| 2786 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2787 | lemma cov_invertible_real: | 
| 67998 | 2788 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
 | 
| 2789 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2790 | and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)" | |
| 2791 | and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x" | |
| 2792 | and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y" | |
| 2793 | and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id" | |
| 2794 | shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S \<and> | |
| 2795 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) = b \<longleftrightarrow> | |
| 2796 | f absolutely_integrable_on T \<and> integral T f = b" | |
| 2797 | (is "?lhs = ?rhs") | |
| 2798 | proof - | |
| 2799 | have Teq: "T = g`S" and Seq: "S = h`T" | |
| 2800 | using hg gh image_iff by fastforce+ | |
| 2801 | let ?DP = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)" and ?DN = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * -f(g x)" | |
| 2802 |   have "+": "(?DP has_integral b) {x \<in> S. f (g x) > 0} \<longleftrightarrow> (f has_integral b) {y \<in> T. f y > 0}" for b
 | |
| 2803 | proof (rule cov_invertible_nonneg_eq) | |
| 2804 |     have *: "(\<lambda>x. f (g x)) -` Y \<inter> {x \<in> S. f (g x) > 0}
 | |
| 2805 |           = ((\<lambda>x. f (g x)) -` Y \<inter> S) \<inter> {x \<in> S. f (g x) > 0}" for Y
 | |
| 2806 | by auto | |
| 2807 |     show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) > 0})" if "x \<in> {x \<in> S. f (g x) > 0}" for x
 | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2808 | using that der_g has_derivative_subset by fastforce | 
| 67998 | 2809 |     show "(h has_derivative h' y) (at y within {y \<in> T. f y > 0})" if "y \<in> {y \<in> T. f y > 0}" for y
 | 
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2810 | using that der_h has_derivative_subset by fastforce | 
| 67998 | 2811 | qed (use gh hg id in auto) | 
| 2812 |   have "-": "(?DN has_integral b) {x \<in> S. f (g x) < 0} \<longleftrightarrow> ((\<lambda>x. - f x) has_integral b) {y \<in> T. f y < 0}" for b
 | |
| 2813 | proof (rule cov_invertible_nonneg_eq) | |
| 2814 |     have *: "(\<lambda>x. - f (g x)) -` y \<inter> {x \<in> S. f (g x) < 0}
 | |
| 2815 |           = ((\<lambda>x. f (g x)) -` uminus ` y \<inter> S) \<inter> {x \<in> S. f (g x) < 0}" for y
 | |
| 2816 | using image_iff by fastforce | |
| 2817 |     show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) < 0})" if "x \<in> {x \<in> S. f (g x) < 0}" for x
 | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2818 | using that der_g has_derivative_subset by fastforce | 
| 67998 | 2819 |     show "(h has_derivative h' y) (at y within {y \<in> T. f y < 0})" if "y \<in> {y \<in> T. f y < 0}" for y
 | 
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2820 | using that der_h has_derivative_subset by fastforce | 
| 67998 | 2821 | qed (use gh hg id in auto) | 
| 2822 | show ?thesis | |
| 2823 | proof | |
| 2824 | assume LHS: ?lhs | |
| 2825 |     have eq: "{x. (if x \<in> S then ?DP x else 0) > 0} = {x \<in> S. ?DP x > 0}"
 | |
| 2826 |       "{x. (if x \<in> S then ?DP x else 0) < 0} = {x \<in> S. ?DP x < 0}"
 | |
| 2827 | by auto | |
| 2828 | have "?DP integrable_on S" | |
| 2829 | using LHS absolutely_integrable_on_def by blast | |
| 2830 | then have "(\<lambda>x. if x \<in> S then ?DP x else 0) integrable_on UNIV" | |
| 2831 | by (simp add: integrable_restrict_UNIV) | |
| 2832 | then have D_borel: "(\<lambda>x. if x \<in> S then ?DP x else 0) \<in> borel_measurable (lebesgue_on UNIV)" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2833 | using integrable_imp_measurable lebesgue_on_UNIV_eq by blast | 
| 67998 | 2834 |     then have SN: "{x \<in> S. ?DP x < 0} \<in> sets lebesgue"
 | 
| 2835 | unfolding borel_measurable_vimage_halfspace_component_lt | |
| 2836 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2837 |     from D_borel have SP: "{x \<in> S. ?DP x > 0} \<in> sets lebesgue"
 | |
| 2838 | unfolding borel_measurable_vimage_halfspace_component_gt | |
| 2839 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2840 |     have "?DP absolutely_integrable_on {x \<in> S. ?DP x > 0}"
 | |
| 2841 | using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) | |
| 2842 |     then have aP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
 | |
| 2843 | by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) | |
| 2844 |     have "?DP absolutely_integrable_on {x \<in> S. ?DP x < 0}"
 | |
| 2845 | using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) | |
| 2846 |     then have aN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
 | |
| 2847 | by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) | |
| 2848 |     have fN: "f integrable_on {y \<in> T. f y < 0}"
 | |
| 2849 |              "integral {y \<in> T. f y < 0} f = integral {x \<in> S. f (g x) < 0} ?DP"
 | |
| 2850 |       using "-" [of "integral {x \<in> S. f(g x) < 0} ?DN"] aN
 | |
| 2851 | by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) | |
| 2852 |     have faN: "f absolutely_integrable_on {y \<in> T. f y < 0}"
 | |
| 2853 | apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. - f x"]) | |
| 2854 | using fN by (auto simp: integrable_neg_iff) | |
| 2855 |     have fP: "f integrable_on {y \<in> T. f y > 0}"
 | |
| 2856 |              "integral {y \<in> T. f y > 0} f = integral {x \<in> S. f (g x) > 0} ?DP"
 | |
| 2857 |       using "+" [of "integral {x \<in> S. f(g x) > 0} ?DP"] aP
 | |
| 2858 | by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) | |
| 2859 |     have faP: "f absolutely_integrable_on {y \<in> T. f y > 0}"
 | |
| 2860 | apply (rule absolutely_integrable_integrable_bound [where g = f]) | |
| 2861 | using fP by auto | |
| 2862 |     have fa: "f absolutely_integrable_on ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0})"
 | |
| 2863 | by (rule absolutely_integrable_Un [OF faN faP]) | |
| 2864 | show ?rhs | |
| 2865 | proof | |
| 2866 | have eq: "((if x \<in> T \<and> f x < 0 \<or> x \<in> T \<and> 0 < f x then 1 else 0) * f x) | |
| 2867 | = (if x \<in> T then 1 else 0) * f x" for x | |
| 2868 | by auto | |
| 2869 | show "f absolutely_integrable_on T" | |
| 73536 | 2870 | using fa by (simp add: indicator_def of_bool_def set_integrable_def eq) | 
| 67998 | 2871 |       have [simp]: "{y \<in> T. f y < 0} \<inter> {y \<in> T. 0 < f y} = {}" for T and f :: "(real^'n::_) \<Rightarrow> real"
 | 
| 2872 | by auto | |
| 2873 |       have "integral T f = integral ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0}) f"
 | |
| 2874 | by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) | |
| 2875 |       also have "\<dots> = integral {y \<in> T. f y < 0} f + integral {y \<in> T. f y > 0} f"
 | |
| 2876 | using fN fP by simp | |
| 2877 |       also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
 | |
| 2878 | by (simp add: fN fP) | |
| 2879 |       also have "\<dots> = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. 0 < f (g x)}) ?DP"
 | |
| 2880 | using aP aN by (simp add: set_lebesgue_integral_eq_integral) | |
| 2881 | also have "\<dots> = integral S ?DP" | |
| 2882 | by (intro empty_imp_negligible integral_spike_set) auto | |
| 2883 | also have "\<dots> = b" | |
| 2884 | using LHS by simp | |
| 2885 | finally show "integral T f = b" . | |
| 2886 | qed | |
| 2887 | next | |
| 2888 | assume RHS: ?rhs | |
| 2889 |     have eq: "{x. (if x \<in> T then f x else 0) > 0} = {x \<in> T. f x > 0}"
 | |
| 2890 |              "{x. (if x \<in> T then f x else 0) < 0} = {x \<in> T. f x < 0}"
 | |
| 2891 | by auto | |
| 2892 | have "f integrable_on T" | |
| 2893 | using RHS absolutely_integrable_on_def by blast | |
| 2894 | then have "(\<lambda>x. if x \<in> T then f x else 0) integrable_on UNIV" | |
| 2895 | by (simp add: integrable_restrict_UNIV) | |
| 2896 | then have D_borel: "(\<lambda>x. if x \<in> T then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2897 | using integrable_imp_measurable lebesgue_on_UNIV_eq by blast | 
| 67998 | 2898 |     then have TN: "{x \<in> T. f x < 0} \<in> sets lebesgue"
 | 
| 2899 | unfolding borel_measurable_vimage_halfspace_component_lt | |
| 2900 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2901 |     from D_borel have TP: "{x \<in> T. f x > 0} \<in> sets lebesgue"
 | |
| 2902 | unfolding borel_measurable_vimage_halfspace_component_gt | |
| 2903 | by (drule_tac x=0 in spec) (auto simp: eq) | |
| 2904 |     have aint: "f absolutely_integrable_on {y. y \<in> T \<and> 0 < (f y)}"
 | |
| 2905 |                "f absolutely_integrable_on {y. y \<in> T \<and> (f y) < 0}"
 | |
| 2906 | and intT: "integral T f = b" | |
| 2907 | using set_integrable_subset [of _ T] TP TN RHS | |
| 2908 | by blast+ | |
| 2909 | show ?lhs | |
| 2910 | proof | |
| 2911 |       have fN: "f integrable_on {v \<in> T. f v < 0}"
 | |
| 2912 | using absolutely_integrable_on_def aint by blast | |
| 2913 |       then have DN: "(?DN has_integral integral {y \<in> T. f y < 0} (\<lambda>x. - f x)) {x \<in> S. f (g x) < 0}"
 | |
| 2914 |         using "-" [of "integral {y \<in> T. f y < 0} (\<lambda>x. - f x)"]
 | |
| 2915 | by (simp add: has_integral_neg_iff integrable_integral) | |
| 2916 |       have aDN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
 | |
| 2917 | apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) | |
| 2918 | using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ | |
| 2919 |       have fP: "f integrable_on {v \<in> T. f v > 0}"
 | |
| 2920 | using absolutely_integrable_on_def aint by blast | |
| 2921 |       then have DP: "(?DP has_integral integral {y \<in> T. f y > 0} f) {x \<in> S. f (g x) > 0}"
 | |
| 2922 |         using "+" [of "integral {y \<in> T. f y > 0} f"]
 | |
| 2923 | by (simp add: has_integral_neg_iff integrable_integral) | |
| 2924 |       have aDP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
 | |
| 2925 | apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) | |
| 2926 | using DP hg by (fastforce simp: integrable_neg_iff)+ | |
| 2927 | have eq: "(if x \<in> S then 1 else 0) * ?DP x = (if x \<in> S \<and> f (g x) < 0 \<or> x \<in> S \<and> f (g x) > 0 then 1 else 0) * ?DP x" for x | |
| 2928 | by force | |
| 2929 |       have "?DP absolutely_integrable_on ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0})"
 | |
| 2930 | by (rule absolutely_integrable_Un [OF aDN aDP]) | |
| 2931 | then show I: "?DP absolutely_integrable_on S" | |
| 73536 | 2932 | by (simp add: indicator_def of_bool_def eq set_integrable_def) | 
| 67998 | 2933 |       have [simp]: "{y \<in> S. f y < 0} \<inter> {y \<in> S. 0 < f y} = {}" for S and f :: "(real^'n::_) \<Rightarrow> real"
 | 
| 2934 | by auto | |
| 2935 |       have "integral S ?DP = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0}) ?DP"
 | |
| 2936 | by (intro empty_imp_negligible integral_spike_set) auto | |
| 2937 |       also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
 | |
| 2938 | using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) | |
| 2939 |       also have "\<dots> = - integral {y \<in> T. f y < 0} (\<lambda>x. - f x) + integral {y \<in> T. f y > 0} f"
 | |
| 2940 | using DN DP by (auto simp: has_integral_iff) | |
| 2941 |       also have "\<dots> = integral ({x \<in> T. f x < 0} \<union> {x \<in> T. 0 < f x}) f"
 | |
| 2942 | by (simp add: fN fP) | |
| 2943 | also have "\<dots> = integral T f" | |
| 2944 | by (intro empty_imp_negligible integral_spike_set) auto | |
| 2945 | also have "\<dots> = b" | |
| 2946 | using intT by simp | |
| 2947 | finally show "integral S ?DP = b" . | |
| 2948 | qed | |
| 2949 | qed | |
| 2950 | qed | |
| 2951 | ||
| 2952 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2953 | lemma cv_inv_version3: | 
| 67998 | 2954 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 2955 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 2956 | and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)" | |
| 2957 | and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x" | |
| 2958 | and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y" | |
| 2959 | and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id" | |
| 2960 | shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 2961 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b | |
| 2962 | \<longleftrightarrow> f absolutely_integrable_on T \<and> integral T f = b" | |
| 2963 | proof - | |
| 2964 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)" | |
| 2965 | have "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x) $ i) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * (f(g x) $ i)) = b $ i) \<longleftrightarrow> | |
| 2966 | ((\<lambda>x. f x $ i) absolutely_integrable_on T \<and> integral T (\<lambda>x. f x $ i) = b $ i)" for i | |
| 2967 | by (rule cov_invertible_real [OF der_g der_h hg gh id]) | |
| 2968 | then have "?D absolutely_integrable_on S \<and> (?D has_integral b) S \<longleftrightarrow> | |
| 2969 | f absolutely_integrable_on T \<and> (f has_integral b) T" | |
| 2970 | unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] | |
| 2971 | absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] | |
| 2972 | by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] | |
| 2973 | has_integral_iff set_lebesgue_integral_eq_integral) | |
| 2974 | then show ?thesis | |
| 2975 | using absolutely_integrable_on_def by blast | |
| 2976 | qed | |
| 2977 | ||
| 2978 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 2979 | lemma cv_inv_version4: | 
| 67998 | 2980 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 2981 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> invertible(matrix(g' x))" | |
| 2982 | and hg: "\<And>x. x \<in> S \<Longrightarrow> continuous_on (g ` S) h \<and> h(g x) = x" | |
| 2983 | shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 2984 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b | |
| 2985 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 2986 | proof - | |
| 2987 | have "\<forall>x. \<exists>h'. x \<in> S | |
| 2988 | \<longrightarrow> (g has_derivative g' x) (at x within S) \<and> linear h' \<and> g' x \<circ> h' = id \<and> h' \<circ> g' x = id" | |
| 2989 | using der_g matrix_invertible has_derivative_linear by blast | |
| 2990 | then obtain h' where h': | |
| 2991 | "\<And>x. x \<in> S | |
| 2992 | \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> | |
| 2993 | linear (h' x) \<and> g' x \<circ> (h' x) = id \<and> (h' x) \<circ> g' x = id" | |
| 2994 | by metis | |
| 2995 | show ?thesis | |
| 2996 | proof (rule cv_inv_version3) | |
| 2997 | show "\<And>y. y \<in> g ` S \<Longrightarrow> (h has_derivative h' (h y)) (at y within g ` S)" | |
| 2998 | using h' hg | |
| 2999 | by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) | |
| 3000 | qed (use h' hg in auto) | |
| 3001 | qed | |
| 3002 | ||
| 3003 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3004 | theorem has_absolute_integral_change_of_variables_invertible: | 
| 67998 | 3005 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3006 | assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 3007 | and hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x" | |
| 3008 | and conth: "continuous_on (g ` S) h" | |
| 3009 | shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b \<longleftrightarrow> | |
| 3010 | f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 3011 | (is "?lhs = ?rhs") | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3012 | proof - | 
| 67998 | 3013 |   let ?S = "{x \<in> S. invertible (matrix (g' x))}" and ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
 | 
| 3014 | have *: "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b | |
| 3015 | \<longleftrightarrow> f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b" | |
| 3016 | proof (rule cv_inv_version4) | |
| 3017 | show "(g has_derivative g' x) (at x within ?S) \<and> invertible (matrix (g' x))" | |
| 3018 | if "x \<in> ?S" for x | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3019 | using der_g that has_derivative_subset that by fastforce | 
| 67998 | 3020 | show "continuous_on (g ` ?S) h \<and> h (g x) = x" | 
| 3021 | if "x \<in> ?S" for x | |
| 3022 | using that continuous_on_subset [OF conth] by (simp add: hg image_mono) | |
| 3023 | qed | |
| 3024 |   have "(g has_derivative g' x) (at x within {x \<in> S. rank (matrix (g' x)) < CARD('m)})" if "x \<in> S" for x
 | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3025 | by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI that) | 
| 67998 | 3026 |   then have "negligible (g ` {x \<in> S. \<not> invertible (matrix (g' x))})"
 | 
| 3027 | by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) | |
| 3028 |   then have neg: "negligible {x \<in> g ` S. x \<notin> g ` ?S \<and> f x \<noteq> 0}"
 | |
| 3029 | by (auto intro: negligible_subset) | |
| 3030 |   have [simp]: "{x \<in> g ` ?S. x \<notin> g ` S \<and> f x \<noteq> 0} = {}"
 | |
| 3031 | by auto | |
| 3032 | have "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b | |
| 3033 | \<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b" | |
| 3034 | apply (intro conj_cong absolutely_integrable_spike_set_eq) | |
| 3035 | apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) | |
| 3036 | done | |
| 3037 | moreover | |
| 3038 | have "f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b | |
| 3039 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 3040 | by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) | |
| 3041 | ultimately | |
| 3042 | show ?thesis | |
| 3043 | using "*" by blast | |
| 3044 | qed | |
| 3045 | ||
| 3046 | ||
| 3047 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3048 | theorem has_absolute_integral_change_of_variables_compact: | 
| 67998 | 3049 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3050 | assumes "compact S" | |
| 3051 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 3052 | and inj: "inj_on g S" | |
| 3053 | shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 3054 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b | |
| 3055 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b)" | |
| 3056 | proof - | |
| 3057 | obtain h where hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x" | |
| 3058 | using inj by (metis the_inv_into_f_f) | |
| 3059 | have conth: "continuous_on (g ` S) h" | |
| 3060 | by (metis \<open>compact S\<close> continuous_on_inv der_g has_derivative_continuous_on hg) | |
| 3061 | show ?thesis | |
| 3062 | by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) | |
| 3063 | qed | |
| 3064 | ||
| 3065 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3066 | lemma has_absolute_integral_change_of_variables_compact_family: | 
| 67998 | 3067 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3068 | assumes compact: "\<And>n::nat. compact (F n)" | |
| 3069 | and der_g: "\<And>x. x \<in> (\<Union>n. F n) \<Longrightarrow> (g has_derivative g' x) (at x within (\<Union>n. F n))" | |
| 3070 | and inj: "inj_on g (\<Union>n. F n)" | |
| 3071 | shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on (\<Union>n. F n) \<and> | |
| 3072 | integral (\<Union>n. F n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b | |
| 3073 | \<longleftrightarrow> f absolutely_integrable_on (g ` (\<Union>n. F n)) \<and> integral (g ` (\<Union>n. F n)) f = b)" | |
| 3074 | proof - | |
| 3075 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)" | |
| 3076 | let ?U = "\<lambda>n. \<Union>m\<le>n. F m" | |
| 3077 | let ?lift = "vec::real\<Rightarrow>real^1" | |
| 3078 | have F_leb: "F m \<in> sets lebesgue" for m | |
| 3079 | by (simp add: compact borel_compact) | |
| 3080 | have iff: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \<and> | |
| 3081 | integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) = b | |
| 3082 | \<longleftrightarrow> f absolutely_integrable_on (g ` (?U n)) \<and> integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \<Rightarrow> real^'k" | |
| 3083 | proof (rule has_absolute_integral_change_of_variables_compact) | |
| 3084 | show "compact (?U n)" | |
| 3085 | by (simp add: compact compact_UN) | |
| 3086 | show "(g has_derivative g' x) (at x within (?U n))" | |
| 3087 | if "x \<in> ?U n" for x | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3088 | using that by (blast intro!: has_derivative_subset [OF der_g]) | 
| 67998 | 3089 | show "inj_on g (?U n)" | 
| 3090 | using inj by (auto simp: inj_on_def) | |
| 3091 | qed | |
| 3092 | show ?thesis | |
| 3093 | unfolding image_UN | |
| 3094 | proof safe | |
| 3095 | assume DS: "?D absolutely_integrable_on (\<Union>n. F n)" | |
| 3096 | and b: "b = integral (\<Union>n. F n) ?D" | |
| 3097 | have DU: "\<And>n. ?D absolutely_integrable_on (?U n)" | |
| 3098 | "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D" | |
| 3099 | using integral_countable_UN [OF DS F_leb] by auto | |
| 3100 | with iff have fag: "f absolutely_integrable_on g ` (?U n)" | |
| 3101 | and fg_int: "integral (\<Union>m\<le>n. g ` F m) f = integral (?U n) ?D" for n | |
| 3102 | by (auto simp: image_UN) | |
| 3103 | let ?h = "\<lambda>x. if x \<in> (\<Union>m. g ` F m) then norm(f x) else 0" | |
| 3104 | have "(\<lambda>x. if x \<in> (\<Union>m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" | |
| 3105 | proof (rule dominated_convergence_absolutely_integrable) | |
| 3106 | show "(\<lambda>x. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k | |
| 3107 | unfolding absolutely_integrable_restrict_UNIV | |
| 3108 | using fag by (simp add: image_UN) | |
| 3109 | let ?nf = "\<lambda>n x. if x \<in> (\<Union>m\<le>n. g ` F m) then norm(f x) else 0" | |
| 3110 | show "?h integrable_on UNIV" | |
| 3111 | proof (rule monotone_convergence_increasing [THEN conjunct1]) | |
| 3112 | show "?nf k integrable_on UNIV" for k | |
| 3113 | using fag | |
| 3114 | unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) | |
| 3115 |         { fix n
 | |
| 3116 | have "(norm \<circ> ?D) absolutely_integrable_on ?U n" | |
| 3117 | by (intro absolutely_integrable_norm DU) | |
| 3118 | then have "integral (g ` ?U n) (norm \<circ> f) = integral (?U n) (norm \<circ> ?D)" | |
| 3119 | using iff [of n "vec \<circ> norm \<circ> f" "integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R (?lift \<circ> norm \<circ> f) (g x))"] | |
| 3120 | unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) | |
| 3121 | } | |
| 3122 | moreover have "bounded (range (\<lambda>k. integral (?U k) (norm \<circ> ?D)))" | |
| 3123 | unfolding bounded_iff | |
| 3124 | proof (rule exI, clarify) | |
| 3125 | fix k | |
| 3126 | show "norm (integral (?U k) (norm \<circ> ?D)) \<le> integral (\<Union>n. F n) (norm \<circ> ?D)" | |
| 3127 | unfolding integral_restrict_UNIV [of _ "norm \<circ> ?D", symmetric] | |
| 3128 | proof (rule integral_norm_bound_integral) | |
| 69325 | 3129 |             show "(\<lambda>x. if x \<in> \<Union> (F ` {..k}) then (norm \<circ> ?D) x else 0) integrable_on UNIV"
 | 
| 67998 | 3130 | "(\<lambda>x. if x \<in> (\<Union>n. F n) then (norm \<circ> ?D) x else 0) integrable_on UNIV" | 
| 3131 | using DU(1) DS | |
| 3132 | unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto | |
| 3133 | qed auto | |
| 3134 | qed | |
| 3135 | ultimately show "bounded (range (\<lambda>k. integral UNIV (?nf k)))" | |
| 3136 | by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) | |
| 3137 | next | |
| 3138 | show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then norm (f x) else 0) | |
| 3139 | \<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then norm (f x) else 0)" for x | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70139diff
changeset | 3140 | by (force intro: tendsto_eventually eventually_sequentiallyI) | 
| 67998 | 3141 | qed auto | 
| 3142 | next | |
| 3143 | show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0) | |
| 3144 | \<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then f x else 0)" for x | |
| 3145 | proof clarsimp | |
| 3146 | fix m y | |
| 3147 | assume "y \<in> F m" | |
| 3148 |         show "(\<lambda>k. if \<exists>x\<in>{..k}. g y \<in> g ` F x then f (g y) else 0) \<longlonglongrightarrow> f (g y)"
 | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70139diff
changeset | 3149 | using \<open>y \<in> F m\<close> by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) | 
| 67998 | 3150 | qed | 
| 3151 | qed auto | |
| 3152 | then show fai: "f absolutely_integrable_on (\<Union>m. g ` F m)" | |
| 3153 | using absolutely_integrable_restrict_UNIV by blast | |
| 3154 | show "integral ((\<Union>x. g ` F x)) f = integral (\<Union>n. F n) ?D" | |
| 3155 | proof (rule LIMSEQ_unique) | |
| 3156 | show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f" | |
| 3157 | unfolding fg_int [symmetric] | |
| 3158 | proof (rule integral_countable_UN [OF fai]) | |
| 3159 | show "g ` F m \<in> sets lebesgue" for m | |
| 3160 | proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) | |
| 3161 | show "g differentiable_on F m" | |
| 3162 | by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) | |
| 3163 | qed auto | |
| 3164 | qed | |
| 3165 | next | |
| 3166 | show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D" | |
| 3167 | by (rule DU) | |
| 3168 | qed | |
| 3169 | next | |
| 3170 | assume fs: "f absolutely_integrable_on (\<Union>x. g ` F x)" | |
| 3171 | and b: "b = integral ((\<Union>x. g ` F x)) f" | |
| 3172 | have gF_leb: "g ` F m \<in> sets lebesgue" for m | |
| 3173 | proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) | |
| 3174 | show "g differentiable_on F m" | |
| 3175 | using der_g unfolding differentiable_def differentiable_on_def | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3176 | by (meson Sup_upper UNIV_I UnionI has_derivative_subset image_eqI) | 
| 67998 | 3177 | qed auto | 
| 3178 | have fgU: "\<And>n. f absolutely_integrable_on (\<Union>m\<le>n. g ` F m)" | |
| 3179 | "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>m. g ` F m) f" | |
| 3180 | using integral_countable_UN [OF fs gF_leb] by auto | |
| 3181 | with iff have DUn: "?D absolutely_integrable_on ?U n" | |
| 3182 | and D_int: "integral (?U n) ?D = integral (\<Union>m\<le>n. g ` F m) f" for n | |
| 3183 | by (auto simp: image_UN) | |
| 3184 | let ?h = "\<lambda>x. if x \<in> (\<Union>n. F n) then norm(?D x) else 0" | |
| 3185 | have "(\<lambda>x. if x \<in> (\<Union>n. F n) then ?D x else 0) absolutely_integrable_on UNIV" | |
| 3186 | proof (rule dominated_convergence_absolutely_integrable) | |
| 3187 | show "(\<lambda>x. if x \<in> ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k | |
| 3188 | unfolding absolutely_integrable_restrict_UNIV using DUn by simp | |
| 3189 | let ?nD = "\<lambda>n x. if x \<in> ?U n then norm(?D x) else 0" | |
| 3190 | show "?h integrable_on UNIV" | |
| 3191 | proof (rule monotone_convergence_increasing [THEN conjunct1]) | |
| 3192 | show "?nD k integrable_on UNIV" for k | |
| 3193 | using DUn | |
| 3194 | unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) | |
| 3195 |         { fix n::nat
 | |
| 3196 | have "(norm \<circ> f) absolutely_integrable_on (\<Union>m\<le>n. g ` F m)" | |
| 3197 | apply (rule absolutely_integrable_norm) | |
| 3198 | using fgU by blast | |
| 3199 | then have "integral (?U n) (norm \<circ> ?D) = integral (g ` ?U n) (norm \<circ> f)" | |
| 3200 | using iff [of n "?lift \<circ> norm \<circ> f" "integral (g ` ?U n) (?lift \<circ> norm \<circ> f)"] | |
| 3201 | unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) | |
| 3202 | } | |
| 3203 | moreover have "bounded (range (\<lambda>k. integral (g ` ?U k) (norm \<circ> f)))" | |
| 3204 | unfolding bounded_iff | |
| 3205 | proof (rule exI, clarify) | |
| 3206 | fix k | |
| 3207 | show "norm (integral (g ` ?U k) (norm \<circ> f)) \<le> integral (g ` (\<Union>n. F n)) (norm \<circ> f)" | |
| 3208 | unfolding integral_restrict_UNIV [of _ "norm \<circ> f", symmetric] | |
| 3209 | proof (rule integral_norm_bound_integral) | |
| 3210 | show "(\<lambda>x. if x \<in> g ` ?U k then (norm \<circ> f) x else 0) integrable_on UNIV" | |
| 3211 | "(\<lambda>x. if x \<in> g ` (\<Union>n. F n) then (norm \<circ> f) x else 0) integrable_on UNIV" | |
| 3212 | using fgU fs | |
| 3213 | unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV | |
| 3214 | by (auto simp: image_UN) | |
| 3215 | qed auto | |
| 3216 | qed | |
| 3217 | ultimately show "bounded (range (\<lambda>k. integral UNIV (?nD k)))" | |
| 3218 | unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp | |
| 3219 | next | |
| 3220 | show "(\<lambda>k. if x \<in> ?U k then norm (?D x) else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then norm (?D x) else 0)" for x | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70139diff
changeset | 3221 | by (force intro: tendsto_eventually eventually_sequentiallyI) | 
| 67998 | 3222 | qed auto | 
| 3223 | next | |
| 3224 | show "(\<lambda>k. if x \<in> ?U k then ?D x else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then ?D x else 0)" for x | |
| 3225 | proof clarsimp | |
| 3226 | fix n | |
| 3227 | assume "x \<in> F n" | |
| 3228 |         show "(\<lambda>m. if \<exists>j\<in>{..m}. x \<in> F j then ?D x else 0) \<longlonglongrightarrow> ?D x"
 | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70139diff
changeset | 3229 | using \<open>x \<in> F n\<close> by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) | 
| 67998 | 3230 | qed | 
| 3231 | qed auto | |
| 3232 | then show Dai: "?D absolutely_integrable_on (\<Union>n. F n)" | |
| 3233 | unfolding absolutely_integrable_restrict_UNIV by simp | |
| 3234 | show "integral (\<Union>n. F n) ?D = integral ((\<Union>x. g ` F x)) f" | |
| 3235 | proof (rule LIMSEQ_unique) | |
| 3236 | show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f" | |
| 3237 | by (rule fgU) | |
| 3238 | show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D" | |
| 3239 | unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) | |
| 3240 | qed | |
| 3241 | qed | |
| 3242 | qed | |
| 3243 | ||
| 3244 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3245 | theorem has_absolute_integral_change_of_variables: | 
| 67998 | 3246 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3247 | assumes S: "S \<in> sets lebesgue" | |
| 3248 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 3249 | and inj: "inj_on g S" | |
| 3250 | shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 3251 | integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b | |
| 3252 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3253 | proof - | 
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 3254 | obtain C N where "fsigma C" and N: "N \<in> null_sets lebesgue" and CNS: "C \<union> N = S" and "disjnt C N" | 
| 67998 | 3255 | using lebesgue_set_almost_fsigma [OF S] . | 
| 3256 | then obtain F :: "nat \<Rightarrow> (real^'m::_) set" | |
| 3257 | where F: "range F \<subseteq> Collect compact" and Ceq: "C = Union(range F)" | |
| 3258 | using fsigma_Union_compact by metis | |
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 3259 | have "negligible N" | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 3260 | using N by (simp add: negligible_iff_null_sets) | 
| 67998 | 3261 | let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)" | 
| 3262 | have "?D absolutely_integrable_on C \<and> integral C ?D = b | |
| 3263 | \<longleftrightarrow> f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b" | |
| 3264 | unfolding Ceq | |
| 3265 | proof (rule has_absolute_integral_change_of_variables_compact_family) | |
| 3266 | fix n x | |
| 69313 | 3267 | assume "x \<in> \<Union>(F ` UNIV)" | 
| 3268 | then show "(g has_derivative g' x) (at x within \<Union>(F ` UNIV))" | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3269 | using Ceq \<open>C \<union> N = S\<close> der_g has_derivative_subset by blast | 
| 67998 | 3270 | next | 
| 69313 | 3271 | have "\<Union>(F ` UNIV) \<subseteq> S" | 
| 67998 | 3272 | using Ceq \<open>C \<union> N = S\<close> by blast | 
| 69313 | 3273 | then show "inj_on g (\<Union>(F ` UNIV))" | 
| 67998 | 3274 | using inj by (meson inj_on_subset) | 
| 3275 | qed (use F in auto) | |
| 3276 | moreover | |
| 3277 | have "?D absolutely_integrable_on C \<and> integral C ?D = b | |
| 3278 | \<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b" | |
| 3279 | proof (rule conj_cong) | |
| 3280 |     have neg: "negligible {x \<in> C - S. ?D x \<noteq> 0}" "negligible {x \<in> S - C. ?D x \<noteq> 0}"
 | |
| 3281 | using CNS by (blast intro: negligible_subset [OF \<open>negligible N\<close>])+ | |
| 3282 | then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" | |
| 3283 | by (rule absolutely_integrable_spike_set_eq) | |
| 3284 | show "(integral C ?D = b) \<longleftrightarrow> (integral S ?D = b)" | |
| 3285 | using integral_spike_set [OF neg] by simp | |
| 3286 | qed | |
| 3287 | moreover | |
| 3288 | have "f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b | |
| 3289 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 3290 | proof (rule conj_cong) | |
| 3291 | have "g differentiable_on N" | |
| 3292 | by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) | |
| 3293 | with \<open>negligible N\<close> | |
| 3294 | have neg_gN: "negligible (g ` N)" | |
| 3295 | by (blast intro: negligible_differentiable_image_negligible) | |
| 3296 |     have neg: "negligible {x \<in> g ` C - g ` S. f x \<noteq> 0}"
 | |
| 3297 |               "negligible {x \<in> g ` S - g ` C. f x \<noteq> 0}"
 | |
| 3298 | using CNS by (blast intro: negligible_subset [OF neg_gN])+ | |
| 3299 | then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" | |
| 3300 | by (rule absolutely_integrable_spike_set_eq) | |
| 3301 | show "(integral (g ` C) f = b) \<longleftrightarrow> (integral (g ` S) f = b)" | |
| 3302 | using integral_spike_set [OF neg] by simp | |
| 3303 | qed | |
| 3304 | ultimately show ?thesis | |
| 3305 | by simp | |
| 3306 | qed | |
| 3307 | ||
| 3308 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3309 | corollary absolutely_integrable_change_of_variables: | 
| 67998 | 3310 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3311 | assumes "S \<in> sets lebesgue" | |
| 3312 | and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 3313 | and "inj_on g S" | |
| 3314 | shows "f absolutely_integrable_on (g ` S) | |
| 3315 | \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S" | |
| 70136 | 3316 | using assms has_absolute_integral_change_of_variables by blast | 
| 67998 | 3317 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3318 | corollary integral_change_of_variables: | 
| 67998 | 3319 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3320 | assumes S: "S \<in> sets lebesgue" | |
| 3321 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)" | |
| 3322 | and inj: "inj_on g S" | |
| 3323 | and disj: "(f absolutely_integrable_on (g ` S) \<or> | |
| 3324 | (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)" | |
| 3325 | shows "integral (g ` S) f = integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x))" | |
| 70136 | 3326 | using has_absolute_integral_change_of_variables [OF S der_g inj] disj | 
| 3327 | by blast | |
| 67998 | 3328 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3329 | lemma has_absolute_integral_change_of_variables_1: | 
| 67998 | 3330 |   fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
 | 
| 3331 | assumes S: "S \<in> sets lebesgue" | |
| 3332 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)" | |
| 3333 | and inj: "inj_on g S" | |
| 3334 | shows "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 3335 | integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) = b | |
| 3336 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 3337 | proof - | |
| 3338 | let ?lift = "vec :: real \<Rightarrow> real^1" | |
| 3339 | let ?drop = "(\<lambda>x::real^1. x $ 1)" | |
| 3340 | have S': "?lift ` S \<in> sets lebesgue" | |
| 3341 | by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 3342 | have "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" | 
| 67998 | 3343 | if "z \<in> S" for z | 
| 3344 | using der_g [OF that] | |
| 3345 | by (simp add: has_vector_derivative_def has_derivative_vector_1) | |
| 3346 | then have der': "\<And>x. x \<in> ?lift ` S \<Longrightarrow> | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 3347 | (?lift \<circ> g \<circ> ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" | 
| 67998 | 3348 | by (auto simp: o_def) | 
| 3349 | have inj': "inj_on (vec \<circ> g \<circ> ?drop) (vec ` S)" | |
| 3350 | using inj by (simp add: inj_on_def) | |
| 3351 | let ?fg = "\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)" | |
| 3352 | have "((\<lambda>x. ?fg x $ i) absolutely_integrable_on S \<and> ((\<lambda>x. ?fg x $ i) has_integral b $ i) S | |
| 3353 | \<longleftrightarrow> (\<lambda>x. f x $ i) absolutely_integrable_on g ` S \<and> ((\<lambda>x. f x $ i) has_integral b $ i) (g ` S))" for i | |
| 3354 | using has_absolute_integral_change_of_variables [OF S' der' inj', of "\<lambda>x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] | |
| 3355 | unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def | |
| 3356 | by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) | |
| 3357 | then have "?fg absolutely_integrable_on S \<and> (?fg has_integral b) S | |
| 3358 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> (f has_integral b) (g ` S)" | |
| 3359 | unfolding has_integral_componentwise_iff [where y=b] | |
| 3360 | absolutely_integrable_componentwise_iff [where f=f] | |
| 3361 | absolutely_integrable_componentwise_iff [where f = ?fg] | |
| 3362 | by (force simp: Basis_vec_def cart_eq_inner_axis) | |
| 3363 | then show ?thesis | |
| 3364 | using absolutely_integrable_on_def by blast | |
| 3365 | qed | |
| 3366 | ||
| 3367 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3368 | corollary absolutely_integrable_change_of_variables_1: | 
| 67998 | 3369 |   fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
 | 
| 3370 | assumes S: "S \<in> sets lebesgue" | |
| 3371 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)" | |
| 3372 | and inj: "inj_on g S" | |
| 3373 | shows "(f absolutely_integrable_on g ` S \<longleftrightarrow> | |
| 3374 | (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)" | |
| 70136 | 3375 | using has_absolute_integral_change_of_variables_1 [OF assms] by auto | 
| 67998 | 3376 | |
| 73928 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3377 | text \<open>when @{term "n=1"}\<close>
 | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3378 | lemma has_absolute_integral_change_of_variables_1': | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3379 | fixes f :: "real \<Rightarrow> real" and g :: "real \<Rightarrow> real" | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3380 | assumes S: "S \<in> sets lebesgue" | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3381 | and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_field_derivative g' x) (at x within S)" | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3382 | and inj: "inj_on g S" | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3383 | shows "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3384 | integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) = b | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3385 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3386 | proof - | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3387 | have "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \<and> | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3388 | integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3389 | \<longleftrightarrow> (\<lambda>x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \<and> | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3390 | integral (g ` S) (\<lambda>x. vec (f x)) = (vec b :: real ^ 1)" | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 3391 | using assms unfolding has_real_derivative_iff_has_vector_derivative | 
| 73928 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3392 | by (intro has_absolute_integral_change_of_variables_1 assms) auto | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3393 | thus ?thesis | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3394 | by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3395 | qed | 
| 
3b76524f5a85
Imported lots of material from Stirling_Formula/Gamma_Asymptotics
 paulson <lp15@cam.ac.uk> parents: 
73648diff
changeset | 3396 | |
| 67998 | 3397 | |
| 69683 | 3398 | subsection\<open>Change of variables for integrals: special case of linear function\<close> | 
| 67998 | 3399 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3400 | lemma has_absolute_integral_change_of_variables_linear: | 
| 67998 | 3401 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3402 | assumes "linear g" | |
| 3403 | shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> | |
| 3404 | integral S (\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) = b | |
| 3405 | \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b" | |
| 3406 | proof (cases "det(matrix g) = 0") | |
| 3407 | case True | |
| 3408 | then have "negligible(g ` S)" | |
| 3409 | using assms det_nz_iff_inj negligible_linear_singular_image by blast | |
| 3410 | with True show ?thesis | |
| 3411 | by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) | |
| 3412 | next | |
| 3413 | case False | |
| 3414 | then obtain h where h: "\<And>x. x \<in> S \<Longrightarrow> h (g x) = x" "linear h" | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 3415 | using assms det_nz_iff_inj linear_injective_isomorphism by metis | 
| 67998 | 3416 | show ?thesis | 
| 3417 | proof (rule has_absolute_integral_change_of_variables_invertible) | |
| 3418 | show "(g has_derivative g) (at x within S)" for x | |
| 3419 | by (simp add: assms linear_imp_has_derivative) | |
| 3420 | show "continuous_on (g ` S) h" | |
| 3421 | using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast | |
| 3422 | qed (use h in auto) | |
| 3423 | qed | |
| 3424 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3425 | lemma absolutely_integrable_change_of_variables_linear: | 
| 67998 | 3426 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3427 | assumes "linear g" | |
| 3428 | shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S | |
| 3429 | \<longleftrightarrow> f absolutely_integrable_on (g ` S)" | |
| 3430 | using assms has_absolute_integral_change_of_variables_linear by blast | |
| 3431 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3432 | lemma absolutely_integrable_on_linear_image: | 
| 67998 | 3433 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3434 | assumes "linear g" | |
| 3435 | shows "f absolutely_integrable_on (g ` S) | |
| 3436 | \<longleftrightarrow> (f \<circ> g) absolutely_integrable_on S \<or> det(matrix g) = 0" | |
| 3437 | unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff | |
| 3438 | by (auto simp: set_integrable_def) | |
| 3439 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3440 | lemma integral_change_of_variables_linear: | 
| 67998 | 3441 |   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
 | 
| 3442 | assumes "linear g" | |
| 3443 | and "f absolutely_integrable_on (g ` S) \<or> (f \<circ> g) absolutely_integrable_on S" | |
| 3444 | shows "integral (g ` S) f = \<bar>det (matrix g)\<bar> *\<^sub>R integral S (f \<circ> g)" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3445 | proof - | 
| 67998 | 3446 | have "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S) \<or> (f absolutely_integrable_on g ` S)" | 
| 3447 | using absolutely_integrable_on_linear_image assms by blast | |
| 3448 | moreover | |
| 3449 | have ?thesis if "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" | |
| 3450 | using has_absolute_integral_change_of_variables_linear [OF \<open>linear g\<close>] that | |
| 3451 | by (auto simp: o_def) | |
| 3452 | ultimately show ?thesis | |
| 3453 | using absolutely_integrable_change_of_variables_linear [OF \<open>linear g\<close>] | |
| 3454 | by blast | |
| 3455 | qed | |
| 3456 | ||
| 69683 | 3457 | subsection\<open>Change of variable for measure\<close> | 
| 67998 | 3458 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3459 | lemma has_measure_differentiable_image: | 
| 67998 | 3460 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 3461 | assumes "S \<in> sets lebesgue" | |
| 3462 | and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 3463 | and "inj_on f S" | |
| 3464 | shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) = m | |
| 3465 | \<longleftrightarrow> ((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral m) S" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3466 | using has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"] | 
| 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3467 | unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def | 
| 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3468 | by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) | 
| 67998 | 3469 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3470 | lemma measurable_differentiable_image_eq: | 
| 67998 | 3471 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 3472 | assumes "S \<in> sets lebesgue" | |
| 3473 | and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 3474 | and "inj_on f S" | |
| 3475 | shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S" | |
| 3476 | using has_measure_differentiable_image [OF assms] | |
| 3477 | by blast | |
| 3478 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3479 | lemma measurable_differentiable_image_alt: | 
| 67998 | 3480 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 3481 | assumes "S \<in> sets lebesgue" | |
| 3482 | and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 3483 | and "inj_on f S" | |
| 3484 | shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3485 | using measurable_differentiable_image_eq [OF assms] | 
| 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3486 | by (simp only: absolutely_integrable_on_iff_nonneg) | 
| 67998 | 3487 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3488 | lemma measure_differentiable_image_eq: | 
| 67998 | 3489 |   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
 | 
| 3490 | assumes S: "S \<in> sets lebesgue" | |
| 3491 | and der_f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 3492 | and inj: "inj_on f S" | |
| 3493 | and intS: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S" | |
| 3494 | shows "measure lebesgue (f ` S) = integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3495 | using measurable_differentiable_image_eq [OF S der_f inj] | 
| 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 3496 | assms has_measure_differentiable_image by blast | 
| 67998 | 3497 | |
| 3498 | end |