diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy +++ b/src/HOL/Analysis/Change_Of_Vars.thy @@ -1,3474 +1,3478 @@ (* Title: HOL/Analysis/Change_Of_Vars.thy Authors: LC Paulson, based on material from HOL Light *) section\Change of Variables Theorems\ theory Change_Of_Vars imports Vitali_Covering_Theorem Determinants begin subsection \Measurable Shear and Stretch\ proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" let ?mn = "axis m 1 - axis n (1::real)" have eq1: "measure lebesgue (cbox a ?c) = measure lebesgue (?f ` cbox a b) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a$m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m})" proof (rule measure_Un3_negligible) show "cbox a ?c \ {x. ?mn \ x \ a$m} \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "negligible {x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ a $ m}) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible ((?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ a $ m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ b$m}) \ {x. ?mn \ x = b$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne apply (auto simp: algebra_simps mem_box_cart inner_axis') apply (drule_tac x=m in spec)+ apply simp done ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") proof show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) apply (auto simp: mem_box_cart split: if_split_asm) done qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" have eq2: "measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a $ m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m}) = measure lebesgue (cbox a (\ i. if i = m then a $ m + b $ n else b $ i))" proof (rule measure_translate_add[of "cbox a ?c \ {x. ?mn \ x \ a$m}" "cbox a ?c \ {x. ?mn \ x \ b$m}" "(\ i. if i = m then a$m - b$m else 0)" "cbox a (\ i. if i = m then a$m + b$n else b$i)"]) show "(cbox a ?c \ {x. ?mn \ x \ a$m}) \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "\x. \x $ n + a $ m \ x $ m\ \ x \ (+) (\ i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \ x $ m}" using \m \ n\ by (rule_tac x="x - (\ i. if i = m then a$m - b$m else 0)" in image_eqI) (simp_all add: mem_box_cart) then have imeq: "(+) ?d ` {x. b $ m \ ?mn \ x} = {x. a $ m \ ?mn \ x}" using \m \ n\ by (auto simp: mem_box_cart inner_axis' algebra_simps) have "\x. \0 \ a $ n; x $ n + a $ m \ x $ m; \i. i \ m \ a $ i \ x $ i \ x $ i \ b $ i\ \ a $ m \ x $ m" using \m \ n\ by force then have "(+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {x. a $ m \ ?mn \ x}" using an ab_ne apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) by (metis (full_types) add_mono mult_2_right) then show "cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") using an \m \ n\ apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) apply (drule_tac x=n in spec)+ by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) have "negligible{x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x})) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}))" by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q using eq1 eq2 eq3 by (simp add: algebra_simps) qed proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True have m0: "0 < \prod m UNIV\" using True by simp have "(indicat_real (?f ` S) has_integral \prod m UNIV\ * measure lebesgue S) UNIV" proof (clarsimp simp add: has_integral_alt [where i=UNIV]) fix e :: "real" assume "e > 0" have "(indicat_real S has_integral (measure lebesgue S)) UNIV" using assms lmeasurable_iff_has_integral by blast then obtain B where "B>0" and B: "\a b. ball 0 B \ cbox a b \ \z. (indicat_real S has_integral z) (cbox a b) \ \z - measure lebesgue S\ < e / \prod m UNIV\" by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \e > 0\) show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (indicat_real (?f ` S) has_integral z) (cbox a b) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" proof (intro exI conjI allI) let ?C = "Max (range (\k. \m k\)) * B" show "?C > 0" using True \B > 0\ by (simp add: Max_gr_iff) show "ball 0 ?C \ cbox u v \ (\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" for u v proof assume uv: "ball 0 ?C \ cbox u v" with \?C > 0\ have cbox_ne: "cbox u v \ {}" using centre_in_ball by blast let ?\ = "\k. u$k / m k" let ?\ = "\k. v$k / m k" have invm0: "\k. inverse (m k) \ 0" using True by auto have "ball 0 B \ (\x. \ k. x $ k / m k) ` ball 0 ?C" proof clarsimp fix x :: "real^'n" assume x: "norm x < B" have [simp]: "\Max (range (\k. \m k\))\ = Max (range (\k. \m k\))" by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) have "norm (\ k. m k * x $ k) \ norm (Max (range (\k. \m k\)) *\<^sub>R x)" by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) also have "\ < ?C" using x by simp (metis \B > 0\ \?C > 0\ mult.commute real_mult_less_iff1 zero_less_mult_pos) finally have "norm (\ k. m k * x $ k) < ?C" . then show "x \ (\x. \ k. x $ k / m k) ` ball 0 ?C" using stretch_Galois [of "inverse \ m"] True by (auto simp: image_iff field_simps) qed then have Bsub: "ball 0 B \ cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))" using cbox_ne uv image_stretch_interval_cart [of "inverse \ m" u v, symmetric] by (force simp: field_simps) obtain z where zint: "(indicat_real S has_integral z) (cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" and zless: "\z - measure lebesgue S\ < e / \prod m UNIV\" using B [OF Bsub] by blast have ind: "indicat_real (?f ` S) = (\x. indicator S (\ k. x$k / m k))" using True stretch_Galois [of m] by (force simp: indicator_def) show "\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e" proof (simp add: ind, intro conjI exI) have "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) ((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" using True has_integral_stretch_cart [OF zint, of "inverse \ m"] by (simp add: field_simps prod_dividef) moreover have "((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))) = cbox u v" using True image_stretch_interval_cart [of "inverse \ m" u v, symmetric] image_stretch_interval_cart [of "\k. 1" u v, symmetric] \cbox u v \ {}\ by (simp add: field_simps image_comp o_def) ultimately show "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) (cbox u v)" by simp have "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ = \prod m UNIV\ * \z - measure lebesgue S\" by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') also have "\ < e" using zless True by (simp add: field_simps) finally show "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ < e" . qed qed qed qed then show ?thesis by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) next case False then obtain k where "m k = 0" and prm: "prod m UNIV = 0" by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) then have "(?f ` S) \ lmeasurable" by (simp add: negligible_iff_measure) with nfS show ?thesis by (simp add: prm negligible_iff_measure0) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" assume "linear f" and "linear g" and f [rule_format]: "\S \ lmeasurable. f ` S \ lmeasurable \ ?Q f S" and g [rule_format]: "\S \ lmeasurable. g ` S \ lmeasurable \ ?Q g S" and S: "S \ lmeasurable" then have gS: "g ` S \ lmeasurable" by blast show "(f \ g) ` S \ lmeasurable \ ?Q (f \ g) S" using f [OF gS] g [OF S] matrix_compose [OF \linear g\ \linear f\] by (simp add: o_def image_comp abs_mult det_mul) next fix f :: "real^'n::_ \ real^'n::_" and i and S :: "(real^'n::_) set" assume "linear f" and 0: "\x. f x $ i = 0" and "S \ lmeasurable" then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" using lmeasurable_iff_indicator_has_integral \linear f\ \\ inj f\ negligible_UNIV negligible_linear_singular_image by blast have "measure lebesgue (f ` S) = 0" by (meson \\ inj f\ \linear f\ negligible_imp_measure0 negligible_linear_singular_image) also have "\ = \det (matrix f)\ * measure lebesgue S" by (simp add: detf) finally show "?Q f S" . qed next fix c and S :: "(real^'n::_) set" assume "S \ lmeasurable" show "(\a. \ i. c i * a $ i) ` S \ lmeasurable \ ?Q (\a. \ i. c i * a $ i) S" proof show "(\a. \ i. c i * a $ i) ` S \ lmeasurable" by (simp add: \S \ lmeasurable\ measurable_stretch) show "?Q (\a. \ i. c i * a $ i) S" by (simp add: measure_stretch [OF \S \ lmeasurable\, of c] axis_def matrix_def det_diagonal) qed next fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have him: "?h ` (cbox a b) \ {}" by blast have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+ show ?thesis using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\i. (b - a)$i", symmetric] by (simp add: eq content_cbox_cart False) qed have "(\ i j. if Fun.swap m n id i = j then 1 else 0) = (\ i j. if j = Fun.swap m n id i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Fun.swap m n id j)" by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" proof show "?h ` S \ lmeasurable" "?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ qed next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" proof cases assume "m < n" have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j < i\ axis_def \m < n\ by auto with \m < n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) next assume "n < m" have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j > i\ axis_def \m > n\ by auto with \m > n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) qed have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have ne: "(+) (\ i. if i = n then - a $ n else 0) ` cbox a b \ {}" by auto let ?v = "\ i. if i = n then - a $ n else 0" have "?h ` cbox a b = (+) (\ i. if i = m \ i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" using \m \ n\ unfolding image_comp o_def by (force simp: vec_eq_iff) then have "measure lebesgue (?h ` (cbox a b)) = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` (+) ?v ` cbox a b)" by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) by (metis cbox_borel cbox_translation measure_completion sets_lborel) also have "\ = measure lebesgue (cbox a b)" by (rule measure_translation) finally show ?thesis . qed show "?h ` S \ lmeasurable \ ?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force qed with assms show "(f ` S) \ lmeasurable" "?Q f S" by metis+ qed lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" proof - have "linear f" by (simp add: f orthogonal_transformation_linear) then show "f ` S \ lmeasurable" by (metis S measurable_linear_image) show "measure lebesgue (f ` S) = measure lebesgue S" by (simp add: measure_linear_image \linear f\ S f) qed proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next case False then have frS: "frontier S \ {}" using \bounded S\ frontier_eq_empty not_bounded_UNIV by blast have "S \ lmeasurable" by (simp add: \bounded S\ measurable_Jordan neg) have null: "(frontier S) \ null_sets lebesgue" by (metis neg negligible_iff_null_sets) have "frontier S \ lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" using neg negligible_imp_measurable negligible_iff_measure by blast+ with \e > 0\ sets_lebesgue_outer_open obtain U where "open U" and U: "frontier S \ U" "U - frontier S \ lmeasurable" "emeasure lebesgue (U - frontier S) < e" by (metis fmeasurableD) with null have "U \ lmeasurable" by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) have "measure lebesgue (U - frontier S) = measure lebesgue U" using mS0 by (simp add: \U \ lmeasurable\ fmeasurableD measure_Diff_null_set null) with U have mU: "measure lebesgue U < e" by (simp add: emeasure_eq_measure2 ennreal_less_iff) show ?thesis proof have "U \ UNIV" using \U \ lmeasurable\ by auto then have "- U \ {}" by blast with \open U\ \frontier S \ U\ show "setdist (frontier S) (- U) > 0" by (auto simp: \bounded S\ open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) fix T assume "T \ lmeasurable" and T: "\t. t \ T \ \y. y \ S \ dist y t < setdist (frontier S) (- U)" then have "measure lebesgue T - measure lebesgue S \ measure lebesgue (T - S)" by (simp add: \S \ lmeasurable\ measure_diff_le_measure_setdiff) also have "\ \ measure lebesgue U" proof - have "T - S \ U" proof clarify fix x assume "x \ T" and "x \ S" then obtain y where "y \ S" and y: "dist y x < setdist (frontier S) (- U)" using T by blast have "closed_segment x y \ frontier S \ {}" using connected_Int_frontier \x \ S\ \y \ S\ by blast then obtain z where z: "z \ closed_segment x y" "z \ frontier S" by auto with y have "dist z x < setdist(frontier S) (- U)" by (auto simp: dist_commute dest!: dist_in_closed_segment) with z have False if "x \ -U" using setdist_le_dist [OF \z \ frontier S\ that] by auto then show "x \ U" by blast qed then show ?thesis by (simp add: \S \ lmeasurable\ \T \ lmeasurable\ \U \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Diff) qed finally have "measure lebesgue T - measure lebesgue S \ measure lebesgue U" . with mU show "measure lebesgue T < measure lebesgue S + e" by linarith qed qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" and bounded: "\x. x \ S \ \det (matrix (f' x))\ \ B" shows measurable_bounded_differentiable_image: "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True then have "S = {}" by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) then show ?thesis by auto next case False then have "B \ 0" by arith let ?\ = "measure lebesgue" have f_diff: "f differentiable_on S" using deriv by (auto simp: differentiable_on_def differentiable_def) have eps: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * ?\ S" (is "?ME") if "e > 0" for e proof - have eps_d: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * (?\ S + d)" (is "?MD") if "d > 0" for d proof - obtain T where T: "open T" "S \ T" and TS: "(T-S) \ lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" using S \d > 0\ sets_lebesgue_outer_open by blast then have "?\ (T-S) < d" by (metis emeasure_eq_measure2 ennreal_leI not_less) with S T TS have "T \ lmeasurable" and Tless: "?\ T < ?\ S + d" by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) have "\r. 0 < r \ r < d \ ball x r \ T \ f ` (S \ ball x r) \ lmeasurable \ ?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" if "x \ S" "d > 0" for x d proof - have lin: "linear (f' x)" and lim0: "((\y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \ 0) (at x within S)" using deriv \x \ S\ by (auto simp: has_derivative_within bounded_linear.linear field_simps) have bo: "bounded (f' x ` ball 0 1)" by (simp add: bounded_linear_image linear_linear lin) have neg: "negligible (frontier (f' x ` ball 0 1))" using deriv has_derivative_linear \x \ S\ by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) - have 0: "0 < e * unit_ball_vol (real CARD('n))" - using \e > 0\ by simp + let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)" + have 0: "0 < e * ?unit_vol" + using \e > 0\ by (simp add: content_ball_pos) obtain k where "k > 0" and k: "\U. \U \ lmeasurable; \y. y \ U \ \z. z \ f' x ` ball 0 1 \ dist z y < k\ - \ ?\ U < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" + \ ?\ U < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast obtain l where "l > 0" and l: "ball x l \ T" using \x \ S\ \open T\ \S \ T\ openE by blast obtain \ where "0 < \" and \: "\y. \y \ S; y \ x; dist y x < \\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" using lim0 \k > 0\ by (simp add: Lim_within) (auto simp add: field_simps) define r where "r \ min (min l (\/2)) (min 1 (d/2))" show ?thesis proof (intro exI conjI) show "r > 0" "r < d" using \l > 0\ \\ > 0\ \d > 0\ by (auto simp: r_def) have "r \ l" by (auto simp: r_def) with l show "ball x r \ T" by auto have ex_lessK: "\x' \ ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" if "y \ S" and "dist x y < r" for y proof (cases "y = x") case True with lin linear_0 \k > 0\ that show ?thesis by (rule_tac x=0 in bexI) (auto simp: linear_0) next case False then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" by (simp add: lin linear_scale) 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)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" using \r > 0\ by (simp add: divide_simps scale_right_diff_distrib [symmetric]) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: algebra_simps field_split_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" using that \0 < \\ by (simp add: dist_commute r_def \ [OF \y \ S\ False]) finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . show "(y - x) /\<^sub>R r \ ball 0 1" using that \r > 0\ by (simp add: dist_norm divide_simps norm_minus_commute) qed qed let ?rfs = "(\x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \ ball x r)" have rfs_mble: "?rfs \ lmeasurable" proof (rule bounded_set_imp_lmeasurable) have "f differentiable_on S \ ball x r" using f_diff by (auto simp: fmeasurableD differentiable_on_subset) with S show "?rfs \ sets lebesgue" by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) let ?B = "(\(x, y). x + y) ` (f' x ` ball 0 1 \ ball 0 k)" have "bounded ?B" by (simp add: bounded_plus [OF bo]) moreover have "?rfs \ ?B" apply (auto simp: dist_norm image_iff dest!: ex_lessK) by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) ultimately show "bounded (?rfs)" by (rule bounded_subset) qed then have "(\x. r *\<^sub>R x) ` ?rfs \ lmeasurable" by (simp add: measurable_linear_image) with \r > 0\ have "(+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) then have "(+) (f x) ` (+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" using measurable_translation by blast then show fsb: "f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) have "?\ (f ` (S \ ball x r)) = ?\ (?rfs) * r ^ CARD('n)" using \r > 0\ fsb by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) - also have "\ \ (\det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))) * r ^ CARD('n)" + also have "\ \ (\det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)" proof - - have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" + have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using rfs_mble by (force intro: k dest!: ex_lessK) - then have "?\ (?rfs) < \det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))" - by (simp add: lin measure_linear_image [of "f' x"] content_ball) + then have "?\ (?rfs) < \det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol" + by (simp add: lin measure_linear_image [of "f' x"]) with \r > 0\ show ?thesis by auto qed also have "\ \ (B + e) * ?\ (ball x r)" - using bounded [OF \x \ S\] \r > 0\ by (simp add: content_ball algebra_simps) + using bounded [OF \x \ S\] \r > 0\ + by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos) finally show "?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" . qed qed then obtain r where r0d: "\x d. \x \ S; d > 0\ \ 0 < r x d \ r x d < d" and rT: "\x d. \x \ S; d > 0\ \ ball x (r x d) \ T" and r: "\x d. \x \ S; d > 0\ \ (f ` (S \ ball x (r x d))) \ lmeasurable \ ?\ (f ` (S \ ball x (r x d))) \ (B + e) * ?\ (ball x (r x d))" by metis obtain C where "countable C" and Csub: "C \ {(x,r x t) |x t. x \ S \ 0 < t}" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible(S - (\i \ C. ball (fst i) (snd i)))" apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \ S \ 0 < t}" fst snd]) apply auto by (metis dist_eq_0_iff r0d) let ?UB = "(\(x,s) \ C. ball x s)" have eq: "f ` (S \ ?UB) = (\(x,s) \ C. f ` (S \ ball x s))" by auto have mle: "?\ (\(x,s) \ K. f ` (S \ ball x s)) \ (B + e) * (?\ S + d)" (is "?l \ ?r") if "K \ C" and "finite K" for K proof - have gt0: "b > 0" if "(a, b) \ K" for a b using Csub that \K \ C\ r0d by auto have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" using pwC that apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) by (metis subsetD fst_conv snd_conv) have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r assume "(x,r) \ K" then have "x \ S" using Csub \K \ C\ by auto show "f ` (S \ ball x r) \ sets lebesgue" by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" apply (rule sum_mono) using Csub r \K \ C\ by auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" using \B \ 0\ \e > 0\ by (simp add: inj sum.reindex prod.case_distrib) also have "\ = (B + e) * ?\ (\(x,s) \ K. ball x s)" using \B \ 0\ \e > 0\ that by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) using Csub rT by force+ also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . qed have fSUB_mble: "(f ` (S \ ?UB)) \ lmeasurable" unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: fmeasurable_UN_bound [OF \countable C\ _ mle]) have fSUB_meas: "?\ (f ` (S \ ?UB)) \ (B + e) * (?\ S + d)" (is "?MUB") unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: measure_UN_bound [OF \countable C\ _ mle]) have neg: "negligible ((f ` (S \ ?UB) - f ` S) \ (f ` S - f ` (S \ ?UB)))" proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) show "f differentiable_on S - (\i\C. ball (fst i) (snd i))" by (meson DiffE differentiable_on_subset subsetI f_diff) qed force show "f ` S \ lmeasurable" by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) show ?MD using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp qed show "f ` S \ lmeasurable" using eps_d [of 1] by simp show ?ME proof (rule field_le_epsilon) fix \ :: real assume "0 < \" then show "?\ (f ` S) \ (B + e) * ?\ S + \" using eps_d [of "\ / (B+e)"] \e > 0\ \B \ 0\ by (auto simp: divide_simps mult_ac) qed qed show ?thesis proof (cases "?\ S = 0") case True with eps have "?\ (f ` S) = 0" by (metis mult_zero_right not_le zero_less_measure_iff) then show ?thesis using eps [of 1] by (simp add: True) next case False have "?\ (f ` S) \ B * ?\ S" proof (rule field_le_epsilon) fix e :: real assume "e > 0" then show "?\ (f ` S) \ B * ?\ S + e" using eps [of "e / ?\ S"] False by (auto simp: algebra_simps zero_less_measure_iff) qed with eps [of 1] show ?thesis by auto qed qed then show "f ` S \ lmeasurable" ?M by blast+ qed lemma m_diff_image_weak: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto define m where "m \ integral S (\x. \det (matrix (f' x))\)" have *: "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" if "e > 0" for e proof - define T where "T \ \n. {x \ S. n * e \ \det (matrix (f' x))\ \ \det (matrix (f' x))\ < (Suc n) * e}" have meas_t: "T n \ lmeasurable" for n proof - have *: "(\x. \det (matrix (f' x))\) \ borel_measurable (lebesgue_on S)" using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) have [intro]: "x \ sets (lebesgue_on S) \ x \ sets lebesgue" for x using S sets_restrict_space_subset by blast have "{x \ S. real n * e \ \det (matrix (f' x))\} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) then have 1: "{x \ S. real n * e \ \det (matrix (f' x))\} \ lmeasurable" using S by (simp add: fmeasurableI2) have "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) then have 2: "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ lmeasurable" using S by (simp add: fmeasurableI2) show ?thesis using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) qed have aint_T: "\k. (\x. \det (matrix (f' x))\) absolutely_integrable_on T k" using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) using that apply auto using of_int_floor_le pos_le_divide_eq apply blast by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" by (simp add: meas_t) next fix x :: "(real,'n) vec" assume "x \ T n" show "(f has_derivative f' x) (at x within T n)" by (metis (no_types, lifting) \x \ T n\ deriv has_derivative_within_subset mem_Collect_eq subsetI T_def) show "\det (matrix (f' x))\ \ (Suc n) * e" using \x \ T n\ T_def by auto next show "(\x. \det (matrix (f' x))\) integrable_on T n" using aint_T absolutely_integrable_on_def by blast qed have disT: "disjoint (range T)" unfolding disjoint_def proof clarsimp show "T m \ T n = {}" if "T m \ T n" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) with \e > 0\ show ?case unfolding T_def proof (clarsimp simp add: Collect_conj_eq [symmetric]) fix x assume "e > 0" "m < n" "n * e \ \det (matrix (f' x))\" "\det (matrix (f' x))\ < (1 + real m) * e" then have "n < 1 + real m" by (metis (no_types, hide_lams) less_le_trans mult.commute not_le real_mult_le_cancel_iff2) then show "False" using less.hyps by linarith qed qed auto qed have injT: "inj_on T ({n. T n \ {}})" unfolding inj_on_def proof clarsimp show "m = n" if "T m = T n" "T n \ {}" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) by (metis add.commute less_le_trans nat_less_real_le not_le real_mult_le_cancel_iff2) then show ?case using less.prems by blast qed auto qed have sum_eq_Tim: "(\k\n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \ real" and n proof (subst sum.reindex_nontrivial) fix i j assume "i \ {..n}" "j \ {..n}" "i \ j" "T i = T j" with that injT [unfolded inj_on_def] show "f (T i) = 0" by simp metis qed (use atMost_atLeast0 in auto) let ?B = "m + e * ?\ S" have "(\k\n. ?\ (f ` T k)) \ ?B" for n proof - have "(\k\n. ?\ (f ` T k)) \ (\k\n. ((k+1) * e) * ?\(T k))" proof (rule sum_mono [OF measure_bounded_differentiable_image]) show "(f has_derivative f' x) (at x within T k)" if "x \ T k" for k x using that unfolding T_def by (blast intro: deriv has_derivative_within_subset) show "(\x. \det (matrix (f' x))\) integrable_on T k" for k using absolutely_integrable_on_def aint_T by blast show "\det (matrix (f' x))\ \ real (k + 1) * e" if "x \ T k" for k x using T_def that by auto qed (use meas_t in auto) also have "\ \ (\k\n. (k * e) * ?\(T k)) + (\k\n. e * ?\(T k))" by (simp add: algebra_simps sum.distrib) also have "\ \ ?B" proof (rule add_mono) have "(\k\n. real k * e * ?\ (T k)) = (\k\n. integral (T k) (\x. k * e))" by (simp add: lmeasure_integral [OF meas_t] flip: integral_mult_right integral_mult_left) also have "\ \ (\k\n. integral (T k) (\x. (abs (det (matrix (f' x))))))" proof (rule sum_mono) fix k assume "k \ {..n}" show "integral (T k) (\x. k * e) \ integral (T k) (\x. \det (matrix (f' x))\)" proof (rule integral_le [OF integrable_on_const [OF meas_t]]) show "(\x. \det (matrix (f' x))\) integrable_on T k" using absolutely_integrable_on_def aint_T by blast next fix x assume "x \ T k" show "k * e \ \det (matrix (f' x))\" using \x \ T k\ T_def by blast qed qed also have "\ = sum (\T. integral T (\x. \det (matrix (f' x))\)) (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = integral (\k\n. T k) (\x. \det (matrix (f' x))\)" proof (rule integral_unique [OF has_integral_Union, symmetric]) fix S assume "S \ T ` {..n}" then show "((\x. \det (matrix (f' x))\) has_integral integral S (\x. \det (matrix (f' x))\)) S" using absolutely_integrable_on_def aint_T by blast next show "pairwise (\S S'. negligible (S \ S')) (T ` {..n})" using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) qed auto also have "\ \ m" unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" apply (rule set_integrable_subset [OF aint_S]) apply (intro measurable meas_t fmeasurableD) apply (force simp: Seq) done then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) finally show "(\k\n. real k * e * ?\ (T k)) \ m" . next have "(\k\n. ?\ (T k)) = sum ?\ (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = ?\ (\k\n. T k)" using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) also have "\ \ ?\ S" using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) finally have "(\k\n. ?\ (T k)) \ ?\ S" . then show "(\k\n. e * ?\ (T k)) \ e * ?\ S" by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) qed finally show "(\k\n. ?\ (f ` T k)) \ ?B" . qed moreover have "measure lebesgue (\k\n. f ` T k) \ (\k\n. ?\ (f ` T k))" for n by (simp add: fmeasurableD meas_ft measure_UNION_le) ultimately have B_ge_m: "?\ (\k\n. (f ` T k)) \ ?B" for n by (meson order_trans) have "(\n. f ` T n) \ lmeasurable" by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) moreover have "?\ (\n. f ` T n) \ m + e * ?\ S" by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) ultimately show "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" by (auto simp: Seq image_Union) qed show ?thesis proof show "f ` S \ lmeasurable" using * linordered_field_no_ub by blast let ?x = "m - ?\ (f ` S)" have False if "?\ (f ` S) > integral S (\x. \det (matrix (f' x))\)" proof - have ml: "m < ?\ (f ` S)" using m_def that by blast then have "?\ S \ 0" using "*"(2) bgauge_existence_lemma by fastforce with ml have 0: "0 < - (m - ?\ (f ` S))/2 / ?\ S" using that zero_less_measure_iff by force then show ?thesis using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) qed then show "?\ (f ` S) \ integral S (\x. \det (matrix (f' x))\)" by fastforce qed qed theorem fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" apply (auto simp: mem_box_cart) apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" and mfIn: "?\ (f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" (is ?MN) for n proof - have In: "?I n \ lmeasurable" by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_within_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" proof - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int absolutely_integrable_integrable_bound by force then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) then show ?thesis using absolutely_integrable_on_def by blast qed ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" by (simp add: int_In int integral_subset_le) ultimately show "f ` ?I n \ lmeasurable" ?MN by auto qed have "?I k \ ?I n" if "k \ n" for k n by (rule Int_mono) (use that in \auto simp: subset_interval_imp_cart\) then have "(\k\n. f ` ?I k) = f ` ?I n" for n by (fastforce simp add:) with mfIn have "?\ (\k\n. f ` ?I k) \ integral S (\x. \det (matrix (f' x))\)" for n by simp then have "(\n. f ` ?I n) \ lmeasurable" "?\ (\n. f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ then show "f ` S \ lmeasurable" ?M by (metis Seq image_UN)+ qed lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \ (\x. (\n. g n x) \ f x))" (is "?lhs = ?rhs") proof assume f: ?lhs have leb_f: "{x. a \ f x \ f x < b} \ sets lebesgue" for a b proof - have "{x. a \ f x \ f x < b} = {x. f x < b} - {x. f x < a}" by auto also have "\ \ sets lebesgue" using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto finally show ?thesis . qed have "g n x \ f x" if inc_g: "\n x. 0 \ g n x \ g n x \ g (Suc n) x" and meas_g: "\n. g n \ borel_measurable lebesgue" and fin: "\n. finite(range (g n))" and lim: "\x. (\n. g n x) \ f x" for g n x proof - have "\r>0. \N. \n\N. dist (g n x) (f x) \ r" if "g n x > f x" proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) have "\na\N. g n x - f x \ dist (g na x) (f x)" for N apply (rule_tac x="N+n" in exI) using g [of N] by (auto simp: dist_norm) with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis apply (auto simp: lim_sequentially) by (meson less_le_not_le not_le_imp_less) qed moreover let ?\ = "\n k. indicator {y. k/2^n \ f y \ f y < (k+1)/2^n}" let ?g = "\n x. (\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x)" have "\g. (\n x. 0 \ g n x \ g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \(\x. (\n. g n x) \ f x)" proof (intro exI allI conjI) show "0 \ ?g n x" for n x proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) apply (drule spec [where x=x]) using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] by linarith qed show "?g n x \ ?g (Suc n) x" for n x proof - have "?g n x = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * (indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x + indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x))" by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x)" by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \ f y \ f y < (2 * k+1)/2 ^ Suc n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2 * k+1) + 1)/2 ^ Suc n} x)" by (force simp: field_simps indicator_def intro: sum.cong) also have "\ \ (\k | k \ \ \ \k\ \ 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x))" (is "?a + _ \ ?b") proof - have *: "\sum f I \ sum h I; a + sum h I \ b\ \ a + sum f I \ b" for I a b f and h :: "real\real" by linarith let ?h = "\k. (2*k+1)/2 ^ Suc n * (indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2*k+1) + 1)/2 ^ Suc n} x)" show ?thesis proof (rule *) show "(\k | k \ \ \ \k\ \ 2 ^ (2*n). 2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < (2 * k+1 + 1)/2 ^ Suc n} x) \ sum ?h {k \ \. \k\ \ 2 ^ (2*n)}" by (rule sum_mono) (simp add: indicator_def field_split_simps) next have \: "?a = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have \: "sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have 0: "(*) 2 ` {k \ \. P k} \ (\x. 2 * x + 1) ` {k \ \. P k} = {}" for P :: "real \ bool" proof - have "2 * i \ 2 * j + 1" for i j :: int by arith thus ?thesis unfolding Ints_def by auto (use of_int_eq_iff in fastforce) qed have "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" unfolding \ \ using finite_abs_int_segment [of "2 ^ (2*n)"] by (subst sum_Un) (auto simp: 0) also have "\ \ ?b" proof (rule sum_mono2) show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" apply auto using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b proof - have "0 \ f x * (2 * 2^n)" by (simp add: f) also have "\ < b+1" by (simp add: that) finally show "0 \ b" using \b \ \\ by (auto simp: elim!: Ints_cases) qed then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b using that by (simp add: indicator_def divide_simps) qed finally show "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} \ ?b" . qed qed finally show ?thesis . qed show "?g n \ borel_measurable lebesgue" for n apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) using leb_f sets_restrict_UNIV by auto show "finite (range (?g n))" for n proof - have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" for x proof (cases "\k. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ f x \ f x < (k+1)/2^n") case True then show ?thesis by (blast intro: indicator_sum_eq) next case False then have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) = 0" by auto then show ?thesis by force qed then have "range (?g n) \ ((\k. (k/2^n)) ` {k. k \ \ \ \k\ \ 2 ^ (2*n)})" by auto moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (intro finite_imageI finite_abs_int_segment) ultimately show ?thesis by (rule finite_subset) qed show "(\n. ?g n x) \ f x" for x proof (clarsimp simp add: lim_sequentially) fix e::real assume "e > 0" obtain N1 where N1: "2 ^ N1 > abs(f x)" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by fastforce have "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) < e" if "N1 + N2 \ n" for n proof - let ?m = "real_of_int \2^n * f x\" have "\?m\ \ 2^n * 2^N1" using N1 apply (simp add: f) by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) also have "\ \ 2 ^ (2*n)" by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff power_add power_increasing_iff semiring_norm(76)) finally have m_le: "\?m\ \ 2 ^ (2*n)" . have "?m/2^n \ f x" "f x < (?m + 1)/2^n" by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) then have eq: "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) = dist (?m/2^n) (f x)" by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) have "\2^n\ * \?m/2^n - f x\ = \2^n * (?m/2^n - f x)\" by (simp add: abs_mult) also have "\ < 2 ^ N2 * e" using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto finally have "dist (?m/2^n) (f x) < e" by (simp add: dist_norm) then show ?thesis using eq by linarith qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force qed qed ultimately show ?rhs by metis next assume RHS: ?rhs with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] show ?lhs by (blast intro: order_trans) qed subsection\Borel measurable Jacobian determinant\ lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" using orthogonal_to_subspace_exists [OF less] orthogonal_def by (metis (mono_tags, lifting) mem_Collect_eq span_base) then obtain k where "k > 0" and k: "\e. e > 0 \ \y. y \ S - {0} \ norm y < e \ k * norm y \ \d \ y\" using lb by blast have "\h. \n. ((h n \ S \ h n \ 0 \ k * norm (h n) \ \d \ h n\) \ norm (h n) < 1 / real (Suc n)) \ norm (h (Suc n)) < norm (h n)" proof (rule dependent_nat_choice) show "\y. (y \ S \ y \ 0 \ k * norm y \ \d \ y\) \ norm y < 1 / real (Suc 0)" by simp (metis DiffE insertCI k not_less not_one_le_zero) qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) then obtain \ where \: "\n. \ n \ S - {0}" and kd: "\n. k * norm(\ n) \ \d \ \ n\" and norm_lt: "\n. norm(\ n) < 1/(Suc n)" by force let ?\ = "\n. \ n /\<^sub>R norm (\ n)" have com: "\g. (\n. g n \ sphere (0::'a) 1) \ \l \ sphere 0 1. \\::nat\nat. strict_mono \ \ (g \ \) \ l" using compact_sphere compact_def by metis moreover have "\n. ?\ n \ sphere 0 1" using \ by auto ultimately obtain l::'a and \::"nat\nat" where l: "l \ sphere 0 1" and "strict_mono \" and to_l: "(?\ \ \) \ l" by meson moreover have "continuous (at l) (\x. (\d \ x\ - k))" by (intro continuous_intros) ultimately have lim_dl: "((\x. (\d \ x\ - k)) \ (?\ \ \)) \ (\d \ l\ - k)" by (meson continuous_imp_tendsto) have "\\<^sub>F i in sequentially. 0 \ ((\x. \d \ x\ - k) \ ((\n. \ n /\<^sub>R norm (\ n)) \ \)) i" using \ kd by (auto simp: field_split_simps) then have "k \ \d \ l\" using tendsto_lowerbound [OF lim_dl, of 0] by auto moreover have "d \ l = 0" proof (rule d) show "f l = 0" proof (rule LIMSEQ_unique [of "f \ ?\ \ \"]) have "isCont f l" using \linear f\ linear_continuous_at linear_conv_bounded_linear by blast then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ f l" unfolding comp_assoc using to_l continuous_imp_tendsto by blast have "\ \ 0" using norm_lt LIMSEQ_norm_0 by metis with \strict_mono \\ have "(\ \ \) \ 0" by (metis LIMSEQ_subseq_LIMSEQ) with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" by (simp add: o_def scale) qed qed ultimately show False using \k > 0\ by auto qed ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis using dim_eq_full by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" shows "f x = 0" proof - have "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within (\x. x-a) ` S)" using lim by (simp add: Lim_within dist_norm) then show ?thesis proof (rule lemma_partial_derivatives0 [OF \linear f\]) fix v :: "'a" assume v: "v \ 0" show "\k>0. \e>0. \x. x \ (\x. x - a) ` S - {0} \ norm x < e \ k * norm x \ \v \ x\" using lb [OF v] by (force simp: norm_minus_commute) qed qed proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b proof (rule sets_negligible_symdiff) let ?T = "{x \ S. \e>0. \d>0. \A. A$m$n < b \ (\i j. A$i$j \ \) \ (\y \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x))}" let ?U = "S \ (\e \ {e \ \. e > 0}. \A \ {A. A$m$n < b \ (\i j. A$i$j \ \)}. \d \ {d \ \. 0 < d}. S \ (\y \ S. {x \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x)}))" have "?T = ?U" proof (intro set_eqI iffI) fix x assume xT: "x \ ?T" then show "x \ ?U" proof (clarsimp simp add:) fix q :: real assume "q \ \" "q > 0" then obtain d A where "d > 0" and A: "A $ m $ n < b" "\i j. A $ i $ j \ \" "\y. \y\S; norm (y - x) < d\ \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)" using xT by auto then obtain \ where "d > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with A show "\A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\s. s \ \ \ 0 < s \ (\y\S. norm (y - x) < s \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)))" by force qed next fix x assume xU: "x \ ?U" then show "x \ ?T" proof clarsimp fix e :: "real" assume "e > 0" then obtain \ where \: "e > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with xU obtain A r where "x \ S" and Ar: "A $ m $ n < b" "\i j. A $ i $ j \ \" "r \ \" "r > 0" and "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ \ * norm (y - x)" by (auto simp: split: if_split_asm) then have "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)" by (meson \e > \\ less_eq_real_def mult_right_mono norm_ge_zero order_trans) then show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" using \x \ S\ Ar by blast qed qed moreover have "?U \ sets lebesgue" proof - have coQ: "countable {e \ \. 0 < e}" using countable_Collect countable_rat by blast have ne: "{e \ \. (0::real) < e} \ {}" using zero_less_one Rats_1 by blast have coA: "countable {A. A $ m $ n < b \ (\i j. A $ i $ j \ \)}" proof (rule countable_subset) show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) qed show ?thesis by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ qed show ?thesis by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto qed ultimately show "?T \ sets lebesgue" by simp let ?M = "(?T - {x \ S. matrix (f' x) $ m $ n \ b} \ ({x \ S. matrix (f' x) $ m $ n \ b} - ?T))" let ?\ = "\x v. \\>0. \e>0. \y \ S-{x}. norm (x - y) < e \ \v \ (y - x)\ < \ * norm (x - y)" have nN: "negligible {x \ S. \v\0. ?\ x v}" unfolding negligible_eq_zero_density proof clarsimp fix x v and r e :: "real" assume "x \ S" "v \ 0" "r > 0" "e > 0" and Theta [rule_format]: "?\ x v" moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0" by (simp add: \v \ 0\ \e > 0\) ultimately obtain d where "d > 0" and dless: "\y. \y \ S - {x}; norm (x - y) < d\ \ \v \ (y - x)\ < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)" by metis let ?W = "ball x (min d r) \ {y. \v \ (y - x)\ < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}" have "open {x. \v \ (x - a)\ < b}" for a b by (intro open_Collect_less continuous_intros) show "\d>0. d \ r \ (\U. {x' \ S. \v\0. ?\ x' v} \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d))" proof (intro exI conjI) show "0 < min d r" "min d r \ r" using \r > 0\ \d > 0\ by auto show "{x' \ S. \v. v \ 0 \ (\\>0. \e>0. \z\S - {x'}. norm (x' - z) < e \ \v \ (z - x')\ < \ * norm (x' - z))} \ ball x (min d r) \ ?W" proof (clarsimp simp: dist_norm norm_minus_commute) fix y w assume "y \ S" "w \ 0" and less [rule_format]: "\\>0. \e>0. \z\S - {y}. norm (y - z) < e \ \w \ (z - y)\ < \ * norm (y - z)" and d: "norm (y - x) < d" and r: "norm (y - x) < r" show "\v \ (y - x)\ < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" proof (cases "y = x") case True with \r > 0\ \d > 0\ \e > 0\ \v \ 0\ show ?thesis by simp next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" apply (rule dless) using False \y \ S\ d by (auto simp: norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . qed qed show "?W \ lmeasurable" by (simp add: fmeasurable_Int_fmeasurable borel_open) obtain k::'m where True by metis obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis define b where "b \ norm v" have "b > 0" using \v \ 0\ by (auto simp: b_def) obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) let ?v = "\ i. min d r / CARD('m)" let ?v' = "\ i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r" let ?x' = "inv T x" let ?W' = "(ball ?x' (min d r) \ {y. \(y - ?x')$k\ < e * min d r / (2 * CARD('m) ^ CARD('m))})" have abs: "x - e \ y \ y \ x + e \ abs(y - x) \ e" for x y e::real by auto have "?W = T ` ?W'" proof - have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) have 2: "{y. \v \ (y - x)\ < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} = T ` {y. \y $ k - ?x' $ k\ < e * min d r / (2 * real CARD('m) ^ CARD('m))}" proof - have *: "\T (b *\<^sub>R axis k 1) \ (y - x)\ = b * \inv T y $ k - ?x' $ k\" for y proof - have "\T (b *\<^sub>R axis k 1) \ (y - x)\ = \(b *\<^sub>R axis k 1) \ inv T (y - x)\" by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v) also have "\ = b * \(axis k 1) \ inv T (y - x)\" using \b > 0\ by (simp add: abs_mult) also have "\ = b * \inv T y $ k - ?x' $ k\" using orthogonal_transformation_linear [OF invT] by (simp add: inner_axis' linear_diff) finally show ?thesis by simp qed show ?thesis using v b_def [symmetric] using \b > 0\ by (simp add: * bij_image_Collect_eq [OF \bij T\] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) qed show ?thesis using \b > 0\ by (simp add: image_Int \inj T\ 1 2 b_def [symmetric]) qed moreover have "?W' \ lmeasurable" by (auto intro: fmeasurable_Int_fmeasurable) ultimately have "measure lebesgue ?W = measure lebesgue ?W'" by (metis measure_orthogonal_image T) also have "\ \ measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" proof (rule measure_mono_fmeasurable) show "?W' \ cbox (?x' - ?v') (?x' + ?v')" apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) qed auto also have "\ \ e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" proof - have "cbox (?x' - ?v) (?x' + ?v) \ {}" using \r > 0\ \d > 0\ by (auto simp: interval_eq_empty_cart divide_less_0_iff) with \r > 0\ \d > 0\ \e > 0\ show ?thesis apply (simp add: content_cbox_if_cart mem_box_cart) apply (auto simp: prod_nonneg) apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) done qed also have "\ \ e/2 * measure lebesgue (cball ?x' (min d r))" proof (rule mult_left_mono [OF measure_mono_fmeasurable]) have *: "norm (?x' - y) \ min d r" if y: "\i. \?x' $ i - y $ i\ \ min d r / real CARD('m)" for y proof - have "norm (?x' - y) \ (\i\UNIV. \(?x' - y) $ i\)" by (rule norm_le_l1_cart) also have "\ \ real CARD('m) * (min d r / real CARD('m))" by (rule sum_bounded_above) (use y in auto) finally show ?thesis by simp qed show "cbox (?x' - ?v) (?x' + ?v) \ cball ?x' (min d r)" apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) by (simp add: abs_diff_le_iff abs_minus_commute) qed (use \e > 0\ in auto) also have "\ < e * content (cball ?x' (min d r))" - using \r > 0\ \d > 0\ \e > 0\ by auto + using \r > 0\ \d > 0\ \e > 0\ by (auto intro: content_cball_pos) also have "\ = e * content (ball x (min d r))" - using \r > 0\ \d > 0\ by (simp add: content_cball content_ball) + using \r > 0\ \d > 0\ content_ball_conv_unit_ball[of "min d r" "inv T x"] + content_ball_conv_unit_ball[of "min d r" x] + by (simp add: content_cball_conv_ball) finally show "measure lebesgue ?W < e * content (ball x (min d r))" . qed qed have *: "(\x. (x \ S) \ (x \ T \ x \ U)) \ (T - U) \ (U - T) \ S" for S T U :: "(real,'m) vec set" by blast have MN: "?M \ {x \ S. \v\0. ?\ x v}" proof (rule *) fix x assume x: "x \ {x \ S. \v\0. ?\ x v}" show "(x \ ?T) \ (x \ {x \ S. matrix (f' x) $ m $ n \ b})" proof (cases "x \ S") case True then have x: "\ ?\ x v" if "v \ 0" for v using x that by force show ?thesis proof (rule iffI; clarsimp) assume b: "\e>0. \d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" (is "\e>0. \d>0. \A. ?\ e d A") then have "\k. \d>0. \A. ?\ (1 / Suc k) d A" by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) then obtain \ A where \: "\k. \ k > 0" and Ab: "\k. A k $ m $ n < b" and A: "\k y. \y \ S; norm (y - x) < \ k\ \ norm (f y - f x - A k *v (y - x)) \ 1/(Suc k) * norm (y - x)" by metis have "\i j. \a. (\n. A n $ i $ j) \ a" proof (intro allI) fix i j have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n by (metis cart_eq_inner_axis matrix_vector_mul_component) let ?CA = "{x. Cauchy (\n. (A n) *v x)}" have "subspace ?CA" unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) with x [OF \d \ 0\] obtain \ where "\ > 0" and \: "\e. e > 0 \ \y \ S - {x}. norm (x - y) < e \ \ * norm (x - y) \ \d \ (y - x)\" by (fastforce simp: not_le Bex_def) obtain \ z where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \x: "\ \ x" and z: "(\n. (\ n - x) /\<^sub>R norm (\ n - x)) \ z" proof - have "\\. (\i. (\ i \ S - {x} \ \ * norm(\ i - x) \ \d \ (\ i - x)\ \ norm(\ i - x) < 1/Suc i) \ norm(\(Suc i) - x) < norm(\ i - x))" proof (rule dependent_nat_choice) show "\y. y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1 / Suc 0" using \ [of 1] by (auto simp: dist_norm norm_minus_commute) next fix y i assume "y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1/Suc i" then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" by auto then obtain y' where "y' \ S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))" "\ * norm (x - y') \ \d \ (y' - x)\" using \ by metis with \ show "\y'. (y' \ S - {x} \ \ * norm (y' - x) \ \d \ (y' - x)\ \ norm (y' - x) < 1/(Suc (Suc i))) \ norm (y' - x) < norm (y - x)" by (auto simp: dist_norm norm_minus_commute) qed then obtain \ where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \conv: "\i. norm(\ i - x) < 1/(Suc i)" by blast let ?f = "\i. (\ i - x) /\<^sub>R norm (\ i - x)" have "?f i \ sphere 0 1" for i using \Sx by auto then obtain l \ where "l \ sphere 0 1" "strict_mono \" and l: "(?f \ \) \ l" using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson show thesis proof show "(\ \ \) i \ S - {x}" "\ * norm ((\ \ \) i - x) \ \d \ ((\ \ \) i - x)\" for i using \Sx \le by auto have "\ \ x" proof (clarsimp simp add: LIMSEQ_def dist_norm) fix r :: "real" assume "r > 0" with real_arch_invD obtain no where "no \ 0" "real no > 1/r" by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) with \conv show "\no. \n\no. norm (\ n - x) < r" by (metis \r > 0\ add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) qed with \strict_mono \\ show "(\ \ \) \ x" by (metis LIMSEQ_subseq_LIMSEQ) show "(\n. ((\ \ \) n - x) /\<^sub>R norm ((\ \ \) n - x)) \ l" using l by (auto simp: o_def) qed qed have "isCont (\x. (\d \ x\ - \)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\y. \d \ ((\ y - x) /\<^sub>R norm (\ y - x))\ - \) \ \d \ z\ - \" by auto moreover have "\\<^sub>F i in sequentially. 0 \ \d \ ((\ i - x) /\<^sub>R norm (\ i - x))\ - \" proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce have "Cauchy (\n. (A n) *v z)" proof (clarsimp simp add: Cauchy_def) fix \ :: "real" assume "0 < \" then obtain N::nat where "N > 0" and N: "\/2 > 1/N" by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) show "\M. \m\M. \n\M. dist (A m *v z) (A n *v z) < \" proof (intro exI allI impI) fix i j assume ij: "N \ i" "N \ j" let ?V = "\i k. A i *v ((\ k - x) /\<^sub>R norm (\ k - x))" have "\\<^sub>F k in sequentially. dist (\ k) x < min (\ i) (\ j)" using \x [unfolded tendsto_iff] by (meson min_less_iff_conj \) then have even: "\\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \ 0" proof (rule eventually_mono, clarsimp) fix p assume p: "dist (\ p) x < \ i" "dist (\ p) x < \ j" let ?C = "\k. f (\ p) - f x - A k *v (\ p - x)" have "norm ((A i - A j) *v (\ p - x)) = norm (?C j - ?C i)" by (simp add: algebra_simps) also have "\ \ norm (?C j) + norm (?C i)" using norm_triangle_ineq4 by blast also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" apply (intro add_mono mult_right_mono) using ij \N > 0\ by (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . have "norm (?V i p - ?V j p) = norm ((A i - A j) *v ((\ p - x) /\<^sub>R norm (\ p - x)))" by (simp add: algebra_simps) also have "\ = norm ((A i - A j) *v (\ p - x)) / norm (\ p - x)" by (simp add: divide_inverse matrix_vector_mult_scaleR) also have "\ \ 2 / N" using no_le by (auto simp: field_split_simps) finally show "norm (?V i p - ?V j p) \ 2 / N" . qed have "isCont (\w. (norm(A i *v w - A j *v w) - 2 / N)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\w. norm (A i *v ((\ w - x) /\<^sub>R norm (\ w - x)) - A j *v ((\ w - x) /\<^sub>R norm (\ w - x))) - 2 / N) \ norm (A i *v z - A j *v z) - 2 / N" by auto have "dist (A i *v z) (A j *v z) \ 2 / N" using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) with N show "dist (A i *v z) (A j *v z) < \" by linarith qed qed then have "d \ z = 0" using CA_eq d orthogonal_def by auto then show False using \0 < \\ \\ \ \d \ z\\ by auto qed ultimately show ?thesis using dim_eq_full by fastforce qed finally have "?CA = UNIV" . then have "Cauchy (\n. (A n) *v axis j 1)" by auto then obtain L where "(\n. A n *v axis j 1) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) then have "(\x. (A x *v axis j 1) $ i) \ L $ i" by (rule tendsto_vec_nth) then show "\a. (\n. A n $ i $ j) \ a" by (force simp: vax) qed then obtain B where B: "\i j. (\n. A n $ i $ j) \ B $ i $ j" by (auto simp: lambda_skolem) have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) moreover interpret linear "f' x" by fact have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) show "linear ((*v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) have "((\y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" using tendsto_minus [OF lim_df] by (simp add: algebra_simps field_split_simps) then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \ 0) (at x within S)" proof (rule Lim_transform) have "((\y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \ 0) (at x within S)" proof (clarsimp simp add: Lim_within dist_norm) fix e :: "real" assume "e > 0" then obtain q::nat where "q \ 0" and qe2: "1/q < e/2" by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) let ?g = "\p. sum (\i. sum (\j. abs((A p - B)$i$j)) UNIV) UNIV" have "(\k. onorm (\y. (A k - B) *v y)) \ 0" proof (rule Lim_null_comparison) show "\\<^sub>F k in sequentially. norm (onorm (\y. (A k - B) *v y)) \ ?g k" proof (rule eventually_sequentiallyI) fix k :: "nat" assume "0 \ k" have "0 \ onorm ((*v) (A k - B))" using matrix_vector_mul_bounded_linear by (rule onorm_pos_le) then show "norm (onorm ((*v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed next show "?g \ 0" using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) qed with \e > 0\ obtain p where "\n. n \ p \ \onorm ((*v) (A n - B))\ < e/2" unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) then have pqe2: "\onorm ((*v) (A (p + q) - B))\ < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) using le_add1 by blast show "\d>0. \y\S. y \ x \ norm (y - x) < d \ inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" proof (intro exI, safe) show "0 < \(p + q)" by (simp add: \) next fix y assume y: "y \ S" "norm (y - x) < \(p + q)" and "y \ x" have *: "\norm(b - c) < e - d; norm(y - x - b) \ d\ \ norm(y - x - c) < e" for b c d e x and y:: "real^'n" using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" proof (rule *) show "norm (f y - f x - A (p + q) *v (y - x)) \ norm (y - x) / (Suc (p + q))" using A [OF y] by simp have "norm (A (p + q) *v (y - x) - B *v (y - x)) \ onorm(\x. (A(p + q) - B) *v x) * norm(y - x)" by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) also have "\ < (e/2) * norm (y - x)" using \y \ x\ pqe2 by auto also have "\ \ (e - 1 / (Suc (p + q))) * norm (y - x)" proof (rule mult_right_mono) have "1 / Suc (p + q) \ 1 / q" using \q \ 0\ by (auto simp: field_split_simps) also have "\ < e/2" using qe2 by auto finally show "e / 2 \ e - 1 / real (Suc (p + q))" by linarith qed auto finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" by (simp add: algebra_simps) qed then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" using \y \ x\ by (simp add: field_split_simps algebra_simps) qed qed then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = (*v) B" by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" by (simp add: B \f' x = (*v) B\) show "\\<^sub>F i in sequentially. A i $ m $ n \ b" by (simp add: Ab less_eq_real_def) qed auto next fix e :: "real" assume "x \ S" and b: "matrix (f' x) $ m $ n \ b" and "e > 0" then obtain d where "d>0" and d: "\y. y\S \ 0 < dist y x \ dist y x < d \ norm (f y - f x - f' x (y - x)) / (norm (y - x)) < e/2" using f [OF \x \ S\] by (simp add: Deriv.has_derivative_at_within Lim_within) (auto simp add: field_simps dest: spec [of _ "e/2"]) let ?A = "matrix(f' x) - (\ i j. if i = m \ j = n then e / 4 else 0)" obtain B where BRats: "\i j. B$i$j \ \" and Bo_e6: "onorm((*v) (?A - B)) < e/6" using matrix_rational_approximation \e > 0\ by (metis zero_less_divide_iff zero_less_numeral) show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" proof (intro exI conjI ballI allI impI) show "d>0" by (rule \d>0\) show "B $ m $ n < b" proof - have "\matrix ((*v) (?A - B)) $ m $ n\ \ onorm ((*v) (?A - B))" using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis then show ?thesis using b Bo_e6 by simp qed show "B $ i $ j \ \" for i j using BRats by auto show "norm (f y - f x - B *v (y - x)) \ e * norm (y - x)" if "y \ S" and y: "norm (y - x) < d" for y proof (cases "y = x") case True then show ?thesis by simp next case False have *: "norm(d' - d) \ e/2 \ norm(y - (x + d')) < e/2 \ norm(y - x - d) \ e" for d d' e and x y::"real^'n" using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp show ?thesis proof (rule *) have split246: "\norm y \ e / 6; norm(x - y) \ e / 4\ \ norm x \ e/2" if "e > 0" for e and x y :: "real^'n" using norm_triangle_le [of y "x-y" "e/2"] \e > 0\ by simp have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . then show "norm ((?A - B) *v (y - x)) \ e * norm (y - x) / 6" by (simp add: field_split_simps False) have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x))" by (simp add: algebra_simps) also have "\ = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" proof - have "(\j\UNIV. (if i = m \ j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i proof (cases "i=m") case True then show ?thesis by (auto simp: if_distrib [of "\z. z * _"] cong: if_cong) next case False then show ?thesis by (simp add: axis_def) qed then have "(\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" by (auto simp: vec_eq_iff matrix_vector_mult_def) then show ?thesis by metis qed also have "\ \ e * norm (y - x) / 4" using \e > 0\ apply (simp add: norm_mult abs_mult) by (metis component_le_norm_cart vector_minus_component) finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) qed finally show "norm (f' x (y - x) - B *v (y - x)) \ (e * norm (y - x)) / 2" . show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" using False d [OF \y \ S\] y by (simp add: dist_norm field_simps) qed qed qed qed qed auto qed show "negligible ?M" using negligible_subset [OF nN MN] . qed then show ?thesis by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) qed theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" unfolding det_def by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") case True then have "{x \ S. f x \ T} = {x. (if x \ S then f x else 0) \ T} \ S" "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T} \ -S" by auto then show ?thesis by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) next case False then have "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T}" by auto then show ?thesis by auto qed then show ?thesis unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] by blast qed lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" using sets_completionE [OF assms] by auto then show thesis by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ (\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue)" (is "?lhs \ _ \ ?rhs") unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] proof (intro iffI allI conjI impI, safe) fix V :: "'b set" assume *: "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "V \ sets borel" then have V: "V \ sets lebesgue" by simp have "{x \ S. f x \ V} = {x \ S. f x \ T \ V}" using fim by blast also have "{x \ S. f x \ T \ V} \ sets lebesgue" using T V * le_inf_iff by blast finally show "{x \ S. f x \ V} \ sets lebesgue" . next fix U :: "'b set" assume "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" "negligible U" "U \ T" then show "{x \ S. f x \ U} \ sets lebesgue" using negligible_imp_sets by blast next fix U :: "'b set" assume 1 [rule_format]: "(\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" using C \U \ T\ by (blast intro: 2) moreover have "{x \ S. f x \ C \ N} = {x \ S. f x \ C} \ {x \ S. f x \ N}" by auto ultimately show "{x \ S. f x \ U} \ sets lebesgue" using C by auto qed subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" show thesis proof have "- e \ x \ i" "x \ i \ e" if "t \ P" "norm (x - t) \ e" for x t using \a > 0\ that Basis_le_norm [of i "x-t"] P i by (auto simp: inner_commute algebra_simps) moreover have "- m \ x \ j" "x \ j \ m" if "norm x \ m" "t \ P" "norm (x - t) \ e" "j \ Basis" and "j \ i" for x t j using that Basis_le_norm [of j x] by auto ultimately show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ cbox (-?v) ?v" by (auto simp: mem_box) have *: "\k\Basis. - ?v \ k \ ?v \ k" using \0 \ m\ \0 \ e\ by (auto simp: inner_Basis) have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)" by (metis DIM_positive Suc_pred power_Suc) show "measure lebesgue (cbox (-?v) ?v) \ 2 * e * (2 * m) ^ (DIM('b) - 1)" using \i \ Basis\ by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) qed blast qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis have Tinv [simp]: "T (inv T x) = x" for x by (simp add: T orthogonal_transformation_surj surj_f_inv_f) obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ T-`P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x proof - have "a \ T x = 0" using P that by blast then show ?thesis by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) qed then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) show thesis proof show "T ` S \ lmeasurable" using S measurable_orthogonal_image T by blast have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t assume "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" using T linear_diff orthogonal_transformation_def by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" using \norm (x - t) \ e\ by linarith next show "inv T t \ T -` P" using \t \ P\ by force qed ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed then show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` S" using image_mono [OF subS] by (rule order_trans) show "measure lebesgue (T ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" using mS T by (simp add: S measure_orthogonal_image) qed qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" by (rule Sard_lemma0 [OF _ _ \0 \ m\ \0 \ e\]) show thesis proof show "(+)w ` S \ lmeasurable" by (metis measurable_translation S) show "{z. norm (z - w) \ m \ (\t\P. norm (z - w - t) \ e)} \ (+)w ` S" using subS by force show "measure lebesgue ((+)w ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" by (metis measure_translation mS) qed qed lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" and derS: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis proof (clarsimp simp add: negligible_outer_le) fix e :: "real" assume "e > 0" obtain c where csub: "S \ cbox (- (vec c)) (vec c)" and "c > 0" proof - obtain b where b: "\x. x \ S \ norm x \ b" using \bounded S\ by (auto simp: bounded_iff) show thesis proof have "- \b\ - 1 \ x $ i \ x $ i \ \b\ + 1" if "x \ S" for x i using component_le_norm_cart [of x i] b [OF that] by auto then show "S \ cbox (- vec (\b\ + 1)) (vec (\b\ + 1))" by (auto simp: mem_box_cart) qed auto qed then have box_cc: "box (- (vec c)) (vec c) \ {}" and cbox_cc: "cbox (- (vec c)) (vec c) \ {}" by (auto simp: interval_eq_empty_cart) obtain d where "d > 0" "d \ B" and d: "(d * 2) * (4 * B) ^ (?n - 1) \ e / (2*c) ^ ?m / ?m ^ ?m" apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) using \B > 0\ \c > 0\ \e > 0\ by (simp_all add: divide_simps min_mult_distrib_right) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. y \ S \ norm(y - x) < r \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)))" for x proof (cases "x \ S") case True then obtain r where "r > 0" and "\y. \y \ S; norm (y - x) < r\ \ norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using derS \d > 0\ by (force simp: has_derivative_within_alt) then show ?thesis by (rule_tac x="min r (1/2)" in exI) simp next case False then show ?thesis by (rule_tac x="1/2" in exI) simp qed then obtain r where r12: "\x. 0 < r x \ r x \ 1/2" and r: "\x y. \x \ S; y \ S; norm(y - x) < r x\ \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)" by metis then have ga: "gauge (\x. ball x (r x))" by (auto simp: gauge_def) obtain \ where \: "countable \" and sub_cc: "\\ \ cbox (- vec c) (vec c)" and cbox: "\K. K \ \ \ interior K \ {} \ (\u v. K = cbox u v)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ ball x (r x)" and close: "\u v. cbox u v \ \ \ \n. \i::'m. v $ i - u $ i = 2*c / 2^n" and covers: "S \ \\" apply (rule covering_lemma [OF csub box_cc ga]) apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) done let ?\ = "measure lebesgue" have "\T. T \ lmeasurable \ f ` (K \ S) \ T \ ?\ T \ e / (2*c) ^ ?m * ?\ K" if "K \ \" for K proof - obtain u v where uv: "K = cbox u v" using cbox \K \ \\ by blast then have uv_ne: "cbox u v \ {}" using cbox that by fastforce obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" using rank_dim_range [of "matrix (f' x)"] x rank[of x] by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) using \B > 0\ \d > 0\ by simp_all show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" unfolding uv proof (clarsimp simp: mult.assoc, intro conjI) fix y assume y: "y \ cbox u v" and "y \ S" then have "norm (y - x) < r x" by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) then have le_dyx: "norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B using norm_triangle_ineq2 [of "y - x" z] by auto show "norm (f y - f x) \ 2 * (B * norm (v - u))" proof (rule * [OF le_dyx]) have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" proof (rule mult_mono) show "onorm (f' x) \ B" using B x by blast qed (use \B > 0\ yx_le in auto) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" apply (rule_tac x="y-x" in exI) using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] by (meson order_trans real_mult_le_cancel_iff2) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" proof (rule order_trans [OF measT]) have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" using \c > 0\ apply (simp add: algebra_simps power_mult_distrib) by (metis Suc_pred power_Suc zero_less_card_finite) also have "\ \ (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" by (rule mult_right_mono [OF d]) auto also have "\ \ e / (2*c) ^ ?m * ?\ K" proof - have "u \ ball (x) (r x)" "v \ ball x (r x)" using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ moreover have "r x \ 1/2" using r12 by auto ultimately have "norm (v - u) \ 1" using norm_triangle_half_r [of x u 1 v] by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) then have "norm (v - u) ^ ?n \ norm (v - u) ^ ?m" by (simp add: power_decreasing [OF mlen]) also have "\ \ ?\ K * real (?m ^ ?m)" proof - obtain n where n: "\i. v$i - u$i = 2 * c / 2^n" using close [of u v] \K \ \\ uv by blast have "norm (v - u) ^ ?m \ (\i\UNIV. \(v - u) $ i\) ^ ?m" by (intro norm_le_l1_cart power_mono) auto also have "\ \ (\i\UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)" by (simp add: n field_simps \c > 0\ less_eq_real_def) also have "\ = ?\ K * real (?m ^ ?m)" by (simp add: uv uv_ne content_cbox_cart) finally show ?thesis . qed finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \ ?\ K" by (simp add: field_split_simps) show ?thesis using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \c > 0\ \e > 0\ by auto qed finally show "?DVU \ e / (2*c) ^ ?m * ?\ K" . qed qed (use T in auto) qed then obtain g where meas_g: "\K. K \ \ \ g K \ lmeasurable" and sub_g: "\K. K \ \ \ f ` (K \ S) \ g K" and le_g: "\K. K \ \ \ ?\ (g K) \ e / (2*c)^?m * ?\ K" by metis have le_e: "?\ (\i\\. g i) \ e" if "\ \ \" "finite \" for \ proof - have "?\ (\i\\. g i) \ (\i\\. ?\ (g i))" using meas_g \\ \ \\ by (auto intro: measure_UNION_le [OF \finite \\]) also have "\ \ (\K\\. e / (2*c) ^ ?m * ?\ K)" using \\ \ \\ sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) also have "\ = e / (2*c) ^ ?m * (\K\\. ?\ K)" by (simp add: sum_distrib_left) also have "\ \ e" proof - have "\ division_of \\" proof (rule division_ofI) show "K \ \\" "K \ {}" "\a b. K = cbox a b" if "K \ \" for K using \K \ \\ covered cbox \\ \ \\ by (auto simp: Union_upper) show "interior K \ interior L = {}" if "K \ \" and "L \ \" and "K \ L" for K L by (metis (mono_tags, lifting) \\ \ \\ pairwiseD djointish pairwise_subset that) qed (use that in auto) then have "sum ?\ \ \ ?\ (\\)" by (simp add: content_division) also have "\ \ ?\ (cbox (- vec c) (vec c) :: (real, 'm) vec set)" proof (rule measure_mono_fmeasurable) show "\\ \ cbox (- vec c) (vec c)" by (meson Sup_subset_mono sub_cc order_trans \\ \ \\) qed (use \\ division_of \\\ lmeasurable_division in auto) also have "\ = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" by simp also have "\ \ (2 ^ ?m * c ^ ?m)" using \c > 0\ by (simp add: content_cbox_if_cart) finally have "sum ?\ \ \ (2 ^ ?m * c ^ ?m)" . then show ?thesis using \e > 0\ \c > 0\ by (auto simp: field_split_simps mult_less_0_iff) qed finally show ?thesis . qed show "\T. f ` S \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "f ` S \ \ (g ` \)" using covers sub_g by force show "\ (g ` \) \ lmeasurable" by (rule fmeasurable_UN_bound [OF \countable \\ meas_g le_e]) show "?\ (\ (g ` \)) \ e" by (rule measure_UN_bound [OF \countable \\ meas_g le_e]) qed qed qed theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) then have eq: "S = (\n. ?U n)" by auto have "negligible (f ` ?U n)" for n proof (rule Sard_lemma2 [OF mlen]) show "0 < real n + 1" by auto show "bounded (?U n)" using bounded_iff by blast show "(f has_derivative f' x) (at x within ?U n)" if "x \ ?U n" for x using der that by (force intro: has_derivative_subset) qed (use rank in auto) then show ?thesis by (subst eq) (simp add: image_Union negligible_Union_nat) qed subsection\A one-way version of change-of-variables not assuming injectivity. \ lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" and nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and det_int_fg: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" and meas_gim: "\T. \T \ g ` S; T \ sets lebesgue\ \ {x \ S. g x \ T} \ sets lebesgue" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast have [simp]: "space (lebesgue_on S) = S" by (simp add: S) have gS_in_sets_leb: "g ` S \ sets lebesgue" apply (rule differentiable_image_in_sets_lebesgue) using der_g by (auto simp: S differentiable_def differentiable_on_def) obtain h where nonneg_h: "\n x. 0 \ h n x" and h_le_f: "\n x. x \ S \ h n (g x) \ f (g x)" and h_inc: "\n x. h n x \ h (Suc n) x" and h_meas: "\n. h n \ borel_measurable lebesgue" and fin_R: "\n. finite(range (h n))" and lim: "\x. x \ g ` S \ (\n. h n x) \ f x" proof - let ?f = "\x. if x \ g ` S then f x else 0" have "?f \ borel_measurable lebesgue \ (\x. 0 \ ?f x)" by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) then show ?thesis apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) apply (rename_tac h) by (rule_tac h=h in that) (auto split: if_split_asm) qed have h_lmeas: "{t. h n (g t) = y} \ S \ sets lebesgue" for y n proof - have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" by simp then have "((h n) -`{y} \ g ` S) \ sets (lebesgue_on (g ` S))" by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) then have "({u. h n u = y} \ g ` S) \ sets lebesgue" using gS_in_sets_leb by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) then have "{x \ S. g x \ ({u. h n u = y} \ g ` S)} \ sets lebesgue" using meas_gim[of "({u. h n u = y} \ g ` S)"] by force moreover have "{t. h n (g t) = y} \ S = {x \ S. g x \ ({u. h n u = y} \ g ` S)}" by blast ultimately show ?thesis by auto qed have hint: "h n integrable_on g ` S \ integral (g ` S) (h n) \ integral S (\x. ?D x * h n (g x))" (is "?INT \ ?lhs \ ?rhs") for n proof - let ?R = "range (h n)" have hn_eq: "h n = (\x. \y\?R. y * indicat_real {x. h n x = y} x)" by (simp add: indicator_def if_distrib fin_R cong: if_cong) have yind: "(\t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \ (integral (g ` S) (\t. y * indicator {x. h n x = y} t)) \ integral S (\t. \det (matrix (g' t))\ * y * indicator {x. h n x = y} (g t))" if y: "y \ ?R" for y::real proof (cases "y=0") case True then show ?thesis using gS_in_sets_leb integrable_0 by force next case False with that have "y > 0" using less_eq_real_def nonneg_h by fastforce have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) by (meson der_g IntD2 has_derivative_within_subset inf_le2) then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) show "(\x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" by (rule integrable_cmul) (use det_int_fg in auto) show "norm (if x \ {t. h n (g t) = y} then ?D x else 0) \ ?D x *\<^sub>R f (g x) /\<^sub>R y" if "x \ S" for x using nonneg_h [of n x] \y > 0\ nonneg_fg [of x] h_le_f [of x n] that by (auto simp: divide_simps mult_left_mono) qed (use S in auto) then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" apply (rule measurable_differentiable_image [OF h_lmeas]) apply (blast intro: has_derivative_within_subset [OF der_g]) apply (rule int_det) done moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) apply (blast intro: has_derivative_within_subset [OF der_g]) done ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integrable_on_cmult_iff integral_indicator) apply (simp add: indicator_def if_distrib cong: if_cong) done qed have hn_int: "h n integrable_on g ` S" apply (subst hn_eq) using yind by (force intro: integrable_sum [OF fin_R]) then show ?thesis proof have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" by (rule integral_sum [OF fin_R]) (use yind in blast) also have "\ \ (\y\range (h n). integral S (\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)))" using yind by (force intro: sum_mono) also have "\ = integral S (\u. \y\range (h n). \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u))" proof (rule integral_sum [OF fin_R, symmetric]) fix y assume y: "y \ ?R" with nonneg_h have "y \ 0" by auto show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. indicat_real {x. h n x = y} (g x)) \ borel_measurable (lebesgue_on S)" using h_lmeas S by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) then show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) \ borel_measurable (lebesgue_on S)" by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) next fix x assume "x \ S" have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" by (metis (full_types) \x \ S\ h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) qed also have "\ = integral S (\T. \det (matrix (g' T))\ * (\y\range (h n). y * indicat_real {x. h n x = y} (g T)))" by (simp add: sum_distrib_left mult.assoc) also have "\ = ?rhs" by (metis hn_eq) finally show "integral (g ` S) (h n) \ ?rhs" . qed qed have le: "integral S (\T. \det (matrix (g' T))\ * h n (g T)) \ ?b" for n proof (rule integral_le) show "(\T. \det (matrix (g' T))\ * h n (g T)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\T. \det (matrix (g' T))\ *\<^sub>R h n (g T)) \ borel_measurable (lebesgue_on S)" proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \S \ sets lebesgue\) have eq: "{x \ S. f x \ a} = (\b \ (f ` S) \ atMost a. {x. f x = b} \ S)" for f and a::real by auto have "finite ((\x. h n (g x)) ` S \ {..a})" for a by (force intro: finite_subset [OF _ fin_R]) with h_lmeas [of n] show "(\x. h n (g x)) \ borel_measurable (lebesgue_on S)" apply (simp add: borel_measurable_vimage_halfspace_component_le \S \ sets lebesgue\ sets_restrict_space_iff eq) by (metis (mono_tags) SUP_inf sets.finite_UN) qed (use der_g in blast) then show "(\T. \det (matrix (g' T))\ * h n (g T)) \ borel_measurable (lebesgue_on S)" by simp show "norm (?D x * h n (g x)) \ ?D x *\<^sub>R f (g x)" if "x \ S" for x by (simp add: h_le_f mult_left_mono nonneg_h that) qed (use S det_int_fg in auto) show "?D x * h n (g x) \ ?D x * f (g x)" if "x \ S" for x by (simp add: \x \ S\ h_le_f mult_left_mono) show "(\x. ?D x * f (g x)) integrable_on S" using det_int_fg by blast qed have "f integrable_on g ` S \ (\k. integral (g ` S) (h k)) \ integral (g ` S) f" proof (rule monotone_convergence_increasing) have "\integral (g ` S) (h n)\ \ integral S (\x. ?D x * f (g x))" for n proof - have "\integral (g ` S) (h n)\ = integral (g ` S) (h n)" using hint by (simp add: integral_nonneg nonneg_h) also have "\ \ integral S (\x. ?D x * f (g x))" using hint le by (meson order_trans) finally show ?thesis . qed then show "bounded (range (\k. integral (g ` S) (h k)))" by (force simp: bounded_iff) qed (use h_inc lim hint in auto) moreover have "integral (g ` S) (h n) \ integral S (\x. ?D x * f (g x))" for n using hint by (blast intro: le order_trans) ultimately show ?thesis by (auto intro: Lim_bounded) qed lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" by (metis (mono_tags, lifting) der_g has_derivative_within_subset mem_Collect_eq subset_iff) have "(\x. if x \ S then \?D x\ * f (g x) else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV intS) then have Df_borel: "(\x. if x \ S then \?D x\ * f (g x) else 0) \ borel_measurable lebesgue" using integrable_imp_measurable lebesgue_on_UNIV_eq by force have S': "S' \ sets lebesgue" proof - from Df_borel borel_measurable_vimage_open [of _ UNIV] have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ T} \ sets lebesgue" if "open T" for T using that unfolding lebesgue_on_UNIV_eq by (fastforce simp add: dest!: spec) then have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ -{0}} \ sets lebesgue" using open_Compl by blast then show ?thesis by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) qed then have gS': "g ` S' \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "g differentiable_on S'" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto have f: "f \ borel_measurable (lebesgue_on (g ` S'))" proof (clarsimp simp add: borel_measurable_vimage_open) fix T :: "real set" assume "open T" have "{x \ g ` S'. f x \ T} = g ` {x \ S'. f(g x) \ T}" by blast moreover have "g ` {x \ S'. f(g x) \ T} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) let ?h = "\x. \?D x\ * f (g x) /\<^sub>R \?D x\" have "(\x. if x \ S' then \?D x\ * f (g x) else 0) = (\x. if x \ S then \?D x\ * f (g x) else 0)" by (auto simp: S'_def) also have "\ \ borel_measurable lebesgue" by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) have "?h \ borel_measurable (lebesgue_on S')" by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" by (metis (no_types, lifting) measurable_lebesgue_cong) then show "{x \ S'. f (g x) \ T} \ sets lebesgue" by (simp add: \S' \ sets lebesgue\ \open T\ borel_measurable_vimage_open sets_restrict_space_iff) show "g differentiable_on {x \ S'. f (g x) \ T}" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto ultimately have "{x \ g ` S'. f x \ T} \ sets lebesgue" by metis then show "{x \ g ` S'. f x \ T} \ sets (lebesgue_on (g ` S'))" by (simp add: \g ` S' \ sets lebesgue\ sets_restrict_space_iff) qed have intS': "(\x. \?D x\ * f (g x)) integrable_on S'" using intS by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) have lebS': "{x \ S'. g x \ T} \ sets lebesgue" if "T \ g ` S'" "T \ sets lebesgue" for T proof - have "g \ borel_measurable (lebesgue_on S')" using der_gS' has_derivative_continuous_on S' by (blast intro: continuous_imp_measurable_on_sets_lebesgue) moreover have "{x \ S'. g x \ U} \ sets lebesgue" if "negligible U" "U \ g ` S'" for U proof (intro negligible_imp_sets negligible_differentiable_vimage that) fix x assume x: "x \ S'" then have "linear (g' x)" using der_gS' has_derivative_linear by blast with x show "inj (g' x)" by (auto simp: S'_def det_nz_iff_inj) qed (use der_gS' in auto) ultimately show ?thesis using double_lebesgue_sets [OF S' gS' order_refl] that by blast qed have int_gS': "f integrable_on g ` S' \ integral (g ` S') f \ integral S' (\x. \?D x\ * f(g x))" using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast have "negligible (g ` {x \ S. det(matrix(g' x)) = 0})" proof (rule baby_Sard, simp_all) fix x assume x: "x \ S \ det (matrix (g' x)) = 0" then show "(g has_derivative g' x) (at x within {x \ S. det (matrix (g' x)) = 0})" by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI) then show "rank (matrix (g' x)) < CARD('n)" using det_nz_iff_inj matrix_vector_mul_linear x by (fastforce simp add: less_rank_noninjective) qed then have negg: "negligible (g ` S - g ` {x \ S. ?D x \ 0})" by (rule negligible_subset) (auto simp: S'_def) have null: "g ` {x \ S. ?D x \ 0} - g ` S = {}" by (auto simp: S'_def) let ?F = "{x \ S. f (g x) \ 0}" have eq: "g ` S' = g ` ?F \ g ` {x \ S. ?D x \ 0}" by (auto simp: S'_def image_iff) show ?thesis proof have "((\x. if x \ g ` ?F then f x else 0) integrable_on g ` {x \ S. ?D x \ 0})" using int_gS' eq integrable_restrict_Int [where f=f] by simp then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) using negg null by auto have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" by (auto simp: image_iff intro!: integral_cong) also have "\ = integral (g ` S') f" using eq integral_restrict_Int by simp also have "\ \ integral S' (\x. \?D x\ * f(g x))" by (metis int_gS') also have "\ \ ?b" by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) finally show "integral (g ` S) f \ ?b" . qed qed lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" proof - let ?D = "\x. \det (matrix (g' x))\ * f (g x)" let ?N = "{x \ S. f (g x) < 0}" and ?P = "{x \ S. f (g x) > 0}" have eq: "{x. (if x \ S then ?D x else 0) > 0} = {x \ S. ?D x > 0}" "{x. (if x \ S then ?D x else 0) < 0} = {x \ S. ?D x < 0}" by auto have "?D integrable_on S" using intS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?D x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?D x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have Dlt: "{x \ S. ?D x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have Dgt: "{x \ S. ?D x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have dfgbm: "?D \ borel_measurable (lebesgue_on S)" using intS absolutely_integrable_on_def integrable_imp_measurable by blast have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \ ?N" for x using der_g has_derivative_within_subset that by force have "(\x. - f x) integrable_on g ` ?N \ integral (g ` ?N) (\x. - f x) \ integral ?N (\x. \det (matrix (g' x))\ * - f (g x))" proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) have 1: "?D integrable_on {x \ S. ?D x < 0}" using Dlt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) have "uminus \ (\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) then show "(\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: integrable_neg_iff o_def) qed auto then have "f integrable_on g ` ?N" by (simp add: integrable_neg_iff) moreover have "g ` ?N = {y \ g ` S. f y < 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y < 0}" by simp then have N: "f absolutely_integrable_on {y \ g ` S. f y < 0}" by (rule absolutely_integrable_absolutely_integrable_ubound) auto have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_within_subset that by force have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) have "?D integrable_on {x \ S. 0 < ?D x}" using Dgt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) then show "?D integrable_on ?P" apply (rule integrable_spike_set) by (auto simp: zero_less_mult_iff empty_imp_negligible) qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y > 0}" by simp then have P: "f absolutely_integrable_on {y \ g ` S. f y > 0}" by (rule absolutely_integrable_absolutely_integrable_lbound) auto have "(\x. if x \ g ` S \ f x < 0 \ x \ g ` S \ 0 < f x then f x else 0) = (\x. if x \ g ` S then f x else 0)" by auto then show ?thesis using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] by simp qed proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using absolutely_integrable_component [OF intS] by auto proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" using integral_on_image_ubound_nonneg [OF assms] by simp subsection\Change-of-variables theorem\ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and f0: "\y. y \ T \ 0 \ f y" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "f integrable_on T \ (integral T f) \ b \ (\x. \det (matrix (g' x))\ * f(g x)) integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) \ b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ have gS: "g differentiable_on S" by (meson der_g differentiable_def differentiable_on_def) let ?D = "\x. \det (matrix (g' x))\ * f (g x)" show ?thesis proof assume ?lhs then have fT: "f integrable_on T" and intf: "integral T f \ b" by blast+ show ?rhs proof let ?fgh = "\x. \det (matrix (h' x))\ * (\det (matrix (g' (h x)))\ * f (g (h x)))" have ddf: "?fgh x = f x" if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis by (simp add: gh that) qed have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" proof (subst absolutely_integrable_on_iff_nonneg) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (simp add: zero_le_mult_iff f0 gh) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp have "integral S (\x. ?D x) \ integral T (\x. ?fgh x)" unfolding Seq proof (rule integral_on_image_ubound) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) also have "\ \ b" by (rule intf) finally show "integral S (\x. ?D x) \ b" . qed next assume R: ?rhs then have "f integrable_on g ` S" using der_g f0 hg integral_on_image_ubound_nonneg by blast moreover have "integral (g ` S) f \ integral S (\x. ?D x)" by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) ultimately show ?lhs using R by (simp add: Teq) qed qed lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" and "\y. y \ T \ 0 \ f y" and "\x. x \ S \ g x \ T \ h(g x) = x" and "\y. y \ T \ h y \ S \ g(h y) = y" and "\y. y \ T \ h' y \ g'(h y) = id" shows "((\x. \det (matrix (g' x))\ * f(g x)) has_integral b) S \ (f has_integral b) T" using cov_invertible_nonneg_le [OF assms] by (simp add: has_integral_iff) (meson eq_iff) lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ let ?DP = "\x. \det (matrix (g' x))\ * f(g x)" and ?DN = "\x. \det (matrix (g' x))\ * -f(g x)" have "+": "(?DP has_integral b) {x \ S. f (g x) > 0} \ (f has_integral b) {y \ T. f y > 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. f (g x)) -` Y \ {x \ S. f (g x) > 0} = ((\x. f (g x)) -` Y \ S) \ {x \ S. f (g x) > 0}" for Y by auto show "(g has_derivative g' x) (at x within {x \ S. f (g x) > 0})" if "x \ {x \ S. f (g x) > 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y > 0})" if "y \ {y \ T. f y > 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) have "-": "(?DN has_integral b) {x \ S. f (g x) < 0} \ ((\x. - f x) has_integral b) {y \ T. f y < 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. - f (g x)) -` y \ {x \ S. f (g x) < 0} = ((\x. f (g x)) -` uminus ` y \ S) \ {x \ S. f (g x) < 0}" for y using image_iff by fastforce show "(g has_derivative g' x) (at x within {x \ S. f (g x) < 0})" if "x \ {x \ S. f (g x) < 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y < 0})" if "y \ {y \ T. f y < 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) show ?thesis proof assume LHS: ?lhs have eq: "{x. (if x \ S then ?DP x else 0) > 0} = {x \ S. ?DP x > 0}" "{x. (if x \ S then ?DP x else 0) < 0} = {x \ S. ?DP x < 0}" by auto have "?DP integrable_on S" using LHS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?DP x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?DP x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have SN: "{x \ S. ?DP x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have SP: "{x \ S. ?DP x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have "?DP absolutely_integrable_on {x \ S. ?DP x > 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) then have aP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) have "?DP absolutely_integrable_on {x \ S. ?DP x < 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) then have aN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) have fN: "f integrable_on {y \ T. f y < 0}" "integral {y \ T. f y < 0} f = integral {x \ S. f (g x) < 0} ?DP" using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) using fN by (auto simp: integrable_neg_iff) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" apply (rule absolutely_integrable_integrable_bound [where g = f]) using fP by auto have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs proof have eq: "((if x \ T \ f x < 0 \ x \ T \ 0 < f x then 1 else 0) * f x) = (if x \ T then 1 else 0) * f x" for x by auto show "f absolutely_integrable_on T" using fa by (simp add: indicator_def set_integrable_def eq) have [simp]: "{y \ T. f y < 0} \ {y \ T. 0 < f y} = {}" for T and f :: "(real^'n::_) \ real" by auto have "integral T f = integral ({y \ T. f y < 0} \ {y \ T. f y > 0}) f" by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) also have "\ = integral {y \ T. f y < 0} f + integral {y \ T. f y > 0} f" using fN fP by simp also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" by (simp add: fN fP) also have "\ = integral ({x \ S. f (g x) < 0} \ {x \ S. 0 < f (g x)}) ?DP" using aP aN by (simp add: set_lebesgue_integral_eq_integral) also have "\ = integral S ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using LHS by simp finally show "integral T f = b" . qed next assume RHS: ?rhs have eq: "{x. (if x \ T then f x else 0) > 0} = {x \ T. f x > 0}" "{x. (if x \ T then f x else 0) < 0} = {x \ T. f x < 0}" by auto have "f integrable_on T" using RHS absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have TN: "{x \ T. f x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have TP: "{x \ T. f x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" using absolutely_integrable_on_def aint by blast then have DN: "(?DN has_integral integral {y \ T. f y < 0} (\x. - f x)) {x \ S. f (g x) < 0}" using "-" [of "integral {y \ T. f y < 0} (\x. - f x)"] by (simp add: has_integral_neg_iff integrable_integral) have aDN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ have fP: "f integrable_on {v \ T. f v > 0}" using absolutely_integrable_on_def aint by blast then have DP: "(?DP has_integral integral {y \ T. f y > 0} f) {x \ S. f (g x) > 0}" using "+" [of "integral {y \ T. f y > 0} f"] by (simp add: has_integral_neg_iff integrable_integral) have aDP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) using DP hg by (fastforce simp: integrable_neg_iff)+ have eq: "(if x \ S then 1 else 0) * ?DP x = (if x \ S \ f (g x) < 0 \ x \ S \ f (g x) > 0 then 1 else 0) * ?DP x" for x by force have "?DP absolutely_integrable_on ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0})" by (rule absolutely_integrable_Un [OF aDN aDP]) then show I: "?DP absolutely_integrable_on S" by (simp add: indicator_def eq set_integrable_def) have [simp]: "{y \ S. f y < 0} \ {y \ S. 0 < f y} = {}" for S and f :: "(real^'n::_) \ real" by auto have "integral S ?DP = integral ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0}) ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) also have "\ = - integral {y \ T. f y < 0} (\x. - f x) + integral {y \ T. f y > 0} f" using DN DP by (auto simp: has_integral_iff) also have "\ = integral ({x \ T. f x < 0} \ {x \ T. 0 < f x}) f" by (simp add: fN fP) also have "\ = integral T f" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using intT by simp finally show "integral S ?DP = b" . qed qed qed lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have "((\x. \det (matrix (g' x))\ * f(g x) $ i) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * (f(g x) $ i)) = b $ i) \ ((\x. f x $ i) absolutely_integrable_on T \ integral T (\x. f x $ i) = b $ i)" for i by (rule cov_invertible_real [OF der_g der_h hg gh id]) then have "?D absolutely_integrable_on S \ (?D has_integral b) S \ f absolutely_integrable_on T \ (f has_integral b) T" unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] has_integral_iff set_lebesgue_integral_eq_integral) then show ?thesis using absolutely_integrable_on_def by blast qed lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "\x. \h'. x \ S \ (g has_derivative g' x) (at x within S) \ linear h' \ g' x \ h' = id \ h' \ g' x = id" using der_g matrix_invertible has_derivative_linear by blast then obtain h' where h': "\x. x \ S \ (g has_derivative g' x) (at x within S) \ linear (h' x) \ g' x \ (h' x) = id \ (h' x) \ g' x = id" by metis show ?thesis proof (rule cv_inv_version3) show "\y. y \ g ` S \ (h has_derivative h' (h y)) (at y within g ` S)" using h' hg by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) qed (use h' hg in auto) qed theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" and conth: "continuous_on (g ` S) h" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" proof (rule cv_inv_version4) show "(g has_derivative g' x) (at x within ?S) \ invertible (matrix (g' x))" if "x \ ?S" for x using der_g that has_derivative_within_subset that by fastforce show "continuous_on (g ` ?S) h \ h (g x) = x" if "x \ ?S" for x using that continuous_on_subset [OF conth] by (simp add: hg image_mono) qed have "(g has_derivative g' x) (at x within {x \ S. rank (matrix (g' x)) < CARD('m)})" if "x \ S" for x by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI that) then have "negligible (g ` {x \ S. \ invertible (matrix (g' x))})" by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) then have neg: "negligible {x \ g ` S. x \ g ` ?S \ f x \ 0}" by (auto intro: negligible_subset) have [simp]: "{x \ g ` ?S. x \ g ` S \ f x \ 0} = {}" by auto have "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" apply (intro conj_cong absolutely_integrable_spike_set_eq) apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) done moreover have "f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) ultimately show ?thesis using "*" by blast qed theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b)" proof - obtain h where hg: "\x. x \ S \ h(g x) = x" using inj by (metis the_inv_into_f_f) have conth: "continuous_on (g ` S) h" by (metis \compact S\ continuous_on_inv der_g has_derivative_continuous_on hg) show ?thesis by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) qed lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" and inj: "inj_on g (\n. F n)" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on (\n. F n) \ integral (\n. F n) (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` (\n. F n)) \ integral (g ` (\n. F n)) f = b)" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" let ?U = "\n. \m\n. F m" let ?lift = "vec::real\real^1" have F_leb: "F m \ sets lebesgue" for m by (simp add: compact borel_compact) have iff: "(\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \ integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) = b \ f absolutely_integrable_on (g ` (?U n)) \ integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \ real^'k" proof (rule has_absolute_integral_change_of_variables_compact) show "compact (?U n)" by (simp add: compact compact_UN) show "(g has_derivative g' x) (at x within (?U n))" if "x \ ?U n" for x using that by (blast intro!: has_derivative_within_subset [OF der_g]) show "inj_on g (?U n)" using inj by (auto simp: inj_on_def) qed show ?thesis unfolding image_UN proof safe assume DS: "?D absolutely_integrable_on (\n. F n)" and b: "b = integral (\n. F n) ?D" have DU: "\n. ?D absolutely_integrable_on (?U n)" "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" using integral_countable_UN [OF DS F_leb] by auto with iff have fag: "f absolutely_integrable_on g ` (?U n)" and fg_int: "integral (\m\n. g ` F m) f = integral (?U n) ?D" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\m. g ` F m) then norm(f x) else 0" have "(\x. if x \ (\m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ (\m\k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using fag by (simp add: image_UN) let ?nf = "\n x. if x \ (\m\n. g ` F m) then norm(f x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nf k integrable_on UNIV" for k using fag unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n have "(norm \ ?D) absolutely_integrable_on ?U n" by (intro absolutely_integrable_norm DU) then have "integral (g ` ?U n) (norm \ f) = integral (?U n) (norm \ ?D)" using iff [of n "vec \ norm \ f" "integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R (?lift \ norm \ f) (g x))"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) } moreover have "bounded (range (\k. integral (?U k) (norm \ ?D)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (?U k) (norm \ ?D)) \ integral (\n. F n) (norm \ ?D)" unfolding integral_restrict_UNIV [of _ "norm \ ?D", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ \ (F ` {..k}) then (norm \ ?D) x else 0) integrable_on UNIV" "(\x. if x \ (\n. F n) then (norm \ ?D) x else 0) integrable_on UNIV" using DU(1) DS unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nf k)))" by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) next show "(\k. if x \ (\m\k. g ` F m) then norm (f x) else 0) \ (if x \ (\m. g ` F m) then norm (f x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ (\m\k. g ` F m) then f x else 0) \ (if x \ (\m. g ` F m) then f x else 0)" for x proof clarsimp fix m y assume "y \ F m" show "(\k. if \x\{..k}. g y \ g ` F x then f (g y) else 0) \ f (g y)" using \y \ F m\ by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) qed qed auto then show fai: "f absolutely_integrable_on (\m. g ` F m)" using absolutely_integrable_restrict_UNIV by blast show "integral ((\x. g ` F x)) f = integral (\n. F n) ?D" proof (rule LIMSEQ_unique) show "(\n. integral (?U n) ?D) \ integral (\x. g ` F x) f" unfolding fg_int [symmetric] proof (rule integral_countable_UN [OF fai]) show "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed next show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" by (rule DU) qed next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" have gF_leb: "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" using der_g unfolding differentiable_def differentiable_on_def by (meson Sup_upper UNIV_I UnionI has_derivative_within_subset image_eqI) qed auto have fgU: "\n. f absolutely_integrable_on (\m\n. g ` F m)" "(\n. integral (\m\n. g ` F m) f) \ integral (\m. g ` F m) f" using integral_countable_UN [OF fs gF_leb] by auto with iff have DUn: "?D absolutely_integrable_on ?U n" and D_int: "integral (?U n) ?D = integral (\m\n. g ` F m) f" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\n. F n) then norm(?D x) else 0" have "(\x. if x \ (\n. F n) then ?D x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using DUn by simp let ?nD = "\n x. if x \ ?U n then norm(?D x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nD k integrable_on UNIV" for k using DUn unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" apply (rule absolutely_integrable_norm) using fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) } moreover have "bounded (range (\k. integral (g ` ?U k) (norm \ f)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (g ` ?U k) (norm \ f)) \ integral (g ` (\n. F n)) (norm \ f)" unfolding integral_restrict_UNIV [of _ "norm \ f", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ g ` ?U k then (norm \ f) x else 0) integrable_on UNIV" "(\x. if x \ g ` (\n. F n) then (norm \ f) x else 0) integrable_on UNIV" using fgU fs unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by (auto simp: image_UN) qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nD k)))" unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp next show "(\k. if x \ ?U k then norm (?D x) else 0) \ (if x \ (\n. F n) then norm (?D x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ ?U k then ?D x else 0) \ (if x \ (\n. F n) then ?D x else 0)" for x proof clarsimp fix n assume "x \ F n" show "(\m. if \j\{..m}. x \ F j then ?D x else 0) \ ?D x" using \x \ F n\ by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) qed qed auto then show Dai: "?D absolutely_integrable_on (\n. F n)" unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) qed qed qed theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - obtain C N where "fsigma C" and N: "N \ null_sets lebesgue" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" where F: "range F \ Collect compact" and Ceq: "C = Union(range F)" using fsigma_Union_compact by metis have "negligible N" using N by (simp add: negligible_iff_null_sets) let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" have "?D absolutely_integrable_on C \ integral C ?D = b \ f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b" unfolding Ceq proof (rule has_absolute_integral_change_of_variables_compact_family) fix n x assume "x \ \(F ` UNIV)" then show "(g has_derivative g' x) (at x within \(F ` UNIV))" using Ceq \C \ N = S\ der_g has_derivative_within_subset by blast next have "\(F ` UNIV) \ S" using Ceq \C \ N = S\ by blast then show "inj_on g (\(F ` UNIV))" using inj by (meson inj_on_subset) qed (use F in auto) moreover have "?D absolutely_integrable_on C \ integral C ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" proof (rule conj_cong) have neg: "negligible {x \ C - S. ?D x \ 0}" "negligible {x \ S - C. ?D x \ 0}" using CNS by (blast intro: negligible_subset [OF \negligible N\])+ then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" by (rule absolutely_integrable_spike_set_eq) show "(integral C ?D = b) \ (integral S ?D = b)" using integral_spike_set [OF neg] by simp qed moreover have "f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (rule conj_cong) have "g differentiable_on N" by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) with \negligible N\ have neg_gN: "negligible (g ` N)" by (blast intro: negligible_differentiable_image_negligible) have neg: "negligible {x \ g ` C - g ` S. f x \ 0}" "negligible {x \ g ` S - g ` C. f x \ 0}" using CNS by (blast intro: negligible_subset [OF neg_gN])+ then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" by (rule absolutely_integrable_spike_set_eq) show "(integral (g ` C) f = b) \ (integral (g ` S) f = b)" using integral_spike_set [OF neg] by simp qed ultimately show ?thesis by simp qed corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using assms has_absolute_integral_change_of_variables by blast corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x))" using has_absolute_integral_change_of_variables [OF S der_g inj] disj by blast lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - let ?lift = "vec :: real \ real^1" let ?drop = "(\x::real^1. x $ 1)" have S': "?lift ` S \ sets lebesgue" by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) have "((\x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" if "z \ S" for z using der_g [OF that] by (simp add: has_vector_derivative_def has_derivative_vector_1) then have der': "\x. x \ ?lift ` S \ (?lift \ g \ ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" by (auto simp: o_def) have inj': "inj_on (vec \ g \ ?drop) (vec ` S)" using inj by (simp add: inj_on_def) let ?fg = "\x. \g' x\ *\<^sub>R f(g x)" have "((\x. ?fg x $ i) absolutely_integrable_on S \ ((\x. ?fg x $ i) has_integral b $ i) S \ (\x. f x $ i) absolutely_integrable_on g ` S \ ((\x. f x $ i) has_integral b $ i) (g ` S))" for i using has_absolute_integral_change_of_variables [OF S' der' inj', of "\x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) then have "?fg absolutely_integrable_on S \ (?fg has_integral b) S \ f absolutely_integrable_on (g ` S) \ (f has_integral b) (g ` S)" unfolding has_integral_componentwise_iff [where y=b] absolutely_integrable_componentwise_iff [where f=f] absolutely_integrable_componentwise_iff [where f = ?fg] by (force simp: Basis_vec_def cart_eq_inner_axis) then show ?thesis using absolutely_integrable_on_def by blast qed corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S)" using has_absolute_integral_change_of_variables_1 [OF assms] by auto subsection\Change of variables for integrals: special case of linear function\ lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix g)\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (cases "det(matrix g) = 0") case True then have "negligible(g ` S)" using assms det_nz_iff_inj negligible_linear_singular_image by blast with True show ?thesis by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) next case False then obtain h where h: "\x. x \ S \ h (g x) = x" "linear h" using assms det_nz_iff_inj linear_injective_isomorphism by metis show ?thesis proof (rule has_absolute_integral_change_of_variables_invertible) show "(g has_derivative g) (at x within S)" for x by (simp add: assms linear_imp_has_derivative) show "continuous_on (g ` S) h" using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast qed (use h in auto) qed lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S \ det(matrix g) = 0" unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover have ?thesis if "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" using has_absolute_integral_change_of_variables_linear [OF \linear g\] that by (auto simp: o_def) ultimately show ?thesis using absolutely_integrable_change_of_variables_linear [OF \linear g\] by blast qed subsection\Change of variable for measure\ lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) integrable_on S" using has_measure_differentiable_image [OF assms] by blast lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" using measurable_differentiable_image_eq [OF assms] by (simp only: absolutely_integrable_on_iff_nonneg) lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" using measurable_differentiable_image_eq [OF S der_f inj] assms has_measure_differentiable_image by blast end diff --git a/src/HOL/Analysis/Elementary_Metric_Spaces.thy b/src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy @@ -1,3258 +1,3313 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) chapter \Functional Analysis\ theory Elementary_Metric_Spaces imports Abstract_Topology_2 Metric_Arith begin section \Elementary Metric Spaces\ subsection \Open and closed balls\ definition\<^marker>\tag important\ ball :: "'a::metric_space \ real \ 'a set" where "ball x e = {y. dist x y < e}" definition\<^marker>\tag important\ cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" definition\<^marker>\tag important\ sphere :: "'a::metric_space \ real \ 'a set" where "sphere x e = {y. dist x y = e}" lemma mem_ball [simp, metric_unfold]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) lemma mem_cball [simp, metric_unfold]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) lemma mem_sphere [simp]: "y \ sphere x e \ dist x y = e" by (simp add: sphere_def) lemma ball_trivial [simp]: "ball x 0 = {}" by (simp add: ball_def) lemma cball_trivial [simp]: "cball x 0 = {x}" by (simp add: cball_def) lemma sphere_trivial [simp]: "sphere x 0 = {x}" by (simp add: sphere_def) lemma disjoint_ballI: "dist x y \ r+s \ ball x r \ ball y s = {}" using dist_triangle_less_add not_le by fastforce lemma disjoint_cballI: "dist x y > r + s \ cball x r \ cball y s = {}" by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball) lemma sphere_empty [simp]: "r < 0 \ sphere a r = {}" for a :: "'a::metric_space" by auto lemma centre_in_ball [simp]: "x \ ball x e \ 0 < e" by simp lemma centre_in_cball [simp]: "x \ cball x e \ 0 \ e" by simp lemma ball_subset_cball [simp, intro]: "ball x e \ cball x e" by (simp add: subset_eq) lemma mem_ball_imp_mem_cball: "x \ ball y e \ x \ cball y e" by (auto) lemma sphere_cball [simp,intro]: "sphere z r \ cball z r" by force lemma cball_diff_sphere: "cball a r - sphere a r = ball a r" by auto lemma subset_ball[intro]: "d \ e \ ball x d \ ball x e" by (simp add: subset_eq) lemma subset_cball[intro]: "d \ e \ cball x d \ cball x e" by (simp add: subset_eq) lemma mem_ball_leI: "x \ ball y e \ e \ f \ x \ ball y f" by (auto) lemma mem_cball_leI: "x \ cball y e \ e \ f \ x \ cball y f" by (auto) lemma cball_trans: "y \ cball z b \ x \ cball y a \ x \ cball z (b + a)" by metric lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" by (simp add: set_eq_iff) arith lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" by (simp add: set_eq_iff) lemma cball_max_Un: "cball a (max r s) = cball a r \ cball a s" by (simp add: set_eq_iff) arith lemma cball_min_Int: "cball a (min r s) = cball a r \ cball a s" by (simp add: set_eq_iff) lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" by (auto simp: cball_def ball_def dist_commute) lemma open_ball [intro, simp]: "open (ball x e)" proof - have "open (dist x -` {.. (\x\S. \e>0. ball x e \ S)" by (simp add: open_dist subset_eq mem_ball Ball_def dist_commute) lemma openI [intro?]: "(\x. x\S \ \e>0. ball x e \ S) \ open S" by (auto simp: open_contains_ball) lemma openE[elim?]: assumes "open S" "x\S" obtains e where "e>0" "ball x e \ S" using assms unfolding open_contains_ball by auto lemma open_contains_ball_eq: "open S \ x\S \ (\e>0. ball x e \ S)" by (metis open_contains_ball subset_eq centre_in_ball) lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" unfolding mem_ball set_eq_iff by (simp add: not_less) metric lemma ball_empty: "e \ 0 \ ball x e = {}" by simp lemma closed_cball [iff]: "closed (cball x e)" proof - have "closed (dist x -` {..e})" by (intro closed_vimage closed_atMost continuous_intros) also have "dist x -` {..e} = cball x e" by auto finally show ?thesis . qed lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" proof - { fix x and e::real assume "x\S" "e>0" "ball x e \ S" then have "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) } moreover { fix x and e::real assume "x\S" "e>0" "cball x e \ S" then have "\d>0. ball x d \ S" unfolding subset_eq apply (rule_tac x="e/2" in exI, auto) done } ultimately show ?thesis unfolding open_contains_ball by auto qed lemma open_contains_cball_eq: "open S \ (\x. x \ S \ (\e>0. cball x e \ S))" by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) lemma eventually_nhds_ball: "d > 0 \ eventually (\x. x \ ball z d) (nhds z)" by (rule eventually_nhds_in_open) simp_all lemma eventually_at_ball: "d > 0 \ eventually (\t. t \ ball z d \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma eventually_at_ball': "d > 0 \ eventually (\t. t \ ball z d \ t \ z \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma at_within_ball: "e > 0 \ dist x y < e \ at y within ball x e = at y" by (subst at_within_open) auto lemma atLeastAtMost_eq_cball: fixes a b::real shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma cball_eq_atLeastAtMost: fixes a b::real shows "cball a b = {a - b .. a + b}" by (auto simp: dist_real_def) lemma greaterThanLessThan_eq_ball: fixes a b::real shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma ball_eq_greaterThanLessThan: fixes a b::real shows "ball a b = {a - b <..< a + b}" by (auto simp: dist_real_def) lemma interior_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) lemma cball_eq_empty [simp]: "cball x e = {} \ e < 0" apply (simp add: set_eq_iff not_le) apply (metis zero_le_dist dist_self order_less_le_trans) done lemma cball_empty [simp]: "e < 0 \ cball x e = {}" by simp lemma cball_sing: fixes x :: "'a::metric_space" shows "e = 0 \ cball x e = {x}" by (auto simp: set_eq_iff) lemma ball_divide_subset: "d \ 1 \ ball x (e/d) \ ball x e" apply (cases "e \ 0") apply (simp add: ball_empty field_split_simps) apply (rule subset_ball) apply (simp add: field_split_simps) done lemma ball_divide_subset_numeral: "ball x (e / numeral w) \ ball x e" using ball_divide_subset one_le_numeral by blast lemma cball_divide_subset: "d \ 1 \ cball x (e/d) \ cball x e" apply (cases "e < 0") apply (simp add: field_split_simps) apply (rule subset_cball) apply (metis div_by_1 frac_le not_le order_refl zero_less_one) done lemma cball_divide_subset_numeral: "cball x (e / numeral w) \ cball x e" using cball_divide_subset one_le_numeral by blast +lemma cball_scale: + assumes "a \ 0" + shows "(\x. a *\<^sub>R x) ` cball c r = cball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" +proof - + have 1: "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a + proof safe + fix x + assume x: "x \ cball c r" + have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" + by (auto simp: dist_norm) + also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" + by (simp add: algebra_simps) + finally show "a *\<^sub>R x \ cball (a *\<^sub>R c) (\a\ * r)" + using that x by (auto simp: dist_norm) + qed + + have "cball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` cball (a *\<^sub>R c) (\a\ * r)" + unfolding image_image using assms by simp + also have "\ \ (\x. a *\<^sub>R x) ` cball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" + using assms by (intro image_mono 1) auto + also have "\ = (\x. a *\<^sub>R x) ` cball c r" + using assms by (simp add: algebra_simps) + finally have "cball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` cball c r" . + moreover from assms have "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" + by (intro 1) auto + ultimately show ?thesis by blast +qed + +lemma ball_scale: + assumes "a \ 0" + shows "(\x. a *\<^sub>R x) ` ball c r = ball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" +proof - + have 1: "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a + proof safe + fix x + assume x: "x \ ball c r" + have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" + by (auto simp: dist_norm) + also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" + by (simp add: algebra_simps) + finally show "a *\<^sub>R x \ ball (a *\<^sub>R c) (\a\ * r)" + using that x by (auto simp: dist_norm) + qed + + have "ball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` ball (a *\<^sub>R c) (\a\ * r)" + unfolding image_image using assms by simp + also have "\ \ (\x. a *\<^sub>R x) ` ball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" + using assms by (intro image_mono 1) auto + also have "\ = (\x. a *\<^sub>R x) ` ball c r" + using assms by (simp add: algebra_simps) + finally have "ball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` ball c r" . + moreover from assms have "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" + by (intro 1) auto + ultimately show ?thesis by blast +qed subsection \Limit Points\ lemma islimpt_approachable: fixes x :: "'a::metric_space" shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" unfolding islimpt_iff_eventually eventually_at by fast lemma islimpt_approachable_le: "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x \ e)" for x :: "'a::metric_space" unfolding islimpt_approachable using approachable_lt_le [where f="\y. dist y x" and P="\y. y \ S \ y = x", THEN arg_cong [where f=Not]] by (simp add: Bex_def conj_commute conj_left_commute) lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \ x islimpt S" for x :: "'a::metric_space" apply (clarsimp simp add: islimpt_approachable) apply (drule_tac x="e/2" in spec) apply (auto simp: simp del: less_divide_eq_numeral1) apply (drule_tac x="dist x' x" in spec) apply (auto simp del: less_divide_eq_numeral1) apply metric done lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}" using closed_limpt limpt_of_limpts by blast lemma limpt_of_closure: "x islimpt closure S \ x islimpt S" for x :: "'a::metric_space" by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts) lemma islimpt_eq_infinite_ball: "x islimpt S \ (\e>0. infinite(S \ ball x e))" apply (simp add: islimpt_eq_acc_point, safe) apply (metis Int_commute open_ball centre_in_ball) by (metis open_contains_ball Int_mono finite_subset inf_commute subset_refl) lemma islimpt_eq_infinite_cball: "x islimpt S \ (\e>0. infinite(S \ cball x e))" apply (simp add: islimpt_eq_infinite_ball, safe) apply (meson Int_mono ball_subset_cball finite_subset order_refl) by (metis open_ball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq) subsection \Perfect Metric Spaces\ lemma perfect_choose_dist: "0 < r \ \a. a \ x \ dist a x < r" for x :: "'a::{perfect_space,metric_space}" using islimpt_UNIV [of x] by (simp add: islimpt_approachable) lemma cball_eq_sing: fixes x :: "'a::{metric_space,perfect_space}" shows "cball x e = {x} \ e = 0" proof (rule linorder_cases) assume e: "0 < e" obtain a where "a \ x" "dist a x < e" using perfect_choose_dist [OF e] by auto then have "a \ x" "dist x a \ e" by (auto simp: dist_commute) with e show ?thesis by (auto simp: set_eq_iff) qed auto subsection \?\ lemma finite_ball_include: fixes a :: "'a::metric_space" assumes "finite S" shows "\e>0. S \ ball a e" using assms proof induction case (insert x S) then obtain e0 where "e0>0" and e0:"S \ ball a e0" by auto define e where "e = max e0 (2 * dist a x)" have "e>0" unfolding e_def using \e0>0\ by auto moreover have "insert x S \ ball a e" using e0 \e>0\ unfolding e_def by auto ultimately show ?case by auto qed (auto intro: zero_less_one) lemma finite_set_avoid: fixes a :: "'a::metric_space" assumes "finite S" shows "\d>0. \x\S. x \ a \ d \ dist a x" using assms proof induction case (insert x S) then obtain d where "d > 0" and d: "\x\S. x \ a \ d \ dist a x" by blast show ?case proof (cases "x = a") case True with \d > 0 \d show ?thesis by auto next case False let ?d = "min d (dist a x)" from False \d > 0\ have dp: "?d > 0" by auto from d have d': "\x\S. x \ a \ ?d \ dist a x" by auto with dp False show ?thesis by (metis insert_iff le_less min_less_iff_conj not_less) qed qed (auto intro: zero_less_one) lemma discrete_imp_closed: fixes S :: "'a::metric_space set" assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" shows "closed S" proof - have False if C: "\e. e>0 \ \x'\S. x' \ x \ dist x' x < e" for x proof - from e have e2: "e/2 > 0" by arith from C[rule_format, OF e2] obtain y where y: "y \ S" "y \ x" "dist y x < e/2" by blast from e2 y(2) have mp: "min (e/2) (dist x y) > 0" by simp from d y C[OF mp] show ?thesis by metric qed then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) qed subsection \Interior\ lemma mem_interior: "x \ interior S \ (\e>0. ball x e \ S)" using open_contains_ball_eq [where S="interior S"] by (simp add: open_subset_interior) lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" by (meson ball_subset_cball interior_subset mem_interior open_contains_cball open_interior subset_trans) subsection \Frontier\ lemma frontier_straddle: fixes a :: "'a::metric_space" shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" unfolding frontier_def closure_interior by (auto simp: mem_interior subset_eq ball_def) subsection \Limits\ proposition Lim: "(f \ l) net \ trivial_limit net \ (\e>0. eventually (\x. dist (f x) l < e) net)" by (auto simp: tendsto_iff trivial_limit_eq) text \Show that they yield usual definitions in the various cases.\ proposition Lim_within_le: "(f \ l)(at a within S) \ (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a \ d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at_le) proposition Lim_within: "(f \ l) (at a within S) \ (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) corollary Lim_withinI [intro?]: assumes "\e. e > 0 \ \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l \ e" shows "(f \ l) (at a within S)" apply (simp add: Lim_within, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done proposition Lim_at: "(f \ l) (at a) \ (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) lemma Lim_transform_within_set: fixes a :: "'a::metric_space" and l :: "'b::metric_space" shows "\(f \ l) (at a within S); eventually (\x. x \ S \ x \ T) (at a)\ \ (f \ l) (at a within T)" apply (clarsimp simp: eventually_at Lim_within) apply (drule_tac x=e in spec, clarify) apply (rename_tac k) apply (rule_tac x="min d k" in exI, simp) done text \Another limit point characterization.\ lemma limpt_sequential_inj: fixes x :: "'a::metric_space" shows "x islimpt S \ (\f. (\n::nat. f n \ S - {x}) \ inj f \ (f \ x) sequentially)" (is "?lhs = ?rhs") proof assume ?lhs then have "\e>0. \x'\S. x' \ x \ dist x' x < e" by (force simp: islimpt_approachable) then obtain y where y: "\e. e>0 \ y e \ S \ y e \ x \ dist (y e) x < e" by metis define f where "f \ rec_nat (y 1) (\n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))" have [simp]: "f 0 = y 1" "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n by (simp_all add: f_def) have f: "f n \ S \ (f n \ x) \ dist (f n) x < inverse(2 ^ n)" for n proof (induction n) case 0 show ?case by (simp add: y) next case (Suc n) then show ?case apply (auto simp: y) by (metis half_gt_zero_iff inverse_positive_iff_positive less_divide_eq_numeral1(1) min_less_iff_conj y zero_less_dist_iff zero_less_numeral zero_less_power) qed show ?rhs proof (rule_tac x=f in exI, intro conjI allI) show "\n. f n \ S - {x}" using f by blast have "dist (f n) x < dist (f m) x" if "m < n" for m n using that proof (induction n) case 0 then show ?case by simp next case (Suc n) then consider "m < n" | "m = n" using less_Suc_eq by blast then show ?case proof cases assume "m < n" have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x" by simp also have "\ < dist (f n) x" by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y) also have "\ < dist (f m) x" using Suc.IH \m < n\ by blast finally show ?thesis . next assume "m = n" then show ?case by simp (metis dist_pos_lt f half_gt_zero_iff inverse_positive_iff_positive min_less_iff_conj y zero_less_numeral zero_less_power) qed qed then show "inj f" by (metis less_irrefl linorder_injI) show "f \ x" apply (rule tendstoI) apply (rule_tac c="nat (ceiling(1/e))" in eventually_sequentiallyI) apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]]) apply (simp add: field_simps) by (meson le_less_trans mult_less_cancel_left not_le of_nat_less_two_power) qed next assume ?rhs then show ?lhs by (fastforce simp add: islimpt_approachable lim_sequentially) qed lemma Lim_dist_ubound: assumes "\(trivial_limit net)" and "(f \ l) net" and "eventually (\x. dist a (f x) \ e) net" shows "dist a l \ e" using assms by (fast intro: tendsto_le tendsto_intros) subsection \Continuity\ text\Derive the epsilon-delta forms, which we often use as "definitions"\ proposition continuous_within_eps_delta: "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" unfolding continuous_within and Lim_within by fastforce corollary continuous_at_eps_delta: "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" using continuous_within_eps_delta [of x UNIV f] by simp lemma continuous_at_right_real_increasing: fixes f :: "real \ real" assumes nondecF: "\x y. x \ y \ f x \ f y" shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong, safe) apply (erule_tac x="a + d" in allE, simp) apply (simp add: nondecF field_simps) apply (drule nondecF, simp) done lemma continuous_at_left_real_increasing: assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" shows "(continuous (at_left (a :: real)) f) = (\e > 0. \delta > 0. f a - f (a - delta) < e)" apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong, safe) apply (erule_tac x="a - d" in allE, simp) apply (simp add: nondecF field_simps) apply (cut_tac x="a - d" and y=x in nondecF, simp_all) done text\Versions in terms of open balls.\ lemma continuous_within_ball: "continuous (at x within s) f \ (\e > 0. \d > 0. f ` (ball x d \ s) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs { fix e :: real assume "e > 0" then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using \?lhs\[unfolded continuous_within Lim_within] by auto { fix y assume "y \ f ` (ball x d \ s)" then have "y \ ball (f x) e" using d(2) using \e > 0\ by (auto simp: dist_commute) } then have "\d>0. f ` (ball x d \ s) \ ball (f x) e" using \d > 0\ unfolding subset_eq ball_def by (auto simp: dist_commute) } then show ?rhs by auto next assume ?rhs then show ?lhs unfolding continuous_within Lim_within ball_def subset_eq apply (auto simp: dist_commute) apply (erule_tac x=e in allE, auto) done qed lemma continuous_at_ball: "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball apply auto apply (erule_tac x=e in allE, auto) apply (rule_tac x=d in exI, auto) apply (erule_tac x=xa in allE) apply (auto simp: dist_commute) done next assume ?rhs then show ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball apply auto apply (erule_tac x=e in allE, auto) apply (rule_tac x=d in exI, auto) apply (erule_tac x="f xa" in allE) apply (auto simp: dist_commute) done qed text\Define setwise continuity in terms of limits within the set.\ lemma continuous_on_iff: "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding continuous_on_def Lim_within by (metis dist_pos_lt dist_self) lemma continuous_within_E: assumes "continuous (at x within s) f" "e>0" obtains d where "d>0" "\x'. \x'\ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms apply (simp add: continuous_within_eps_delta) apply (drule spec [of _ e], clarify) apply (rule_tac d="d/2" in that, auto) done lemma continuous_onI [intro?]: assumes "\x e. \e > 0; x \ s\ \ \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) \ e" shows "continuous_on s f" apply (simp add: continuous_on_iff, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done text\Some simple consequential lemmas.\ lemma continuous_onE: assumes "continuous_on s f" "x\s" "e>0" obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms apply (simp add: continuous_on_iff) apply (elim ballE allE) apply (auto intro: that [where d="d/2" for d]) done text\The usual transformation theorems.\ lemma continuous_transform_within: fixes f g :: "'a::metric_space \ 'b::topological_space" assumes "continuous (at x within s) f" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "continuous (at x within s) g" using assms unfolding continuous_within by (force intro: Lim_transform_within) subsection \Closure and Limit Characterization\ lemma closure_approachable: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" apply (auto simp: closure_def islimpt_approachable) apply (metis dist_self) done lemma closure_approachable_le: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x \ e)" unfolding closure_approachable using dense by force lemma closure_approachableD: assumes "x \ closure S" "e>0" shows "\y\S. dist x y < e" using assms unfolding closure_approachable by (auto simp: dist_commute) lemma closed_approachable: fixes S :: "'a::metric_space set" shows "closed S \ (\e>0. \y\S. dist y x < e) \ x \ S" by (metis closure_closed closure_approachable) lemma closure_contains_Inf: fixes S :: "real set" assumes "S \ {}" "bdd_below S" shows "Inf S \ closure S" proof - have *: "\x\S. Inf S \ x" using cInf_lower[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Inf S < Inf S + e" by simp with assms obtain x where "x \ S" "x < Inf S + e" by (subst (asm) cInf_less_iff) auto with * have "\x\S. dist x (Inf S) < e" by (intro bexI[of _ x]) (auto simp: dist_real_def) } then show ?thesis unfolding closure_approachable by auto qed lemma closure_contains_Sup: fixes S :: "real set" assumes "S \ {}" "bdd_above S" shows "Sup S \ closure S" proof - have *: "\x\S. x \ Sup S" using cSup_upper[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Sup S - e < Sup S" by simp with assms obtain x where "x \ S" "Sup S - e < x" by (subst (asm) less_cSup_iff) auto with * have "\x\S. dist x (Sup S) < e" by (intro bexI[of _ x]) (auto simp: dist_real_def) } then show ?thesis unfolding closure_approachable by auto qed lemma not_trivial_limit_within_ball: "\ trivial_limit (at x within S) \ (\e>0. S \ ball x e - {x} \ {})" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S - {x}" and "dist y x < e" using \?lhs\ not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto then have "y \ S \ ball x e - {x}" unfolding ball_def by (simp add: dist_commute) then have "S \ ball x e - {x} \ {}" by blast } then show ?thesis by auto qed show ?lhs if ?rhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S \ ball x e - {x}" using \?rhs\ by blast then have "y \ S - {x}" and "dist y x < e" unfolding ball_def by (simp_all add: dist_commute) then have "\y \ S - {x}. dist y x < e" by auto } then show ?thesis using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto qed qed subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) definition\<^marker>\tag important\ (in metric_space) bounded :: "'a set \ bool" where "bounded S \ (\x e. \y\S. dist x y \ e)" lemma bounded_subset_cball: "bounded S \ (\e x. S \ cball x e \ 0 \ e)" unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist) lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" unfolding bounded_def by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" unfolding bounded_any_center [where a=0] by (simp add: dist_norm) lemma bdd_above_norm: "bdd_above (norm ` X) \ bounded X" by (simp add: bounded_iff bdd_above_def) lemma bounded_norm_comp: "bounded ((\x. norm (f x)) ` S) = bounded (f ` S)" by (simp add: bounded_iff) lemma boundedI: assumes "\x. x \ S \ norm x \ B" shows "bounded S" using assms bounded_iff by blast lemma bounded_empty [simp]: "bounded {}" by (simp add: bounded_def) lemma bounded_subset: "bounded T \ S \ T \ bounded S" by (metis bounded_def subset_eq) lemma bounded_interior[intro]: "bounded S \ bounded(interior S)" by (metis bounded_subset interior_subset) lemma bounded_closure[intro]: assumes "bounded S" shows "bounded (closure S)" proof - from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto { fix y assume "y \ closure S" then obtain f where f: "\n. f n \ S" "(f \ y) sequentially" unfolding closure_sequential by auto have "\n. f n \ S \ dist x (f n) \ a" using a by simp then have "eventually (\n. dist x (f n) \ a) sequentially" by (simp add: f(1)) have "dist x y \ a" apply (rule Lim_dist_ubound [of sequentially f]) apply (rule trivial_limit_sequentially) apply (rule f(2)) apply fact done } then show ?thesis unfolding bounded_def by auto qed lemma bounded_closure_image: "bounded (f ` closure S) \ bounded (f ` S)" by (simp add: bounded_subset closure_subset image_mono) lemma bounded_cball[simp,intro]: "bounded (cball x e)" apply (simp add: bounded_def) apply (rule_tac x=x in exI) apply (rule_tac x=e in exI, auto) done lemma bounded_ball[simp,intro]: "bounded (ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj) lemma bounded_Union[intro]: "finite F \ \S\F. bounded S \ bounded (\F)" by (induct rule: finite_induct[of F]) auto lemma bounded_UN [intro]: "finite A \ \x\A. bounded (B x) \ bounded (\x\A. B x)" by (induct set: finite) auto lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" proof - have "\y\{x}. dist x y \ 0" by simp then have "bounded {x}" unfolding bounded_def by fast then show ?thesis by (metis insert_is_Un bounded_Un) qed lemma bounded_subset_ballI: "S \ ball x r \ bounded S" by (meson bounded_ball bounded_subset) lemma bounded_subset_ballD: assumes "bounded S" shows "\r. 0 < r \ S \ ball x r" proof - obtain e::real and y where "S \ cball y e" "0 \ e" using assms by (auto simp: bounded_subset_cball) then show ?thesis by (intro exI[where x="dist x y + e + 1"]) metric qed lemma finite_imp_bounded [intro]: "finite S \ bounded S" by (induct set: finite) simp_all lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" by (metis Int_lower1 Int_lower2 bounded_subset) lemma bounded_diff[intro]: "bounded S \ bounded (S - T)" by (metis Diff_subset bounded_subset) lemma bounded_dist_comp: assumes "bounded (f ` S)" "bounded (g ` S)" shows "bounded ((\x. dist (f x) (g x)) ` S)" proof - from assms obtain M1 M2 where *: "dist (f x) undefined \ M1" "dist undefined (g x) \ M2" if "x \ S" for x by (auto simp: bounded_any_center[of _ undefined] dist_commute) have "dist (f x) (g x) \ M1 + M2" if "x \ S" for x using *[OF that] by metric then show ?thesis by (auto intro!: boundedI) qed lemma bounded_Times: assumes "bounded s" "bounded t" shows "bounded (s \ t)" proof - obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" using assms [unfolded bounded_def] by auto then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto qed subsection \Compactness\ lemma compact_imp_bounded: assumes "compact U" shows "bounded U" proof - have "compact U" "\x\U. open (ball x 1)" "U \ (\x\U. ball x 1)" using assms by auto then obtain D where D: "D \ U" "finite D" "U \ (\x\D. ball x 1)" by (metis compactE_image) from \finite D\ have "bounded (\x\D. ball x 1)" by (simp add: bounded_UN) then show "bounded U" using \U \ (\x\D. ball x 1)\ by (rule bounded_subset) qed lemma closure_Int_ball_not_empty: assumes "S \ closure T" "x \ S" "r > 0" shows "T \ ball x r \ {}" using assms centre_in_ball closure_iff_nhds_not_empty by blast lemma compact_sup_maxdistance: fixes s :: "'a::metric_space set" assumes "compact s" and "s \ {}" shows "\x\s. \y\s. \u\s. \v\s. dist u v \ dist x y" proof - have "compact (s \ s)" using \compact s\ by (intro compact_Times) moreover have "s \ s \ {}" using \s \ {}\ by auto moreover have "continuous_on (s \ s) (\x. dist (fst x) (snd x))" by (intro continuous_at_imp_continuous_on ballI continuous_intros) ultimately show ?thesis using continuous_attains_sup[of "s \ s" "\x. dist (fst x) (snd x)"] by auto qed subsubsection\Totally bounded\ lemma cauchy_def: "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N \ dist (s m) (s n) < e)" unfolding Cauchy_def by metis proposition seq_compact_imp_totally_bounded: assumes "seq_compact s" shows "\e>0. \k. finite k \ k \ s \ s \ (\x\k. ball x e)" proof - { fix e::real assume "e > 0" assume *: "\k. finite k \ k \ s \ \ s \ (\x\k. ball x e)" let ?Q = "\x n r. r \ s \ (\m < (n::nat). \ (dist (x m) r < e))" have "\x. \n::nat. ?Q x n (x n)" proof (rule dependent_wellorder_choice) fix n x assume "\y. y < n \ ?Q x y (x y)" then have "\ s \ (\x\x ` {0..s" "z \ (\x\x ` {0..r. ?Q x n r" using z by auto qed simp then obtain x where "\n::nat. x n \ s" and x:"\n m. m < n \ \ (dist (x m) (x n) < e)" by blast then obtain l r where "l \ s" and r:"strict_mono r" and "((x \ r) \ l) sequentially" using assms by (metis seq_compact_def) from this(3) have "Cauchy (x \ r)" using LIMSEQ_imp_Cauchy by auto then obtain N::nat where "\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" unfolding cauchy_def using \e > 0\ by blast then have False using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) } then show ?thesis by metis qed subsubsection\Heine-Borel theorem\ proposition seq_compact_imp_Heine_Borel: fixes s :: "'a :: metric_space set" assumes "seq_compact s" shows "compact s" proof - from seq_compact_imp_totally_bounded[OF \seq_compact s\] obtain f where f: "\e>0. finite (f e) \ f e \ s \ s \ (\x\f e. ball x e)" unfolding choice_iff' .. define K where "K = (\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" have "countably_compact s" using \seq_compact s\ by (rule seq_compact_imp_countably_compact) then show "compact s" proof (rule countably_compact_imp_compact) show "countable K" unfolding K_def using f by (auto intro: countable_finite countable_subset countable_rat intro!: countable_image countable_SIGMA countable_UN) show "\b\K. open b" by (auto simp: K_def) next fix T x assume T: "open T" "x \ T" and x: "x \ s" from openE[OF T] obtain e where "0 < e" "ball x e \ T" by auto then have "0 < e / 2" "ball x (e / 2) \ T" by auto from Rats_dense_in_real[OF \0 < e / 2\] obtain r where "r \ \" "0 < r" "r < e / 2" by auto from f[rule_format, of r] \0 < r\ \x \ s\ obtain k where "k \ f r" "x \ ball k r" by auto from \r \ \\ \0 < r\ \k \ f r\ have "ball k r \ K" by (auto simp: K_def) then show "\b\K. x \ b \ b \ s \ T" proof (rule bexI[rotated], safe) fix y assume "y \ ball k r" with \r < e / 2\ \x \ ball k r\ have "dist x y < e" by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute) with \ball x e \ T\ show "y \ T" by auto next show "x \ ball k r" by fact qed qed qed proposition compact_eq_seq_compact_metric: "compact (s :: 'a::metric_space set) \ seq_compact s" using compact_imp_seq_compact seq_compact_imp_Heine_Borel by blast proposition compact_def: \ \this is the definition of compactness in HOL Light\ "compact (S :: 'a::metric_space set) \ (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ (f \ r) \ l))" unfolding compact_eq_seq_compact_metric seq_compact_def by auto subsubsection \Complete the chain of compactness variants\ proposition compact_eq_Bolzano_Weierstrass: fixes s :: "'a::metric_space set" shows "compact s \ (\t. infinite t \ t \ s --> (\x \ s. x islimpt t))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using Heine_Borel_imp_Bolzano_Weierstrass[of s] by auto next assume ?rhs then show ?lhs unfolding compact_eq_seq_compact_metric by (rule Bolzano_Weierstrass_imp_seq_compact) qed proposition Bolzano_Weierstrass_imp_bounded: "\t. infinite t \ t \ s \ (\x \ s. x islimpt t) \ bounded s" using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass . subsection \Banach fixed point theorem\ theorem banach_fix:\ \TODO: rename to \Banach_fix\\ assumes s: "complete s" "s \ {}" and c: "0 \ c" "c < 1" and f: "f ` s \ s" and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" shows "\!x\s. f x = x" proof - from c have "1 - c > 0" by simp from s(2) obtain z0 where z0: "z0 \ s" by blast define z where "z n = (f ^^ n) z0" for n with f z0 have z_in_s: "z n \ s" for n :: nat by (induct n) auto define d where "d = dist (z 0) (z 1)" have fzn: "f (z n) = z (Suc n)" for n by (simp add: z_def) have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat proof (induct n) case 0 then show ?case by (simp add: d_def) next case (Suc m) with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp then show ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] by (simp add: fzn mult_le_cancel_left) qed have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat proof (induct n) case 0 show ?case by simp next case (Suc k) from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" by (simp add: dist_triangle) also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" by simp also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" by (simp add: field_simps) also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" by (simp add: power_add field_simps) also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" by (simp add: field_simps) finally show ?case by simp qed have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e proof (cases "d = 0") case True from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) with c cf_z2[of 0] True have "z n = z0" for n by (simp add: z_def) with \e > 0\ show ?thesis by simp next case False with zero_le_dist[of "z 0" "z 1"] have "d > 0" by (metis d_def less_le) with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" by simp with c obtain N where N: "c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat proof - from c \n \ N\ have *: "c ^ n \ c ^ N" using power_decreasing[OF \n\N\, of c] by simp from c \m > n\ have "1 - c ^ (m - n) > 0" using power_strict_mono[of c 1 "m - n"] by simp with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" by simp from cf_z2[of n "m - n"] \m > n\ have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" using mult_right_mono[OF * order_less_imp_le[OF **]] by (simp add: mult.assoc) also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" by simp also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto finally show ?thesis by simp qed have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat proof (cases "n = m") case True with \e > 0\ show ?thesis by simp next case False with *[of n m] *[of m n] and that show ?thesis by (auto simp: dist_commute nat_neq_iff) qed then show ?thesis by auto qed then have "Cauchy z" by (simp add: cauchy_def) then obtain x where "x\s" and x:"(z \ x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto define e where "e = dist (f x) x" have "e = 0" proof (rule ccontr) assume "e \ 0" then have "e > 0" unfolding e_def using zero_le_dist[of "f x" x] by (metis dist_eq_0_iff dist_nz e_def) then obtain N where N:"\n\N. dist (z n) x < e / 2" using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto then have N':"dist (z N) x < e / 2" by auto have *: "c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 using zero_le_dist[of "z N" x] and c by (metis dist_eq_0_iff dist_nz order_less_asym less_le) have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] using z_in_s[of N] \x\s\ using c by auto also have "\ < e / 2" using N' and c using * by auto finally show False unfolding fzn using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] unfolding e_def by auto qed then have "f x = x" by (auto simp: e_def) moreover have "y = x" if "f y = y" "y \ s" for y proof - from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp with c and zero_le_dist[of x y] have "dist x y = 0" by (simp add: mult_le_cancel_right1) then show ?thesis by simp qed ultimately show ?thesis using \x\s\ by blast qed subsection \Edelstein fixed point theorem\ theorem edelstein_fix:\ \TODO: rename to \Edelstein_fix\\ fixes s :: "'a::metric_space set" assumes s: "compact s" "s \ {}" and gs: "(g ` s) \ s" and dist: "\x\s. \y\s. x \ y \ dist (g x) (g y) < dist x y" shows "\!x\s. g x = x" proof - let ?D = "(\x. (x, x)) ` s" have D: "compact ?D" "?D \ {}" by (rule compact_continuous_image) (auto intro!: s continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) have "\x y e. x \ s \ y \ s \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" using dist by fastforce then have "continuous_on s g" by (auto simp: continuous_on_iff) then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" unfolding continuous_on_eq_continuous_within by (intro continuous_dist ballI continuous_within_compose) (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) obtain a where "a \ s" and le: "\x. x \ s \ dist (g a) a \ dist (g x) x" using continuous_attains_inf[OF D cont] by auto have "g a = a" proof (rule ccontr) assume "g a \ a" with \a \ s\ gs have "dist (g (g a)) (g a) < dist (g a) a" by (intro dist[rule_format]) auto moreover have "dist (g a) a \ dist (g (g a)) (g a)" using \a \ s\ gs by (intro le) auto ultimately show False by auto qed moreover have "\x. x \ s \ g x = x \ x = a" using dist[THEN bspec[where x=a]] \g a = a\ and \a\s\ by auto ultimately show "\!x\s. g x = x" using \a \ s\ by blast qed subsection \The diameter of a set\ definition\<^marker>\tag important\ diameter :: "'a::metric_space set \ real" where "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" lemma diameter_empty [simp]: "diameter{} = 0" by (auto simp: diameter_def) lemma diameter_singleton [simp]: "diameter{x} = 0" by (auto simp: diameter_def) lemma diameter_le: assumes "S \ {} \ 0 \ d" and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" shows "diameter S \ d" using assms by (auto simp: dist_norm diameter_def intro: cSUP_least) lemma diameter_bounded_bound: fixes s :: "'a :: metric_space set" assumes s: "bounded s" "x \ s" "y \ s" shows "dist x y \ diameter s" proof - from s obtain z d where z: "\x. x \ s \ dist z x \ d" unfolding bounded_def by auto have "bdd_above (case_prod dist ` (s\s))" proof (intro bdd_aboveI, safe) fix a b assume "a \ s" "b \ s" with z[of a] z[of b] dist_triangle[of a b z] show "dist a b \ 2 * d" by (simp add: dist_commute) qed moreover have "(x,y) \ s\s" using s by auto ultimately have "dist x y \ (SUP (x,y)\s\s. dist x y)" by (rule cSUP_upper2) simp with \x \ s\ show ?thesis by (auto simp: diameter_def) qed lemma diameter_lower_bounded: fixes s :: "'a :: metric_space set" assumes s: "bounded s" and d: "0 < d" "d < diameter s" shows "\x\s. \y\s. d < dist x y" proof (rule ccontr) assume contr: "\ ?thesis" moreover have "s \ {}" using d by (auto simp: diameter_def) ultimately have "diameter s \ d" by (auto simp: not_less diameter_def intro!: cSUP_least) with \d < diameter s\ show False by auto qed lemma diameter_bounded: assumes "bounded s" shows "\x\s. \y\s. dist x y \ diameter s" and "\d>0. d < diameter s \ (\x\s. \y\s. dist x y > d)" using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms by auto lemma bounded_two_points: "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" apply (rule iffI) subgoal using diameter_bounded(1) by auto subgoal using bounded_any_center[of S] by meson done lemma diameter_compact_attained: assumes "compact s" and "s \ {}" shows "\x\s. \y\s. dist x y = diameter s" proof - have b: "bounded s" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys: "x\s" "y\s" and xy: "\u\s. \v\s. dist u v \ dist x y" using compact_sup_maxdistance[OF assms] by auto then have "diameter s \ dist x y" unfolding diameter_def apply clarsimp apply (rule cSUP_least, fast+) done then show ?thesis by (metis b diameter_bounded_bound order_antisym xys) qed lemma diameter_ge_0: assumes "bounded S" shows "0 \ diameter S" by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) lemma diameter_subset: assumes "S \ T" "bounded T" shows "diameter S \ diameter T" proof (cases "S = {} \ T = {}") case True with assms show ?thesis by (force simp: diameter_ge_0) next case False then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) with False \S \ T\ show ?thesis apply (simp add: diameter_def) apply (rule cSUP_subset_mono, auto) done qed lemma diameter_closure: assumes "bounded S" shows "diameter(closure S) = diameter S" proof (rule order_antisym) have "False" if "diameter S < diameter (closure S)" proof - define d where "d = diameter(closure S) - diameter(S)" have "d > 0" using that by (simp add: d_def) then have "diameter(closure(S)) - d / 2 < diameter(closure(S))" by simp have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" by (simp add: d_def field_split_simps) have bocl: "bounded (closure S)" using assms by blast moreover have "0 \ diameter S" using assms diameter_ge_0 by blast ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \d > 0\ d_def by auto then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" using closure_approachable by (metis \0 < d\ zero_less_divide_iff zero_less_numeral) then have "dist x' y' \ diameter S" using assms diameter_bounded_bound by blast with x'y' have "dist x y \ d / 4 + diameter S + d / 4" by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans) then show ?thesis using xy d_def by linarith qed then show "diameter (closure S) \ diameter S" by fastforce next show "diameter S \ diameter (closure S)" by (simp add: assms bounded_closure closure_subset diameter_subset) qed proposition Lebesgue_number_lemma: assumes "compact S" "\ \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" proof (cases "S = {}") case True then show ?thesis by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) next case False { fix x assume "x \ S" then obtain C where C: "x \ C" "C \ \" using \S \ \\\ by blast then obtain r where r: "r>0" "ball x (2*r) \ C" by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" using C by blast } then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" by metis then have "S \ (\x \ S. ball x (r x))" by auto then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" by (rule compactE [OF \compact S\]) auto then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" by (meson finite_subset_image) then have "S0 \ {}" using False \S \ \\\ by auto define \ where "\ = Inf (r ` S0)" have "\ > 0" using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) show ?thesis proof show "0 < \" by (simp add: \0 < \\) show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T proof (cases "T = {}") case True then show ?thesis using \\ \ {}\ by blast next case False then obtain y where "y \ T" by blast then have "y \ S" using \T \ S\ by auto then obtain x where "x \ S0" and x: "y \ ball x (r x)" using \S \ \\\ S0 that by blast have "ball y \ \ ball y (r x)" by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) also have "... \ ball x (2*r x)" using x by metric finally obtain C where "C \ \" "ball y \ \ C" by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) have "bounded T" using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast then have "T \ ball y \" using \y \ T\ dia diameter_bounded_bound by fastforce then show ?thesis apply (rule_tac x=C in bexI) using \ball y \ \ C\ \C \ \\ by auto qed qed qed subsection \Metric spaces with the Heine-Borel property\ text \ A metric space (or topological vector space) is said to have the Heine-Borel property if every closed and bounded subset is compact. \ class heine_borel = metric_space + assumes bounded_imp_convergent_subsequence: "bounded (range f) \ \l r. strict_mono (r::nat\nat) \ ((f \ r) \ l) sequentially" proposition bounded_closed_imp_seq_compact: fixes s::"'a::heine_borel set" assumes "bounded s" and "closed s" shows "seq_compact s" proof (unfold seq_compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ s" with \bounded s\ have "bounded (range f)" by (auto intro: bounded_subset) obtain l r where r: "strict_mono (r :: nat \ nat)" and l: "((f \ r) \ l) sequentially" using bounded_imp_convergent_subsequence [OF \bounded (range f)\] by auto from f have fr: "\n. (f \ r) n \ s" by simp have "l \ s" using \closed s\ fr l by (rule closed_sequentially) show "\l\s. \r. strict_mono r \ ((f \ r) \ l) sequentially" using \l \ s\ r l by blast qed lemma compact_eq_bounded_closed: fixes s :: "'a::heine_borel set" shows "compact s \ bounded s \ closed s" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using compact_imp_closed compact_imp_bounded by blast next assume ?rhs then show ?lhs using bounded_closed_imp_seq_compact[of s] unfolding compact_eq_seq_compact_metric by auto qed lemma compact_Inter: fixes \ :: "'a :: heine_borel set set" assumes com: "\S. S \ \ \ compact S" and "\ \ {}" shows "compact(\ \)" using assms by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed) lemma compact_closure [simp]: fixes S :: "'a::heine_borel set" shows "compact(closure S) \ bounded S" by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed) instance\<^marker>\tag important\ real :: heine_borel proof fix f :: "nat \ real" assume f: "bounded (range f)" obtain r :: "nat \ nat" where r: "strict_mono r" "monoseq (f \ r)" unfolding comp_def by (metis seq_monosub) then have "Bseq (f \ r)" unfolding Bseq_eq_bounded using f by (metis BseqI' bounded_iff comp_apply rangeI) with r show "\l r. strict_mono r \ (f \ r) \ l" using Bseq_monoseq_convergent[of "f \ r"] by (auto simp: convergent_def) qed lemma compact_lemma_general: fixes f :: "nat \ 'a" fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl "proj" 60) fixes unproj:: "('b \ 'c) \ 'a" assumes finite_basis: "finite basis" assumes bounded_proj: "\k. k \ basis \ bounded ((\x. x proj k) ` range f)" assumes proj_unproj: "\e k. k \ basis \ (unproj e) proj k = e k" assumes unproj_proj: "\x. unproj (\k. x proj k) = x" shows "\d\basis. \l::'a. \ r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof safe fix d :: "'b set" assume d: "d \ basis" with finite_basis have "finite d" by (blast intro: finite_subset) from this d show "\l::'a. \r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof (induct d) case empty then show ?case unfolding strict_mono_def by auto next case (insert k d) have k[intro]: "k \ basis" using insert by auto have s': "bounded ((\x. x proj k) ` range f)" using k by (rule bounded_proj) obtain l1::"'a" and r1 where r1: "strict_mono r1" and lr1: "\e > 0. eventually (\n. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially" using insert(3) using insert(4) by auto have f': "\n. f (r1 n) proj k \ (\x. x proj k) ` range f" by simp have "bounded (range (\i. f (r1 i) proj k))" by (metis (lifting) bounded_subset f' image_subsetI s') then obtain l2 r2 where r2:"strict_mono r2" and lr2:"((\i. f (r1 (r2 i)) proj k) \ l2) sequentially" using bounded_imp_convergent_subsequence[of "\i. f (r1 i) proj k"] by (auto simp: o_def) define r where "r = r1 \ r2" have r:"strict_mono r" using r1 and r2 unfolding r_def o_def strict_mono_def by auto moreover define l where "l = unproj (\i. if i = k then l2 else l1 proj i)" { fix e::real assume "e > 0" from lr1 \e > 0\ have N1: "eventually (\n. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially" by blast from lr2 \e > 0\ have N2:"eventually (\n. dist (f (r1 (r2 n)) proj k) l2 < e) sequentially" by (rule tendstoD) from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e) sequentially" by (rule eventually_subseq) have "eventually (\n. \i\(insert k d). dist (f (r n) proj i) (l proj i) < e) sequentially" using N1' N2 by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def proj_unproj) } ultimately show ?case by auto qed qed lemma bounded_fst: "bounded s \ bounded (fst ` s)" unfolding bounded_def by (metis (erased, hide_lams) dist_fst_le image_iff order_trans) lemma bounded_snd: "bounded s \ bounded (snd ` s)" unfolding bounded_def by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans) instance\<^marker>\tag important\ prod :: (heine_borel, heine_borel) heine_borel proof fix f :: "nat \ 'a \ 'b" assume f: "bounded (range f)" then have "bounded (fst ` range f)" by (rule bounded_fst) then have s1: "bounded (range (fst \ f))" by (simp add: image_comp) obtain l1 r1 where r1: "strict_mono r1" and l1: "(\n. fst (f (r1 n))) \ l1" using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast from f have s2: "bounded (range (snd \ f \ r1))" by (auto simp: image_comp intro: bounded_snd bounded_subset) obtain l2 r2 where r2: "strict_mono r2" and l2: "((\n. snd (f (r1 (r2 n)))) \ l2) sequentially" using bounded_imp_convergent_subsequence [OF s2] unfolding o_def by fast have l1': "((\n. fst (f (r1 (r2 n)))) \ l1) sequentially" using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . have l: "((f \ (r1 \ r2)) \ (l1, l2)) sequentially" using tendsto_Pair [OF l1' l2] unfolding o_def by simp have r: "strict_mono (r1 \ r2)" using r1 r2 unfolding strict_mono_def by simp show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" using l r by fast qed subsection \Completeness\ proposition (in metric_space) completeI: assumes "\f. \n. f n \ s \ Cauchy f \ \l\s. f \ l" shows "complete s" using assms unfolding complete_def by fast proposition (in metric_space) completeE: assumes "complete s" and "\n. f n \ s" and "Cauchy f" obtains l where "l \ s" and "f \ l" using assms unfolding complete_def by fast (* TODO: generalize to uniform spaces *) lemma compact_imp_complete: fixes s :: "'a::metric_space set" assumes "compact s" shows "complete s" proof - { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" from as(1) obtain l r where lr: "l\s" "strict_mono r" "(f \ r) \ l" using assms unfolding compact_def by blast note lr' = seq_suble [OF lr(2)] { fix e :: real assume "e > 0" from as(2) obtain N where N:"\m n. N \ m \ N \ n \ dist (f m) (f n) < e/2" unfolding cauchy_def using \e > 0\ apply (erule_tac x="e/2" in allE, auto) done from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" using \e > 0\ by auto { fix n :: nat assume n: "n \ max N M" have "dist ((f \ r) n) l < e/2" using n M by auto moreover have "r n \ N" using lr'[of n] n by auto then have "dist (f n) ((f \ r) n) < e / 2" using N and n by auto ultimately have "dist (f n) l < e" using n M by metric } then have "\N. \n\N. dist (f n) l < e" by blast } then have "\l\s. (f \ l) sequentially" using \l\s\ unfolding lim_sequentially by auto } then show ?thesis unfolding complete_def by auto qed proposition compact_eq_totally_bounded: "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\x\k. ball x e))" (is "_ \ ?rhs") proof assume assms: "?rhs" then obtain k where k: "\e. 0 < e \ finite (k e)" "\e. 0 < e \ s \ (\x\k e. ball x e)" by (auto simp: choice_iff') show "compact s" proof cases assume "s = {}" then show "compact s" by (simp add: compact_def) next assume "s \ {}" show ?thesis unfolding compact_def proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ s" define e where "e n = 1 / (2 * Suc n)" for n then have [simp]: "\n. 0 < e n" by auto define B where "B n U = (SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U))" for n U { fix n U assume "infinite {n. f n \ U}" then have "\b\k (e n). infinite {i\{n. f n \ U}. f i \ ball b (e n)}" using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) then obtain a where "a \ k (e n)" "infinite {i \ {n. f n \ U}. f i \ ball a (e n)}" .. then have "\b. infinite {i. f i \ b} \ (\x. b \ ball x (e n) \ U)" by (intro exI[of _ "ball a (e n) \ U"] exI[of _ a]) (auto simp: ac_simps) from someI_ex[OF this] have "infinite {i. f i \ B n U}" "\x. B n U \ ball x (e n) \ U" unfolding B_def by auto } note B = this define F where "F = rec_nat (B 0 UNIV) B" { fix n have "infinite {i. f i \ F n}" by (induct n) (auto simp: F_def B) } then have F: "\n. \x. F (Suc n) \ ball x (e n) \ F n" using B by (simp add: F_def) then have F_dec: "\m n. m \ n \ F n \ F m" using decseq_SucI[of F] by (auto simp: decseq_def) obtain sel where sel: "\k i. i < sel k i" "\k i. f (sel k i) \ F k" proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) fix k i have "infinite ({n. f n \ F k} - {.. i})" using \infinite {n. f n \ F k}\ by auto from infinite_imp_nonempty[OF this] show "\x>i. f x \ F k" by (simp add: set_eq_iff not_le conj_commute) qed define t where "t = rec_nat (sel 0 0) (\n i. sel (Suc n) i)" have "strict_mono t" unfolding strict_mono_Suc_iff by (simp add: t_def sel) moreover have "\i. (f \ t) i \ s" using f by auto moreover { fix n have "(f \ t) n \ F n" by (cases n) (simp_all add: t_def sel) } note t = this have "Cauchy (f \ t)" proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) fix r :: real and N n m assume "1 / Suc N < r" "Suc N \ n" "Suc N \ m" then have "(f \ t) n \ F (Suc N)" "(f \ t) m \ F (Suc N)" "2 * e N < r" using F_dec t by (auto simp: e_def field_simps) with F[of N] obtain x where "dist x ((f \ t) n) < e N" "dist x ((f \ t) m) < e N" by (auto simp: subset_eq) with \2 * e N < r\ show "dist ((f \ t) m) ((f \ t) n) < r" by metric qed ultimately show "\l\s. \r. strict_mono r \ (f \ r) \ l" using assms unfolding complete_def by blast qed qed qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)" proof - from assms obtain N :: nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def by force then have N:"\n. N \ n \ dist (s N) (s n) < 1" by auto moreover have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" unfolding bounded_any_center [where a="s N"] by auto ultimately show "?thesis" unfolding bounded_any_center [where a="s N"] apply (rule_tac x="max a 1" in exI, auto) apply (erule_tac x=y in allE) apply (erule_tac x=y in ballE, auto) done qed instance heine_borel < complete_space proof fix f :: "nat \ 'a" assume "Cauchy f" then have "bounded (range f)" by (rule cauchy_imp_bounded) then have "compact (closure (range f))" unfolding compact_eq_bounded_closed by auto then have "complete (closure (range f))" by (rule compact_imp_complete) moreover have "\n. f n \ closure (range f)" using closure_subset [of "range f"] by auto ultimately have "\l\closure (range f). (f \ l) sequentially" using \Cauchy f\ unfolding complete_def by auto then show "convergent f" unfolding convergent_def by auto qed lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)" proof (rule completeI) fix f :: "nat \ 'a" assume "Cauchy f" then have "convergent f" by (rule Cauchy_convergent) then show "\l\UNIV. f \ l" unfolding convergent_def by simp qed lemma complete_imp_closed: fixes S :: "'a::metric_space set" assumes "complete S" shows "closed S" proof (unfold closed_sequential_limits, clarify) fix f x assume "\n. f n \ S" and "f \ x" from \f \ x\ have "Cauchy f" by (rule LIMSEQ_imp_Cauchy) with \complete S\ and \\n. f n \ S\ obtain l where "l \ S" and "f \ l" by (rule completeE) from \f \ x\ and \f \ l\ have "x = l" by (rule LIMSEQ_unique) with \l \ S\ show "x \ S" by simp qed lemma complete_Int_closed: fixes S :: "'a::metric_space set" assumes "complete S" and "closed t" shows "complete (S \ t)" proof (rule completeI) fix f assume "\n. f n \ S \ t" and "Cauchy f" then have "\n. f n \ S" and "\n. f n \ t" by simp_all from \complete S\ obtain l where "l \ S" and "f \ l" using \\n. f n \ S\ and \Cauchy f\ by (rule completeE) from \closed t\ and \\n. f n \ t\ and \f \ l\ have "l \ t" by (rule closed_sequentially) with \l \ S\ and \f \ l\ show "\l\S \ t. f \ l" by fast qed lemma complete_closed_subset: fixes S :: "'a::metric_space set" assumes "closed S" and "S \ t" and "complete t" shows "complete S" using assms complete_Int_closed [of t S] by (simp add: Int_absorb1) lemma complete_eq_closed: fixes S :: "('a::complete_space) set" shows "complete S \ closed S" proof assume "closed S" then show "complete S" using subset_UNIV complete_UNIV by (rule complete_closed_subset) next assume "complete S" then show "closed S" by (rule complete_imp_closed) qed lemma convergent_eq_Cauchy: fixes S :: "nat \ 'a::complete_space" shows "(\l. (S \ l) sequentially) \ Cauchy S" unfolding Cauchy_convergent_iff convergent_def .. lemma convergent_imp_bounded: fixes S :: "nat \ 'a::metric_space" shows "(S \ l) sequentially \ bounded (range S)" by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) lemma frontier_subset_compact: fixes S :: "'a::heine_borel set" shows "compact S \ frontier S \ S" using frontier_subset_closed compact_eq_bounded_closed by blast lemma continuous_closed_imp_Cauchy_continuous: fixes S :: "('a::complete_space) set" shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially) by (meson LIMSEQ_imp_Cauchy complete_def) lemma banach_fix_type: fixes f::"'a::complete_space\'a" assumes c:"0 \ c" "c < 1" and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" shows "\!x. (f x = x)" using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] by auto subsection\<^marker>\tag unimportant\\ Finite intersection property\ text\Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\ lemma closed_imp_fip: fixes S :: "'a::heine_borel set" assumes "closed S" and T: "T \ \" "bounded T" and clof: "\T. T \ \ \ closed T" and none: "\\'. \finite \'; \' \ \\ \ S \ \\' \ {}" shows "S \ \\ \ {}" proof - have "compact (S \ T)" using \closed S\ clof compact_eq_bounded_closed T by blast then have "(S \ T) \ \\ \ {}" apply (rule compact_imp_fip) apply (simp add: clof) by (metis Int_assoc complete_lattice_class.Inf_insert finite_insert insert_subset none \T \ \\) then show ?thesis by blast qed lemma closed_imp_fip_compact: fixes S :: "'a::heine_borel set" shows "\closed S; \T. T \ \ \ compact T; \\'. \finite \'; \' \ \\ \ S \ \\' \ {}\ \ S \ \\ \ {}" by (metis Inf_greatest closed_imp_fip compact_eq_bounded_closed empty_subsetI finite.emptyI inf.orderE) lemma closed_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes "closed S" "T \ \" "bounded T" and "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" proof - have "UNIV \ \\ \ {}" using assms closed_imp_fip [OF closed_UNIV] by auto then show ?thesis by simp qed lemma compact_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes clof: "\T. T \ \ \ compact T" and none: "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" by (metis InterI all_not_in_conv clof closed_fip_Heine_Borel compact_eq_bounded_closed none) lemma compact_sequence_with_limit: fixes f :: "nat \ 'a::heine_borel" shows "(f \ l) sequentially \ compact (insert l (range f))" apply (simp add: compact_eq_bounded_closed, auto) apply (simp add: convergent_imp_bounded) by (simp add: closed_limpt islimpt_insert sequence_unique_limpt) subsection \Properties of Balls and Spheres\ lemma compact_cball[simp]: fixes x :: "'a::heine_borel" shows "compact (cball x e)" using compact_eq_bounded_closed bounded_cball closed_cball by blast lemma compact_frontier_bounded[intro]: fixes S :: "'a::heine_borel set" shows "bounded S \ compact (frontier S)" unfolding frontier_def using compact_eq_bounded_closed by blast lemma compact_frontier[intro]: fixes S :: "'a::heine_borel set" shows "compact S \ compact (frontier S)" using compact_eq_bounded_closed compact_frontier_bounded by blast subsection \Distance from a Set\ lemma distance_attains_sup: assumes "compact s" "s \ {}" shows "\x\s. \y\s. dist a y \ dist a x" proof (rule continuous_attains_sup [OF assms]) { fix x assume "x\s" have "(dist a \ dist a x) (at x within s)" by (intro tendsto_dist tendsto_const tendsto_ident_at) } then show "continuous_on s (dist a)" unfolding continuous_on .. qed text \For \emph{minimal} distance, we only need closure, not compactness.\ lemma distance_attains_inf: fixes a :: "'a::heine_borel" assumes "closed s" and "s \ {}" obtains x where "x\s" "\y. y \ s \ dist a x \ dist a y" proof - from assms obtain b where "b \ s" by auto let ?B = "s \ cball a (dist b a)" have "?B \ {}" using \b \ s\ by (auto simp: dist_commute) moreover have "continuous_on ?B (dist a)" by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const) moreover have "compact ?B" by (intro closed_Int_compact \closed s\ compact_cball) ultimately obtain x where "x \ ?B" "\y\?B. dist a x \ dist a y" by (metis continuous_attains_inf) with that show ?thesis by fastforce qed subsection \Infimum Distance\ definition\<^marker>\tag important\ "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)" by (auto intro!: zero_le_dist) lemma infdist_notempty: "A \ {} \ infdist x A = (INF a\A. dist x a)" by (simp add: infdist_def) lemma infdist_nonneg: "0 \ infdist x A" by (auto simp: infdist_def intro: cINF_greatest) lemma infdist_le: "a \ A \ infdist x A \ dist x a" by (auto intro: cINF_lower simp add: infdist_def) lemma infdist_le2: "a \ A \ dist x a \ d \ infdist x A \ d" by (auto intro!: cINF_lower2 simp add: infdist_def) lemma infdist_zero[simp]: "a \ A \ infdist a A = 0" by (auto intro!: antisym infdist_nonneg infdist_le2) lemma infdist_Un_min: assumes "A \ {}" "B \ {}" shows "infdist x (A \ B) = min (infdist x A) (infdist x B)" using assms by (simp add: infdist_def cINF_union inf_real_def) lemma infdist_triangle: "infdist x A \ infdist y A + dist x y" proof (cases "A = {}") case True then show ?thesis by (simp add: infdist_def) next case False then obtain a where "a \ A" by auto have "infdist x A \ Inf {dist x y + dist y a |a. a \ A}" proof (rule cInf_greatest) from \A \ {}\ show "{dist x y + dist y a |a. a \ A} \ {}" by simp fix d assume "d \ {dist x y + dist y a |a. a \ A}" then obtain a where d: "d = dist x y + dist y a" "a \ A" by auto show "infdist x A \ d" unfolding infdist_notempty[OF \A \ {}\] proof (rule cINF_lower2) show "a \ A" by fact show "dist x a \ d" unfolding d by (rule dist_triangle) qed simp qed also have "\ = dist x y + infdist y A" proof (rule cInf_eq, safe) fix a assume "a \ A" then show "dist x y + infdist y A \ dist x y + dist y a" by (auto intro: infdist_le) next fix i assume inf: "\d. d \ {dist x y + dist y a |a. a \ A} \ i \ d" then have "i - dist x y \ infdist y A" unfolding infdist_notempty[OF \A \ {}\] using \a \ A\ by (intro cINF_greatest) (auto simp: field_simps) then show "i \ dist x y + infdist y A" by simp qed finally show ?thesis by simp qed lemma infdist_triangle_abs: "\infdist x A - infdist y A\ \ dist x y" by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle) lemma in_closure_iff_infdist_zero: assumes "A \ {}" shows "x \ closure A \ infdist x A = 0" proof assume "x \ closure A" show "infdist x A = 0" proof (rule ccontr) assume "infdist x A \ 0" with infdist_nonneg[of x A] have "infdist x A > 0" by auto then have "ball x (infdist x A) \ closure A = {}" apply auto apply (metis \x \ closure A\ closure_approachable dist_commute infdist_le not_less) done then have "x \ closure A" by (metis \0 < infdist x A\ centre_in_ball disjoint_iff_not_equal) then show False using \x \ closure A\ by simp qed next assume x: "infdist x A = 0" then obtain a where "a \ A" by atomize_elim (metis all_not_in_conv assms) show "x \ closure A" unfolding closure_approachable apply safe proof (rule ccontr) fix e :: real assume "e > 0" assume "\ (\y\A. dist y x < e)" then have "infdist x A \ e" using \a \ A\ unfolding infdist_def by (force simp: dist_commute intro: cINF_greatest) with x \e > 0\ show False by auto qed qed lemma in_closed_iff_infdist_zero: assumes "closed A" "A \ {}" shows "x \ A \ infdist x A = 0" proof - have "x \ closure A \ infdist x A = 0" by (rule in_closure_iff_infdist_zero) fact with assms show ?thesis by simp qed lemma infdist_pos_not_in_closed: assumes "closed S" "S \ {}" "x \ S" shows "infdist x S > 0" using in_closed_iff_infdist_zero[OF assms(1) assms(2), of x] assms(3) infdist_nonneg le_less by fastforce lemma infdist_attains_inf: fixes X::"'a::heine_borel set" assumes "closed X" assumes "X \ {}" obtains x where "x \ X" "infdist y X = dist y x" proof - have "bdd_below (dist y ` X)" by auto from distance_attains_inf[OF assms, of y] obtain x where INF: "x \ X" "\z. z \ X \ dist y x \ dist y z" by auto have "infdist y X = dist y x" by (auto simp: infdist_def assms intro!: antisym cINF_lower[OF _ \x \ X\] cINF_greatest[OF assms(2) INF(2)]) with \x \ X\ show ?thesis .. qed text \Every metric space is a T4 space:\ instance metric_space \ t4_space proof fix S T::"'a set" assume H: "closed S" "closed T" "S \ T = {}" consider "S = {}" | "T = {}" | "S \ {} \ T \ {}" by auto then show "\U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" proof (cases) case 1 show ?thesis apply (rule exI[of _ "{}"], rule exI[of _ UNIV]) using 1 by auto next case 2 show ?thesis apply (rule exI[of _ UNIV], rule exI[of _ "{}"]) using 2 by auto next case 3 define U where "U = (\x\S. ball x ((infdist x T)/2))" have A: "open U" unfolding U_def by auto have "infdist x T > 0" if "x \ S" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have B: "S \ U" unfolding U_def by auto define V where "V = (\x\T. ball x ((infdist x S)/2))" have C: "open V" unfolding V_def by auto have "infdist x S > 0" if "x \ T" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have D: "T \ V" unfolding V_def by auto have "(ball x ((infdist x T)/2)) \ (ball y ((infdist y S)/2)) = {}" if "x \ S" "y \ T" for x y proof (auto) fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S" have "2 * dist x y \ 2 * dist x z + 2 * dist y z" by metric also have "... < infdist x T + infdist y S" using H by auto finally have "dist x y < infdist x T \ dist x y < infdist y S" by auto then show False using infdist_le[OF \x \ S\, of y] infdist_le[OF \y \ T\, of x] by (auto simp add: dist_commute) qed then have E: "U \ V = {}" unfolding U_def V_def by auto show ?thesis apply (rule exI[of _ U], rule exI[of _ V]) using A B C D E by auto qed qed lemma tendsto_infdist [tendsto_intros]: assumes f: "(f \ l) F" shows "((\x. infdist (f x) A) \ infdist l A) F" proof (rule tendstoI) fix e ::real assume "e > 0" from tendstoD[OF f this] show "eventually (\x. dist (infdist (f x) A) (infdist l A) < e) F" proof (eventually_elim) fix x from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] have "dist (infdist (f x) A) (infdist l A) \ dist (f x) l" by (simp add: dist_commute dist_real_def) also assume "dist (f x) l < e" finally show "dist (infdist (f x) A) (infdist l A) < e" . qed qed lemma continuous_infdist[continuous_intros]: assumes "continuous F f" shows "continuous F (\x. infdist (f x) A)" using assms unfolding continuous_def by (rule tendsto_infdist) lemma continuous_on_infdist [continuous_intros]: assumes "continuous_on S f" shows "continuous_on S (\x. infdist (f x) A)" using assms unfolding continuous_on by (auto intro: tendsto_infdist) lemma compact_infdist_le: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "e > 0" shows "compact {x. infdist x A \ e}" proof - from continuous_closed_vimage[of "{0..e}" "\x. infdist x A"] continuous_infdist[OF continuous_ident, of _ UNIV A] have "closed {x. infdist x A \ e}" by (auto simp: vimage_def infdist_nonneg) moreover from assms obtain x0 b where b: "\x. x \ A \ dist x0 x \ b" "closed A" by (auto simp: compact_eq_bounded_closed bounded_def) { fix y assume "infdist y A \ e" moreover from infdist_attains_inf[OF \closed A\ \A \ {}\, of y] obtain z where "z \ A" "infdist y A = dist y z" by blast ultimately have "dist x0 y \ b + e" using b by metric } then have "bounded {x. infdist x A \ e}" by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"]) ultimately show "compact {x. infdist x A \ e}" by (simp add: compact_eq_bounded_closed) qed subsection \Separation between Points and Sets\ proposition separate_point_closed: fixes s :: "'a::heine_borel set" assumes "closed s" and "a \ s" shows "\d>0. \x\s. d \ dist a x" proof (cases "s = {}") case True then show ?thesis by(auto intro!: exI[where x=1]) next case False from assms obtain x where "x\s" "\y\s. dist a x \ dist a y" using \s \ {}\ by (blast intro: distance_attains_inf [of s a]) with \x\s\ show ?thesis using dist_pos_lt[of a x] and\a \ s\ by blast qed proposition separate_compact_closed: fixes s t :: "'a::heine_borel set" assumes "compact s" and t: "closed t" "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof cases assume "s \ {} \ t \ {}" then have "s \ {}" "t \ {}" by auto let ?inf = "\x. infdist x t" have "continuous_on s ?inf" by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) then obtain x where x: "x \ s" "\y\s. ?inf x \ ?inf y" using continuous_attains_inf[OF \compact s\ \s \ {}\] by auto then have "0 < ?inf x" using t \t \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) moreover have "\x'\s. \y\t. ?inf x \ dist x' y" using x by (auto intro: order_trans infdist_le) ultimately show ?thesis by auto qed (auto intro!: exI[of _ 1]) proposition separate_closed_compact: fixes s t :: "'a::heine_borel set" assumes "closed s" and "compact t" and "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof - have *: "t \ s = {}" using assms(3) by auto show ?thesis using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute) qed proposition compact_in_open_separated: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "open B" assumes "A \ B" obtains e where "e > 0" "{x. infdist x A \ e} \ B" proof atomize_elim have "closed (- B)" "compact A" "- B \ A = {}" using assms by (auto simp: open_Diff compact_eq_bounded_closed) from separate_closed_compact[OF this] obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" by auto define d where "d = d' / 2" hence "d>0" "d < d'" using d' by auto with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" by force show "\e>0. {x. infdist x A \ e} \ B" proof (rule ccontr) assume "\e. 0 < e \ {x. infdist x A \ e} \ B" with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" by auto from assms have "closed A" "A \ {}" by (auto simp: compact_eq_bounded_closed) from infdist_attains_inf[OF this] obtain y where y: "y \ A" "infdist x A = dist x y" by auto have "dist x y \ d" using x y by simp also have "\ < dist x y" using y d x by auto finally show False by simp qed qed subsection \Uniform Continuity\ lemma uniformly_continuous_onE: assumes "uniformly_continuous_on s f" "0 < e" obtains d where "d>0" "\x x'. \x\s; x'\s; dist x' x < d\ \ dist (f x') (f x) < e" using assms by (auto simp: uniformly_continuous_on_def) lemma uniformly_continuous_on_sequentially: "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ (\n. dist (x n) (y n)) \ 0 \ (\n. dist (f(x n)) (f(y n))) \ 0)" (is "?lhs = ?rhs") proof assume ?lhs { fix x y assume x: "\n. x n \ s" and y: "\n. y n \ s" and xy: "((\n. dist (x n) (y n)) \ 0) sequentially" { fix e :: real assume "e > 0" then obtain d where "d > 0" and d: "\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using \?lhs\[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto obtain N where N: "\n\N. dist (x n) (y n) < d" using xy[unfolded lim_sequentially dist_norm] and \d>0\ by auto { fix n assume "n\N" then have "dist (f (x n)) (f (y n)) < e" using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y by (simp add: dist_commute) } then have "\N. \n\N. dist (f (x n)) (f (y n)) < e" by auto } then have "((\n. dist (f(x n)) (f(y n))) \ 0) sequentially" unfolding lim_sequentially and dist_real_def by auto } then show ?rhs by auto next assume ?rhs { assume "\ ?lhs" then obtain e where "e > 0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto then obtain fa where fa: "\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def by (auto simp: dist_commute) define x where "x n = fst (fa (inverse (real n + 1)))" for n define y where "y n = snd (fa (inverse (real n + 1)))" for n have xyn: "\n. x n \ s \ y n \ s" and xy0: "\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" unfolding x_def and y_def using fa by auto { fix e :: real assume "e > 0" then obtain N :: nat where "N \ 0" and N: "0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inverse[of e] by auto { fix n :: nat assume "n \ N" then have "inverse (real n + 1) < inverse (real N)" using of_nat_0_le_iff and \N\0\ by auto also have "\ < e" using N by auto finally have "inverse (real n + 1) < e" by auto then have "dist (x n) (y n) < e" using xy0[THEN spec[where x=n]] by auto } then have "\N. \n\N. dist (x n) (y n) < e" by auto } then have "\e>0. \N. \n\N. dist (f (x n)) (f (y n)) < e" using \?rhs\[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding lim_sequentially dist_real_def by auto then have False using fxy and \e>0\ by auto } then show ?lhs unfolding uniformly_continuous_on_def by blast qed subsection \Continuity on a Compact Domain Implies Uniform Continuity\ text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ lemma Heine_Borel_lemma: assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" proof - have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" proof - have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n using neg by simp then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" by metis then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" using \compact S\ compact_def that by metis then obtain G where "l \ G" "G \ \" using Ssub by auto then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" using opn open_dist by blast obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" using to_l apply (simp add: lim_sequentially) using \0 < e\ half_gt_zero that by blast obtain N2 where N2: "of_nat N2 > 2/e" using reals_Archimedean2 by blast obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" using fG [OF \G \ \\, of "r (max N1 N2)"] by blast then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" by simp also have "... \ 1 / real (Suc (max N1 N2))" apply (simp add: field_split_simps del: max.bounded_iff) using \strict_mono r\ seq_suble by blast also have "... \ 1 / real (Suc N2)" by (simp add: field_simps) also have "... < e/2" using N2 \0 < e\ by (simp add: field_simps) finally have "dist (f (r (max N1 N2))) x < e / 2" . moreover have "dist (f (r (max N1 N2))) l < e/2" using N1 max.cobounded1 by blast ultimately have "dist x l < e" by metric then show ?thesis using e \x \ G\ by blast qed then show ?thesis by (meson that) qed lemma compact_uniformly_equicontinuous: assumes "compact S" and cont: "\x e. \x \ S; 0 < e\ \ \d. 0 < d \ (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" and "0 < e" obtains d where "0 < d" "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" proof - obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" using cont by metis let ?\ = "((\x. ball x (d x (e / 2))) ` S)" have Ssub: "S \ \ ?\" by clarsimp (metis d_pos \0 < e\ dist_self half_gt_zero_iff) then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" by (rule Heine_Borel_lemma [OF \compact S\]) auto moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v proof - obtain G where "G \ ?\" "u \ G" "v \ G" using k that by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \ S" by auto with that d_dist have "dist (f w) (f v) < e/2" by (metis \0 < e\ dist_commute half_gt_zero) moreover have "dist (f w) (f u) < e/2" using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) ultimately show ?thesis using dist_triangle_half_r by blast qed ultimately show ?thesis using that by blast qed corollary compact_uniformly_continuous: fixes f :: "'a :: metric_space \ 'b :: metric_space" assumes f: "continuous_on S f" and S: "compact S" shows "uniformly_continuous_on S f" using f unfolding continuous_on_iff uniformly_continuous_on_def by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) subsection\<^marker>\tag unimportant\\ Theorems relating continuity and uniform continuity to closures\ lemma continuous_on_closure: "continuous_on (closure S) f \ (\x e. x \ closure S \ 0 < e \ (\d. 0 < d \ (\y. y \ S \ dist y x < d \ dist (f y) (f x) < e)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_on_iff by (metis Un_iff closure_def) next assume R [rule_format]: ?rhs show ?lhs proof fix x and e::real assume "0 < e" and x: "x \ closure S" obtain \::real where "\ > 0" and \: "\y. \y \ S; dist y x < \\ \ dist (f y) (f x) < e/2" using R [of x "e/2"] \0 < e\ x by auto have "dist (f y) (f x) \ e" if y: "y \ closure S" and dyx: "dist y x < \/2" for y proof - obtain \'::real where "\' > 0" and \': "\z. \z \ S; dist z y < \'\ \ dist (f z) (f y) < e/2" using R [of y "e/2"] \0 < e\ y by auto obtain z where "z \ S" and z: "dist z y < min \' \ / 2" using closure_approachable y by (metis \0 < \'\ \0 < \\ divide_pos_pos min_less_iff_conj zero_less_numeral) have "dist (f z) (f y) < e/2" using \' [OF \z \ S\] z \0 < \'\ by metric moreover have "dist (f z) (f x) < e/2" using \[OF \z \ S\] z dyx by metric ultimately show ?thesis by metric qed then show "\d>0. \x'\closure S. dist x' x < d \ dist (f x') (f x) \ e" by (rule_tac x="\/2" in exI) (simp add: \\ > 0\) qed qed lemma continuous_on_closure_sequentially: fixes f :: "'a::metric_space \ 'b :: metric_space" shows "continuous_on (closure S) f \ (\x a. a \ closure S \ (\n. x n \ S) \ x \ a \ (f \ x) \ f a)" (is "?lhs = ?rhs") proof - have "continuous_on (closure S) f \ (\x \ closure S. continuous (at x within S) f)" by (force simp: continuous_on_closure continuous_within_eps_delta) also have "... = ?rhs" by (force simp: continuous_within_sequentially) finally show ?thesis . qed lemma uniformly_continuous_on_closure: fixes f :: "'a::metric_space \ 'b::metric_space" assumes ucont: "uniformly_continuous_on S f" and cont: "continuous_on (closure S) f" shows "uniformly_continuous_on (closure S) f" unfolding uniformly_continuous_on_def proof (intro allI impI) fix e::real assume "0 < e" then obtain d::real where "d>0" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (f x') (f x) < e/3" using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto show "\d>0. \x\closure S. \x'\closure S. dist x' x < d \ dist (f x') (f x) < e" proof (rule exI [where x="d/3"], clarsimp simp: \d > 0\) fix x y assume x: "x \ closure S" and y: "y \ closure S" and dyx: "dist y x * 3 < d" obtain d1::real where "d1 > 0" and d1: "\w. \w \ closure S; dist w x < d1\ \ dist (f w) (f x) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \0 < e\ x by auto obtain x' where "x' \ S" and x': "dist x' x < min d1 (d / 3)" using closure_approachable [of x S] by (metis \0 < d1\ \0 < d\ divide_pos_pos min_less_iff_conj x zero_less_numeral) obtain d2::real where "d2 > 0" and d2: "\w \ closure S. dist w y < d2 \ dist (f w) (f y) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \0 < e\ y by auto obtain y' where "y' \ S" and y': "dist y' y < min d2 (d / 3)" using closure_approachable [of y S] by (metis \0 < d2\ \0 < d\ divide_pos_pos min_less_iff_conj y zero_less_numeral) have "dist x' x < d/3" using x' by auto then have "dist x' y' < d" using dyx y' by metric then have "dist (f x') (f y') < e/3" by (rule d [OF \y' \ S\ \x' \ S\]) moreover have "dist (f x') (f x) < e/3" using \x' \ S\ closure_subset x' d1 by (simp add: closure_def) moreover have "dist (f y') (f y) < e/3" using \y' \ S\ closure_subset y' d2 by (simp add: closure_def) ultimately show "dist (f y) (f x) < e" by metric qed qed lemma uniformly_continuous_on_extension_at_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" assumes "x \ closure X" obtains l where "(f \ l) (at x within X)" proof - from assms obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] obtain l where l: "(\n. f (xs n)) \ l" by atomize_elim (simp only: convergent_eq_Cauchy) have "(f \ l) (at x within X)" proof (safe intro!: Lim_within_LIMSEQ) fix xs' assume "\n. xs' n \ x \ xs' n \ X" and xs': "xs' \ x" then have "xs' n \ x" "xs' n \ X" for n by auto from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \xs' \ x\ \xs' _ \ X\] obtain l' where l': "(\n. f (xs' n)) \ l'" by atomize_elim (simp only: convergent_eq_Cauchy) show "(\n. f (xs' n)) \ l" proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' \ e / 2" have "e' > 0" using \e > 0\ by (simp add: e'_def) have "\\<^sub>F n in sequentially. dist (f (xs n)) l < e'" by (simp add: \0 < e'\ l tendstoD) moreover from uc[unfolded uniformly_continuous_on_def, rule_format, OF \e' > 0\] obtain d where d: "d > 0" "\x x'. x \ X \ x' \ X \ dist x x' < d \ dist (f x) (f x') < e'" by auto have "\\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" by (auto intro!: \0 < d\ order_tendstoD tendsto_eq_intros xs xs') ultimately show "\\<^sub>F n in sequentially. dist (f (xs' n)) l < e" proof eventually_elim case (elim n) have "dist (f (xs' n)) l \ dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" by metric also have "dist (f (xs n)) (f (xs' n)) < e'" by (auto intro!: d xs \xs' _ \ _\ elim) also note \dist (f (xs n)) l < e'\ also have "e' + e' = e" by (simp add: e'_def) finally show ?case by simp qed qed qed thus ?thesis .. qed lemma uniformly_continuous_on_extension_on_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" obtains g where "uniformly_continuous_on (closure X) g" "\x. x \ X \ f x = g x" "\Y h x. X \ Y \ Y \ closure X \ continuous_on Y h \ (\x. x \ X \ f x = h x) \ x \ Y \ h x = g x" proof - from uc have cont_f: "continuous_on X f" by (simp add: uniformly_continuous_imp_continuous) obtain y where y: "(f \ y x) (at x within X)" if "x \ closure X" for x apply atomize_elim apply (rule choice) using uniformly_continuous_on_extension_at_closure[OF assms] by metis let ?g = "\x. if x \ X then f x else y x" have "uniformly_continuous_on (closure X) ?g" unfolding uniformly_continuous_on_def proof safe fix e::real assume "e > 0" define e' where "e' \ e / 3" have "e' > 0" using \e > 0\ by (simp add: e'_def) from uc[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] obtain d where "d > 0" and d: "\x x'. x \ X \ x' \ X \ dist x' x < d \ dist (f x') (f x) < e'" by auto define d' where "d' = d / 3" have "d' > 0" using \d > 0\ by (simp add: d'_def) show "\d>0. \x\closure X. \x'\closure X. dist x' x < d \ dist (?g x') (?g x) < e" proof (safe intro!: exI[where x=d'] \d' > 0\) fix x x' assume x: "x \ closure X" and x': "x' \ closure X" and dist: "dist x' x < d'" then obtain xs xs' where xs: "xs \ x" "\n. xs n \ X" and xs': "xs' \ x'" "\n. xs' n \ X" by (auto simp: closure_sequential) have "\\<^sub>F n in sequentially. dist (xs' n) x' < d'" and "\\<^sub>F n in sequentially. dist (xs n) x < d'" by (auto intro!: \0 < d'\ order_tendstoD tendsto_eq_intros xs xs') moreover have "(\x. f (xs x)) \ y x" if "x \ closure X" "x \ X" "xs \ x" "\n. xs n \ X" for xs x using that not_eventuallyD by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at) then have "(\x. f (xs' x)) \ ?g x'" "(\x. f (xs x)) \ ?g x" using x x' by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) then have "\\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" "\\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" by (auto intro!: \0 < e'\ order_tendstoD tendsto_eq_intros) ultimately have "\\<^sub>F n in sequentially. dist (?g x') (?g x) < e" proof eventually_elim case (elim n) have "dist (?g x') (?g x) \ dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) also from \dist (xs' n) x' < d'\ \dist x' x < d'\ \dist (xs n) x < d'\ have "dist (xs' n) (xs n) < d" unfolding d'_def by metric with \xs _ \ X\ \xs' _ \ X\ have "dist (f (xs' n)) (f (xs n)) < e'" by (rule d) also note \dist (f (xs' n)) (?g x') < e'\ also note \dist (f (xs n)) (?g x) < e'\ finally show ?case by (simp add: e'_def) qed then show "dist (?g x') (?g x) < e" by simp qed qed moreover have "f x = ?g x" if "x \ X" for x using that by simp moreover { fix Y h x assume Y: "x \ Y" "X \ Y" "Y \ closure X" and cont_h: "continuous_on Y h" and extension: "(\x. x \ X \ f x = h x)" { assume "x \ X" have "x \ closure X" using Y by auto then obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y have hx: "(\x. f (xs x)) \ h x" by (auto simp: subsetD extension) then have "(\x. f (xs x)) \ y x" using \x \ X\ not_eventuallyD xs(2) by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at xs) with hx have "h x = y x" by (rule LIMSEQ_unique) } then have "h x = ?g x" using extension by auto } ultimately show ?thesis .. qed lemma bounded_uniformly_continuous_image: fixes f :: "'a :: heine_borel \ 'b :: heine_borel" assumes "uniformly_continuous_on S f" "bounded S" shows "bounded(f ` S)" by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure) subsection \With Abstract Topology (TODO: move and remove dependency?)\ lemma openin_contains_ball: "openin (top_of_set t) s \ s \ t \ (\x \ s. \e. 0 < e \ ball x e \ t \ s)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: openin_open) apply (metis Int_commute Int_mono inf.cobounded2 open_contains_ball order_refl subsetCE) done next assume ?rhs then show ?lhs apply (simp add: openin_euclidean_subtopology_iff) by (metis (no_types) Int_iff dist_commute inf.absorb_iff2 mem_ball) qed lemma openin_contains_cball: "openin (top_of_set t) s \ s \ t \ (\x \ s. \e. 0 < e \ cball x e \ t \ s)" apply (simp add: openin_contains_ball) apply (rule iffI) apply (auto dest!: bspec) apply (rule_tac x="e/2" in exI, force+) done subsection \Closed Nest\ text \Bounded closed nest property (proof does not use Heine-Borel)\ lemma bounded_closed_nest: fixes S :: "nat \ ('a::heine_borel) set" assumes "\n. closed (S n)" and "\n. S n \ {}" and "\m n. m \ n \ S n \ S m" and "bounded (S 0)" obtains a where "\n. a \ S n" proof - from assms(2) obtain x where x: "\n. x n \ S n" using choice[of "\n x. x \ S n"] by auto from assms(4,1) have "seq_compact (S 0)" by (simp add: bounded_closed_imp_seq_compact) then obtain l r where lr: "l \ S 0" "strict_mono r" "(x \ r) \ l" using x and assms(3) unfolding seq_compact_def by blast have "\n. l \ S n" proof fix n :: nat have "closed (S n)" using assms(1) by simp moreover have "\i. (x \ r) i \ S i" using x and assms(3) and lr(2) [THEN seq_suble] by auto then have "\i. (x \ r) (i + n) \ S n" using assms(3) by (fast intro!: le_add2) moreover have "(\i. (x \ r) (i + n)) \ l" using lr(3) by (rule LIMSEQ_ignore_initial_segment) ultimately show "l \ S n" by (rule closed_sequentially) qed then show ?thesis using that by blast qed text \Decreasing case does not even need compactness, just completeness.\ lemma decreasing_closed_nest: fixes S :: "nat \ ('a::complete_space) set" assumes "\n. closed (S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x\S n. \y\S n. dist x y < e" obtains a where "\n. a \ S n" proof - have "\n. \x. x \ S n" using assms(2) by auto then have "\t. \n. t n \ S n" using choice[of "\n x. x \ S n"] by auto then obtain t where t: "\n. t n \ S n" by auto { fix e :: real assume "e > 0" then obtain N where N: "\x\S N. \y\S N. dist x y < e" using assms(4) by blast { fix m n :: nat assume "N \ m \ N \ n" then have "t m \ S N" "t n \ S N" using assms(3) t unfolding subset_eq t by blast+ then have "dist (t m) (t n) < e" using N by auto } then have "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto } then have "Cauchy t" unfolding cauchy_def by auto then obtain l where l:"(t \ l) sequentially" using complete_UNIV unfolding complete_def by auto { fix n :: nat { fix e :: real assume "e > 0" then obtain N :: nat where N: "\n\N. dist (t n) l < e" using l[unfolded lim_sequentially] by auto have "t (max n N) \ S n" by (meson assms(3) contra_subsetD max.cobounded1 t) then have "\y\S n. dist y l < e" using N max.cobounded2 by blast } then have "l \ S n" using closed_approachable[of "S n" l] assms(1) by auto } then show ?thesis using that by blast qed text \Strengthen it to the intersection actually being a singleton.\ lemma decreasing_closed_nest_sing: fixes S :: "nat \ 'a::complete_space set" assumes "\n. closed(S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x \ (S n). \ y\(S n). dist x y < e" shows "\a. \(range S) = {a}" proof - obtain a where a: "\n. a \ S n" using decreasing_closed_nest[of S] using assms by auto { fix b assume b: "b \ \(range S)" { fix e :: real assume "e > 0" then have "dist a b < e" using assms(4) and b and a by blast } then have "dist a b = 0" by (metis dist_eq_0_iff dist_nz less_le) } with a have "\(range S) = {a}" unfolding image_def by auto then show ?thesis .. qed subsection\<^marker>\tag unimportant\ \Making a continuous function avoid some value in a neighbourhood\ lemma continuous_within_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x within s) f" and "f x \ a" shows "\e>0. \y \ s. dist x y < e --> f y \ a" proof - obtain U where "open U" and "f x \ U" and "a \ U" using t1_space [OF \f x \ a\] by fast have "(f \ f x) (at x within s)" using assms(1) by (simp add: continuous_within) then have "eventually (\y. f y \ U) (at x within s)" using \open U\ and \f x \ U\ unfolding tendsto_def by fast then have "eventually (\y. f y \ a) (at x within s)" using \a \ U\ by (fast elim: eventually_mono) then show ?thesis using \f x \ a\ by (auto simp: dist_commute eventually_at) qed lemma continuous_at_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x) f" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms continuous_within_avoid[of x UNIV f a] by simp lemma continuous_on_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "x \ s" and "f x \ a" shows "\e>0. \y \ s. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] using assms(3) by auto lemma continuous_on_open_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "open s" and "x \ s" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] using continuous_at_avoid[of x f a] assms(4) by auto subsection \Consequences for Real Numbers\ lemma closed_contains_Inf: fixes S :: "real set" shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" by (metis closure_contains_Inf closure_closed) lemma closed_subset_contains_Inf: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" by (metis closure_contains_Inf closure_minimal subset_eq) lemma closed_contains_Sup: fixes S :: "real set" shows "S \ {} \ bdd_above S \ closed S \ Sup S \ S" by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) lemma closed_subset_contains_Sup: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_above A \ Sup A \ C" by (metis closure_contains_Sup closure_minimal subset_eq) lemma atLeastAtMost_subset_contains_Inf: fixes A :: "real set" and a b :: real shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" by (rule closed_subset_contains_Inf) (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" by (simp add: bounded_iff) lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" by (auto simp: bounded_def bdd_above_def dist_real_def) (metis abs_le_D1 abs_minus_commute diff_le_eq) lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" by (auto simp: bounded_def bdd_below_def dist_real_def) (metis abs_le_D1 add.commute diff_le_eq) lemma bounded_has_Sup: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Sup S" and "\b. (\x\S. x \ b) \ Sup S \ b" proof show "\b. (\x\S. x \ b) \ Sup S \ b" using assms by (metis cSup_least) qed (metis cSup_upper assms(1) bounded_imp_bdd_above) lemma Sup_insert: fixes S :: "real set" shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) lemma bounded_has_Inf: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Inf S" and "\b. (\x\S. x \ b) \ Inf S \ b" proof show "\b. (\x\S. x \ b) \ Inf S \ b" using assms by (metis cInf_greatest) qed (metis cInf_lower assms(1) bounded_imp_bdd_below) lemma Inf_insert: fixes S :: "real set" shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) lemma open_real: fixes s :: "real set" shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" unfolding open_dist dist_norm by simp lemma islimpt_approachable_real: fixes s :: "real set" shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" unfolding islimpt_approachable dist_norm by simp lemma closed_real: fixes s :: "real set" shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" unfolding closed_limpt islimpt_approachable dist_norm by simp lemma continuous_at_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" unfolding continuous_at unfolding Lim_at unfolding dist_norm apply auto apply (erule_tac x=e in allE, auto) apply (rule_tac x=d in exI, auto) apply (erule_tac x=x' in allE, auto) apply (erule_tac x=e in allE, auto) done lemma continuous_on_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" unfolding continuous_on_iff dist_norm by simp lemma continuous_on_closed_Collect_le: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" shows "closed {x \ s. f x \ g x}" proof - have "closed ((\x. g x - f x) -` {0..} \ s)" using closed_real_atLeast continuous_on_diff [OF g f] by (simp add: continuous_on_closed_vimage [OF s]) also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" by auto finally show ?thesis . qed lemma continuous_le_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. x \ a}" ] assms continuous_on_closed_Collect_le[of "UNIV" "\x. x" "\x. a"] by auto lemma continuous_ge_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. a \ x}"] assms continuous_on_closed_Collect_le[of "UNIV" "\x. a" "\x. x"] by auto subsection\The infimum of the distance between two sets\ definition\<^marker>\tag important\ setdist :: "'a::metric_space set \ 'a set \ real" where "setdist s t \ (if s = {} \ t = {} then 0 else Inf {dist x y| x y. x \ s \ y \ t})" lemma setdist_empty1 [simp]: "setdist {} t = 0" by (simp add: setdist_def) lemma setdist_empty2 [simp]: "setdist t {} = 0" by (simp add: setdist_def) lemma setdist_pos_le [simp]: "0 \ setdist s t" by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) lemma le_setdistI: assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" shows "d \ setdist s t" using assms by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" unfolding setdist_def by (auto intro!: bdd_belowI [where m=0] cInf_lower) lemma le_setdist_iff: "d \ setdist s t \ (\x \ s. \y \ t. d \ dist x y) \ (s = {} \ t = {} \ d \ 0)" apply (cases "s = {} \ t = {}") apply (force simp add: setdist_def) apply (intro iffI conjI) using setdist_le_dist apply fastforce apply (auto simp: intro: le_setdistI) done lemma setdist_ltE: assumes "setdist s t < b" "s \ {}" "t \ {}" obtains x y where "x \ s" "y \ t" "dist x y < b" using assms by (auto simp: not_le [symmetric] le_setdist_iff) lemma setdist_refl: "setdist s s = 0" apply (cases "s = {}") apply (force simp add: setdist_def) apply (rule antisym [OF _ setdist_pos_le]) apply (metis all_not_in_conv dist_self setdist_le_dist) done lemma setdist_sym: "setdist s t = setdist t s" by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) lemma setdist_triangle: "setdist s t \ setdist s {a} + setdist {a} t" proof (cases "s = {} \ t = {}") case True then show ?thesis using setdist_pos_le by fastforce next case False have "\x. x \ s \ setdist s t - dist x a \ setdist {a} t" apply (rule le_setdistI, blast) using False apply (fastforce intro: le_setdistI) apply (simp add: algebra_simps) apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) done then have "setdist s t - setdist {a} t \ setdist s {a}" using False by (fastforce intro: le_setdistI) then show ?thesis by (simp add: algebra_simps) qed lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" by (simp add: setdist_def) lemma setdist_Lipschitz: "\setdist {x} s - setdist {y} s\ \ dist x y" apply (subst setdist_singletons [symmetric]) by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} s))" by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma continuous_on_setdist [continuous_intros]: "continuous_on t (\y. (setdist {y} s))" by (metis continuous_at_setdist continuous_at_imp_continuous_on) lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\y. (setdist {y} s))" by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma setdist_subset_right: "\t \ {}; t \ u\ \ setdist s u \ setdist s t" apply (cases "s = {} \ u = {}", force) apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) done lemma setdist_subset_left: "\s \ {}; s \ t\ \ setdist t u \ setdist s u" by (metis setdist_subset_right setdist_sym) lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t" proof (cases "s = {} \ t = {}") case True then show ?thesis by force next case False { fix y assume "y \ t" have "continuous_on (closure s) (\a. dist a y)" by (auto simp: continuous_intros dist_norm) then have *: "\x. x \ closure s \ setdist s t \ dist x y" apply (rule continuous_ge_on_closure) apply assumption apply (blast intro: setdist_le_dist \y \ t\ ) done } note * = this show ?thesis apply (rule antisym) using False closure_subset apply (blast intro: setdist_subset_left) using False * apply (force intro!: le_setdistI) done qed lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s" by (metis setdist_closure_1 setdist_sym) lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" by (metis antisym dist_self setdist_le_dist setdist_pos_le) lemma setdist_unique: "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ \ setdist S T = dist a b" by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" using setdist_subset_left by auto lemma infdist_eq_setdist: "infdist x A = setdist {x} A" by (simp add: infdist_def setdist_def Setcompr_eq_image) lemma setdist_eq_infdist: "setdist A B = (if A = {} then 0 else INF a\A. infdist a B)" proof - have "Inf {dist x y |x y. x \ A \ y \ B} = (INF x\A. Inf (dist x ` B))" if "b \ B" "a \ A" for a b proof (rule order_antisym) have "Inf {dist x y |x y. x \ A \ y \ B} \ Inf (dist x ` B)" if "b \ B" "a \ A" "x \ A" for x proof - have *: "\b'. b' \ B \ Inf {dist x y |x y. x \ A \ y \ B} \ dist x b'" by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3)) show ?thesis using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: *)+ qed then show "Inf {dist x y |x y. x \ A \ y \ B} \ (INF x\A. Inf (dist x ` B))" using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: bdd_below_def) next have *: "\x y. \b \ B; a \ A; x \ A; y \ B\ \ \a\A. Inf (dist a ` B) \ dist x y" by (meson bdd_below_image_dist cINF_lower) show "(INF x\A. Inf (dist x ` B)) \ Inf {dist x y |x y. x \ A \ y \ B}" proof (rule conditionally_complete_lattice_class.cInf_mono) show "bdd_below ((\x. Inf (dist x ` B)) ` A)" by (metis (no_types, lifting) bdd_belowI2 ex_in_conv infdist_def infdist_nonneg that(1)) qed (use that in \auto simp: *\) qed then show ?thesis by (auto simp: setdist_def infdist_def) qed lemma infdist_mono: assumes "A \ B" "A \ {}" shows "infdist x B \ infdist x A" by (simp add: assms infdist_eq_setdist setdist_subset_right) lemma infdist_singleton [simp]: "infdist x {y} = dist x y" by (simp add: infdist_eq_setdist) proposition setdist_attains_inf: assumes "compact B" "B \ {}" obtains y where "y \ B" "setdist A B = infdist y A" proof (cases "A = {}") case True then show thesis by (metis assms diameter_compact_attained infdist_def setdist_def that) next case False obtain y where "y \ B" and min: "\y'. y' \ B \ infdist y A \ infdist y' A" by (metis continuous_attains_inf [OF assms continuous_on_infdist] continuous_on_id) show thesis proof have "setdist A B = (INF y\B. infdist y A)" by (metis \B \ {}\ setdist_eq_infdist setdist_sym) also have "\ = infdist y A" proof (rule order_antisym) show "(INF y\B. infdist y A) \ infdist y A" proof (rule cInf_lower) show "infdist y A \ (\y. infdist y A) ` B" using \y \ B\ by blast show "bdd_below ((\y. infdist y A) ` B)" by (meson bdd_belowI2 infdist_nonneg) qed next show "infdist y A \ (INF y\B. infdist y A)" by (simp add: \B \ {}\ cINF_greatest min) qed finally show "setdist A B = infdist y A" . qed (fact \y \ B\) qed end diff --git a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy @@ -1,5036 +1,5080 @@ (* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Huge cleanup by LCP *) theory Equivalence_Lebesgue_Henstock_Integration imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral Homeomorphism Cartesian_Euclidean_Space begin lemma le_left_mono: "x \ y \ y \ a \ x \ (a::'a::preorder)" by (auto intro: order_trans) lemma ball_trans: assumes "y \ ball z q" "r + q \ s" shows "ball y r \ ball z s" using assms by metric lemma has_integral_implies_lebesgue_measurable_cbox: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) (cbox x y)" shows "f \ lebesgue_on (cbox x y) \\<^sub>M borel" proof (rule cld_measure.borel_measurable_cld) let ?L = "lebesgue_on (cbox x y)" let ?\ = "emeasure ?L" let ?\' = "outer_measure_of ?L" interpret L: finite_measure ?L proof show "?\ (space ?L) \ \" by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq) qed show "cld_measure ?L" proof fix B A assume "B \ A" "A \ null_sets ?L" then show "B \ sets ?L" using null_sets_completion_subset[OF \B \ A\, of lborel] by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: ) next fix A assume "A \ space ?L" "\B. B \ sets ?L \ ?\ B < \ \ A \ B \ sets ?L" from this(1) this(2)[of "space ?L"] show "A \ sets ?L" by (auto simp: Int_absorb2 less_top[symmetric]) qed auto then interpret cld_measure ?L . have content_eq_L: "A \ sets borel \ A \ cbox x y \ content A = measure ?L A" for A by (subst measure_restrict_space) (auto simp: measure_def) fix E and a b :: real assume "E \ sets ?L" "a < b" "0 < ?\ E" "?\ E < \" then obtain M :: real where "?\ E = M" "0 < M" by (cases "?\ E") auto define e where "e = M / (4 + 2 / (b - a))" from \a < b\ \0 have "0 < e" by (auto intro!: divide_pos_pos simp: field_simps e_def) have "e < M / (3 + 2 / (b - a))" using \a < b\ \0 < M\ unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps) then have "2 * e < (b - a) * (M - e * 3)" using \0 \0 < e\ \a < b\ by (simp add: field_simps) have e_less_M: "e < M / 1" unfolding e_def using \a < b\ \0 by (intro divide_strict_left_mono) (auto simp: field_simps) obtain d where "gauge d" and integral_f: "\p. p tagged_division_of cbox x y \ d fine p \ norm ((\(x,k) \ p. content k *\<^sub>R f x) - I) < e" using \0 f unfolding has_integral by auto define C where "C X m = X \ {x. ball x (1/Suc m) \ d x}" for X m have "incseq (C X)" for X unfolding C_def [abs_def] by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto { fix X assume "X \ space ?L" and eq: "?\' X = ?\ E" have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\m. C X m)" using \X \ space ?L\ by (intro SUP_outer_measure_of_incseq \incseq (C X)\) (auto simp: C_def) also have "(\m. C X m) = X" proof - { fix x obtain e where "0 < e" "ball x e \ d x" using gaugeD[OF \gauge d\, of x] unfolding open_contains_ball by auto moreover obtain n where "1 / (1 + real n) < e" using reals_Archimedean[OF \0] by (auto simp: inverse_eq_divide) then have "ball x (1 / (1 + real n)) \ ball x e" by (intro subset_ball) auto ultimately have "\n. ball x (1 / (1 + real n)) \ d x" by blast } then show ?thesis by (auto simp: C_def) qed finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\ E" using eq by auto also have "\ > M - e" using \0 < M\ \?\ E = M\ \0 by (auto intro!: ennreal_lessI) finally have "\m. M - e < outer_measure_of ?L (C X m)" unfolding less_SUP_iff by auto } note C = this let ?E = "{x\E. f x \ a}" and ?F = "{x\E. b \ f x}" have "\ (?\' ?E = ?\ E \ ?\' ?F = ?\ E)" proof assume eq: "?\' ?E = ?\ E \ ?\' ?F = ?\ E" with C[of ?E] C[of ?F] \E \ sets ?L\[THEN sets.sets_into_space] obtain ma mb where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)" by auto moreover define m where "m = max ma mb" ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)" using incseqD[OF \incseq (C ?E)\, of ma m, THEN outer_measure_of_mono] incseqD[OF \incseq (C ?F)\, of mb m, THEN outer_measure_of_mono] by (auto intro: less_le_trans) define d' where "d' x = d x \ ball x (1 / (3 * Suc m))" for x have "gauge d'" unfolding d'_def by (intro gauge_Int \gauge d\ gauge_ball) auto then obtain p where p: "p tagged_division_of cbox x y" "d' fine p" by (rule fine_division_exists) then have "d fine p" unfolding d'_def[abs_def] fine_def by auto define s where "s = {(x::'a, k). k \ (C ?E m) \ {} \ k \ (C ?F m) \ {}}" define T where "T E k = (SOME x. x \ k \ C E m)" for E k let ?A = "(\(x, k). (T ?E k, k)) ` (p \ s) \ (p - s)" let ?B = "(\(x, k). (T ?F k, k)) ` (p \ s) \ (p - s)" { fix X assume X_eq: "X = ?E \ X = ?F" let ?T = "(\(x, k). (T X k, k))" let ?p = "?T ` (p \ s) \ (p - s)" have in_s: "(x, k) \ s \ T X k \ k \ C X m" for x k using someI_ex[of "\x. x \ k \ C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def) { fix x k assume "(x, k) \ p" "(x, k) \ s" have k: "k \ ball x (1 / (3 * Suc m))" using \d' fine p\[THEN fineD, OF \(x, k) \ p\] by (auto simp: d'_def) then have "x \ ball (T X k) (1 / (3 * Suc m))" using in_s[OF \(x, k) \ s\] by (auto simp: C_def subset_eq dist_commute) then have "ball x (1 / (3 * Suc m)) \ ball (T X k) (1 / Suc m)" by (rule ball_trans) (auto simp: field_split_simps) with k in_s[OF \(x, k) \ s\] have "k \ d (T X k)" by (auto simp: C_def) } then have "d fine ?p" using \d fine p\ by (auto intro!: fineI) moreover have "?p tagged_division_of cbox x y" proof (rule tagged_division_ofI) show "finite ?p" using p(1) by auto next fix z k assume *: "(z, k) \ ?p" then consider "(z, k) \ p" "(z, k) \ s" | x' where "(x', k) \ p" "(x', k) \ s" "z = T X k" by (auto simp: T_def) then have "z \ k \ k \ cbox x y \ (\a b. k = cbox a b)" using p(1) by cases (auto dest: in_s) then show "z \ k" "k \ cbox x y" "\a b. k = cbox a b" by auto next fix z k z' k' assume "(z, k) \ ?p" "(z', k') \ ?p" "(z, k) \ (z', k')" with tagged_division_ofD(5)[OF p(1), of _ k _ k'] show "interior k \ interior k' = {}" by (auto simp: T_def dest: in_s) next have "{k. \x. (x, k) \ ?p} = {k. \x. (x, k) \ p}" by (auto simp: T_def image_iff Bex_def) then show "\{k. \x. (x, k) \ ?p} = cbox x y" using p(1) by auto qed ultimately have I: "norm ((\(x,k) \ ?p. content k *\<^sub>R f x) - I) < e" using integral_f by auto have "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) + (\(x,k) \ p - s. content k *\<^sub>R f x)" using p(1)[THEN tagged_division_ofD(1)] by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def) also have "(\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k))" proof (subst sum.reindex_nontrivial, safe) fix x1 x2 k assume 1: "(x1, k) \ p" "(x1, k) \ s" and 2: "(x2, k) \ p" "(x2, k) \ s" and eq: "content k *\<^sub>R f (T X k) \ 0" with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k] show "x1 = x2" by (auto simp: content_eq_0_interior) qed (use p in \auto intro!: sum.cong\) finally have eq: "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k)) + (\(x,k) \ p - s. content k *\<^sub>R f x)" . have in_T: "(x, k) \ s \ T X k \ X" for x k using in_s[of x k] by (auto simp: C_def) note I eq in_T } note parts = this have p_in_L: "(x, k) \ p \ k \ sets ?L" for x k using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space) have [simp]: "finite p" using tagged_division_ofD(1)[OF p(1)] . have "(M - 3*e) * (b - a) \ (\(x,k) \ p \ s. content k) * (b - a)" proof (intro mult_right_mono) have fin: "?\ (E \ \{k\snd`p. k \ C X m = {}}) < \" for X using \?\ E < \\ by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \E \ sets ?L\) have sets: "(E \ \{k\snd`p. k \ C X m = {}}) \ sets ?L" for X using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \E \ sets ?L\ sets.finite_Union sets.Int) (auto intro: p_in_L) { fix X assume "X \ E" "M - e < ?\' (C X m)" have "M - e \ ?\' (C X m)" by (rule less_imp_le) fact also have "\ \ ?\' (E - (E \ \{k\snd`p. k \ C X m = {}}))" proof (intro outer_measure_of_mono subsetI) fix v assume "v \ C X m" then have "v \ cbox x y" "v \ E" using \E \ space ?L\ \X \ E\ by (auto simp: space_restrict_space C_def) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto then show "v \ E - E \ (\{k\snd`p. k \ C X m = {}})" using \v \ C X m\ \v \ E\ by auto qed also have "\ = ?\ E - ?\ (E \ \{k\snd`p. k \ C X m = {}})" using \E \ sets ?L\ fin[of X] sets[of X] by (auto intro!: emeasure_Diff) finally have "?\ (E \ \{k\snd`p. k \ C X m = {}}) \ e" using \0 < e\ e_less_M apply (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") by (auto simp add: \?\ E = M\ ennreal_minus ennreal_le_iff2) note this } note upper_bound = this have "?\ (E \ \(snd`(p - s))) = ?\ ((E \ \{k\snd`p. k \ C ?E m = {}}) \ (E \ \{k\snd`p. k \ C ?F m = {}}))" by (intro arg_cong[where f="?\"]) (auto simp: s_def image_def Bex_def) also have "\ \ ?\ (E \ \{k\snd`p. k \ C ?E m = {}}) + ?\ (E \ \{k\snd`p. k \ C ?F m = {}})" using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto also have "\ \ e + ennreal e" using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto finally have "?\ E - 2*e \ ?\ (E - (E \ \(snd`(p - s))))" using \0 < e\ \E \ sets ?L\ tagged_division_ofD(1)[OF p(1)] by (subst emeasure_Diff) (auto simp: top_unique simp flip: ennreal_plus intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ \ ?\ (\x\p \ s. snd x)" proof (safe intro!: emeasure_mono subsetI) fix v assume "v \ E" and not: "v \ (\x\p \ s. snd x)" then have "v \ cbox x y" using \E \ space ?L\ by (auto simp: space_restrict_space) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto with not show "v \ \(snd ` (p - s))" by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"]) qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ = measure ?L (\x\p \ s. snd x)" by (auto intro!: emeasure_eq_ennreal_measure) finally have "M - 2 * e \ measure ?L (\x\p \ s. snd x)" unfolding \?\ E = M\ using \0 < e\ by (simp add: ennreal_minus) also have "measure ?L (\x\p \ s. snd x) = content (\x\p \ s. snd x)" using tagged_division_ofD(1,3,4) [OF p(1)] by (intro content_eq_L[symmetric]) (fastforce intro!: sets.finite_UN UN_least del: subsetI)+ also have "content (\x\p \ s. snd x) \ (\k\p \ s. content (snd k))" using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite dest!: p(1)[THEN tagged_division_ofD(4)]) finally show "M - 3 * e \ (\(x, y)\p \ s. content y)" using \0 < e\ by (simp add: split_beta) qed (use \a < b\ in auto) also have "\ = (\(x,k) \ p \ s. content k * (b - a))" by (simp add: sum_distrib_right split_beta') also have "\ \ (\(x,k) \ p \ s. content k * (f (T ?F k) - f (T ?E k)))" using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono) also have "\ = (\(x,k) \ p \ s. content k * f (T ?F k)) - (\(x,k) \ p \ s. content k * f (T ?E k))" by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric]) also have "\ = (\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x)" by (subst (1 2) parts) auto also have "\ \ norm ((\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x))" by auto also have "\ \ e + e" using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto finally show False using \2 * e < (b - a) * (M - e * 3)\ by (auto simp: field_simps) qed moreover have "?\' ?E \ ?\ E" "?\' ?F \ ?\ E" unfolding outer_measure_of_eq[OF \E \ sets ?L\, symmetric] by (auto intro!: outer_measure_of_mono) ultimately show "min (?\' ?E) (?\' ?F) < ?\ E" unfolding min_less_iff_disj by (auto simp: less_le) qed lemma has_integral_implies_lebesgue_measurable_real: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) \" shows "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof - define B :: "nat \ 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n show "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof (rule measurable_piecewise_restrict) have "(\n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \ \(B ` UNIV)" unfolding B_def by (intro UN_mono box_subset_cbox order_refl) then show "countable (range B)" "space lebesgue \ \(B ` UNIV)" by (auto simp: B_def UN_box_eq_UNIV) next fix \' assume "\' \ range B" then obtain n where \': "\' = B n" by auto then show "\' \ space lebesgue \ sets lebesgue" by (auto simp: B_def) have "f integrable_on \" using f by auto then have "(\x. f x * indicator \ x) integrable_on \" by (auto simp: integrable_on_def cong: has_integral_cong) then have "(\x. f x * indicator \ x) integrable_on (\ \ B n)" by (rule integrable_on_superset) auto then have "(\x. f x * indicator \ x) integrable_on B n" unfolding B_def by (rule integrable_on_subcbox) auto then show "(\x. f x * indicator \ x) \ lebesgue_on \' \\<^sub>M borel" unfolding B_def \' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def) qed qed lemma has_integral_implies_lebesgue_measurable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "(f has_integral I) \" shows "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI) fix i :: "'b" assume "i \ Basis" have "(\x. (f x \ i) * indicator \ x) \ borel_measurable (completion lborel)" using has_integral_linear[OF f bounded_linear_inner_left, of i] by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def) then show "(\x. indicator \ x *\<^sub>R f x \ i) \ borel_measurable (completion lborel)" by (simp add: ac_simps) qed subsection \Equivalence Lebesgue integral on \<^const>\lborel\ and HK-integral\ lemma has_integral_measure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and finite: "emeasure lborel A < \" shows "((\x. 1) has_integral measure lborel A) A" proof - { fix l u :: 'a have "((\x. 1) has_integral measure lborel (box l u)) (box l u)" proof cases assume "\b\Basis. l \ b \ u \ b" then show ?thesis apply simp apply (subst has_integral_restrict[symmetric, OF box_subset_cbox]) apply (subst has_integral_spike_interior_eq[where g="\_. 1"]) using has_integral_const[of "1::real" l u] apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases) done next assume "\ (\b\Basis. l \ b \ u \ b)" then have "box l u = {}" unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le) then show ?thesis by simp qed } note has_integral_box = this { fix a b :: 'a let ?M = "\A. measure lborel (A \ box a b)" have "Int_stable (range (\(a, b). box a b))" by (auto simp: Int_stable_def box_Int_box) moreover have "(range (\(a, b). box a b)) \ Pow UNIV" by auto moreover have "A \ sigma_sets UNIV (range (\(a, b). box a b))" using A unfolding borel_eq_box by simp ultimately have "((\x. 1) has_integral ?M A) (A \ box a b)" proof (induction rule: sigma_sets_induct_disjoint) case (basic A) then show ?case by (auto simp: box_Int_box has_integral_box) next case empty then show ?case by simp next case (compl A) then have [measurable]: "A \ sets borel" by (simp add: borel_eq_box) have "((\x. 1) has_integral ?M (box a b)) (box a b)" by (simp add: has_integral_box) moreover have "((\x. if x \ A \ box a b then 1 else 0) has_integral ?M A) (box a b)" by (subst has_integral_restrict) (auto intro: compl) ultimately have "((\x. 1 - (if x \ A \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_diff) then have "((\x. (if x \ (UNIV - A) \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_cong[THEN iffD1, rotated 1]) auto then have "((\x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \ box a b)" by (subst (asm) has_integral_restrict) auto also have "?M (box a b) - ?M A = ?M (UNIV - A)" by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2) finally show ?case . next case (union F) then have [measurable]: "\i. F i \ sets borel" by (simp add: borel_eq_box subset_eq) have "((\x. if x \ \(F ` UNIV) \ box a b then 1 else 0) has_integral ?M (\i. F i)) (box a b)" proof (rule has_integral_monotone_convergence_increasing) let ?f = "\k x. \i F i \ box a b then 1 else 0 :: real" show "\k. (?f k has_integral (\ik x. ?f k x \ ?f (Suc k) x" by (intro sum_mono2) auto from union(1) have *: "\x i j. x \ F i \ x \ F j \ j = i" by (auto simp add: disjoint_family_on_def) show "\x. (\k. ?f k x) \ (if x \ \(F ` UNIV) \ box a b then 1 else 0)" apply (auto simp: * sum.If_cases Iio_Int_singleton) apply (rule_tac k="Suc xa" in LIMSEQ_offset) apply simp done have *: "emeasure lborel ((\x. F x) \ box a b) \ emeasure lborel (box a b)" by (intro emeasure_mono) auto with union(1) show "(\k. \i ?M (\i. F i)" unfolding sums_def[symmetric] UN_extend_simps by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique) qed then show ?case by (subst (asm) has_integral_restrict) auto qed } note * = this show ?thesis proof (rule has_integral_monotone_convergence_increasing) let ?B = "\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set" let ?f = "\n::nat. \x. if x \ A \ ?B n then 1 else 0 :: real" let ?M = "\n. measure lborel (A \ ?B n)" show "\n::nat. (?f n has_integral ?M n) A" using * by (subst has_integral_restrict) simp_all show "\k x. ?f k x \ ?f (Suc k) x" by (auto simp: box_def) { fix x assume "x \ A" moreover have "(\k. indicator (A \ ?B k) x :: real) \ indicator (\k::nat. A \ ?B k) x" by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def) ultimately show "(\k. if x \ A \ ?B k then 1 else 0::real) \ 1" by (simp add: indicator_def UN_box_eq_UNIV) } have "(\n. emeasure lborel (A \ ?B n)) \ emeasure lborel (\n::nat. A \ ?B n)" by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def) also have "(\n. emeasure lborel (A \ ?B n)) = (\n. measure lborel (A \ ?B n))" proof (intro ext emeasure_eq_ennreal_measure) fix n have "emeasure lborel (A \ ?B n) \ emeasure lborel (?B n)" by (intro emeasure_mono) auto then show "emeasure lborel (A \ ?B n) \ top" by (auto simp: top_unique) qed finally show "(\n. measure lborel (A \ ?B n)) \ measure lborel A" using emeasure_eq_ennreal_measure[of lborel A] finite by (simp add: UN_box_eq_UNIV less_top) qed qed lemma nn_integral_has_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" shows "(f has_integral r) UNIV" using f proof (induct f arbitrary: r rule: borel_measurable_induct_real) case (set A) then have "((\x. 1) has_integral measure lborel A) A" by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator) with set show ?case by (simp add: ennreal_indicator measure_def) (simp add: indicator_def) next case (mult g c) then have "ennreal c * (\\<^sup>+ x. g x \lborel) = ennreal r" by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult) with \0 \ r\ \0 \ c\ obtain r' where "(c = 0 \ r = 0) \ (0 \ r' \ (\\<^sup>+ x. ennreal (g x) \lborel) = ennreal r' \ r = c * r')" by (cases "\\<^sup>+ x. ennreal (g x) \lborel" rule: ennreal_cases) (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric]) with mult show ?case by (auto intro!: has_integral_cmult_real) next case (add g h) then have "(\\<^sup>+ x. h x + g x \lborel) = (\\<^sup>+ x. h x \lborel) + (\\<^sup>+ x. g x \lborel)" by (simp add: nn_integral_add) with add obtain a b where "0 \ a" "0 \ b" "(\\<^sup>+ x. h x \lborel) = ennreal a" "(\\<^sup>+ x. g x \lborel) = ennreal b" "r = a + b" by (cases "\\<^sup>+ x. h x \lborel" "\\<^sup>+ x. g x \lborel" rule: ennreal2_cases) (auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus) with add show ?case by (auto intro!: has_integral_add) next case (seq U) note seq(1)[measurable] and f[measurable] { fix i x have "U i x \ f x" using seq(5) apply (rule LIMSEQ_le_const) using seq(4) apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def) done } note U_le_f = this { fix i have "(\\<^sup>+x. U i x \lborel) \ (\\<^sup>+x. f x \lborel)" using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp then obtain p where "(\\<^sup>+x. U i x \lborel) = ennreal p" "p \ r" "0 \ p" using seq(6) \0\r\ by (cases "\\<^sup>+x. U i x \lborel" rule: ennreal_cases) (auto simp: top_unique) moreover note seq ultimately have "\p. (\\<^sup>+x. U i x \lborel) = ennreal p \ 0 \ p \ p \ r \ (U i has_integral p) UNIV" by auto } then obtain p where p: "\i. (\\<^sup>+x. ennreal (U i x) \lborel) = ennreal (p i)" and bnd: "\i. p i \ r" "\i. 0 \ p i" and U_int: "\i.(U i has_integral (p i)) UNIV" by metis have int_eq: "\i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) \ integral UNIV f" proof (rule monotone_convergence_increasing) show "\k. U k integrable_on UNIV" using U_int by auto show "\k x. x\UNIV \ U k x \ U (Suc k) x" using \incseq U\ by (auto simp: incseq_def le_fun_def) then show "bounded (range (\k. integral UNIV (U k)))" using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) show "\x. x\UNIV \ (\k. U k x) \ f x" using seq by auto qed moreover have "(\i. (\\<^sup>+x. U i x \lborel)) \ (\\<^sup>+x. f x \lborel)" using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto ultimately have "integral UNIV f = r" by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique) with * show ?case by (simp add: has_integral_integral) qed lemma nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "(\\<^sup>+x. f x \lborel) = integral UNIV f" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique) qed lemma nn_integral_integrable_on: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "f integrable_on UNIV" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f) qed lemma nn_integral_has_integral_lborel: fixes f :: "'a::euclidean_space \ real" assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" shows "integral\<^sup>N lborel f = I" proof - from f_borel have "(\x. ennreal (f x)) \ borel_measurable lborel" by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain F where F: "\i. simple_function lborel (F i)" "incseq F" "\i x. F i x < top" "\x. (SUP i. F i x) = ennreal (f x)" by blast then have [measurable]: "\i. F i \ borel_measurable lborel" by (metis borel_measurable_simple_function) let ?B = "\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set" have "0 \ I" using I by (rule has_integral_nonneg) (simp add: nonneg) have F_le_f: "enn2real (F i x) \ f x" for i x using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\i. F i x"] by (cases "F i x" rule: ennreal_cases) auto let ?F = "\i x. F i x * indicator (?B i) x" have "(\\<^sup>+ x. ennreal (f x) \lborel) = (SUP i. integral\<^sup>N lborel (\x. ?F i x))" proof (subst nn_integral_monotone_convergence_SUP[symmetric]) { fix x obtain j where j: "x \ ?B j" using UN_box_eq_UNIV by auto have "ennreal (f x) = (SUP i. F i x)" using F(4)[of x] nonneg[of x] by (simp add: max_def) also have "\ = (SUP i. ?F i x)" proof (rule SUP_eq) fix i show "\j\UNIV. F i x \ ?F j x" using j F(2) by (intro bexI[of _ "max i j"]) (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def) qed (auto intro!: F split: split_indicator) finally have "ennreal (f x) = (SUP i. ?F i x)" . } then show "(\\<^sup>+ x. ennreal (f x) \lborel) = (\\<^sup>+ x. (SUP i. ?F i x) \lborel)" by simp qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator) also have "\ \ ennreal I" proof (rule SUP_least) fix i :: nat have finite_F: "(\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel) < \" proof (rule nn_integral_bound_simple_function) have "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} \ emeasure lborel (?B i)" by (intro emeasure_mono) (auto split: split_indicator) then show "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} < \" by (auto simp: less_top[symmetric] top_unique) qed (auto split: split_indicator intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal) have int_F: "(\x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV" using F(4) finite_F by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg) have "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) = (\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel)" using F(3,4) by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator) also have "\ = ennreal (integral UNIV (\x. enn2real (F i x) * indicator (?B i) x))" using F by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) (auto split: split_indicator intro: enn2real_nonneg) also have "\ \ ennreal I" by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f simp: \0 \ I\ split: split_indicator ) finally show "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) \ ennreal I" . qed finally have "(\\<^sup>+ x. ennreal (f x) \lborel) < \" by (auto simp: less_top[symmetric] top_unique) from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis by (simp add: integral_unique) qed lemma has_integral_iff_emeasure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and [simp]: "0 \ r" shows "((\x. 1) has_integral r) A \ emeasure lborel A = ennreal r" proof (cases "emeasure lborel A = \") case emeasure_A: True have "\ (\x. 1::real) integrable_on A" proof assume int: "(\x. 1::real) integrable_on A" then have "(indicator A::'a \ real) integrable_on UNIV" unfolding indicator_def[abs_def] integrable_restrict_UNIV . then obtain r where "((indicator A::'a\real) has_integral r) UNIV" by auto from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False by (simp add: ennreal_indicator) qed with emeasure_A show ?thesis by auto next case False then have "((\x. 1) has_integral measure lborel A) A" by (simp add: has_integral_measure_lborel less_top) with False show ?thesis by (auto simp: emeasure_eq_ennreal_measure has_integral_unique) qed lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x" by (auto simp: max_def ennreal_neg) lemma has_integral_integral_real: fixes f::"'a::euclidean_space \ real" assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - from integrableE[OF f] obtain r q where "0 \ r" "0 \ q" and r: "(\\<^sup>+ x. ennreal (max 0 (f x)) \lborel) = ennreal r" and q: "(\\<^sup>+ x. ennreal (max 0 (- f x)) \lborel) = ennreal q" and f: "f \ borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q" unfolding ennreal_max_0 by auto then have "((\x. max 0 (f x)) has_integral r) UNIV" "((\x. max 0 (- f x)) has_integral q) UNIV" using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto note has_integral_diff[OF this] moreover have "(\x. max 0 (f x) - max 0 (- f x)) = f" by auto ultimately show ?thesis by (simp add: eq) qed lemma has_integral_AE: assumes ae: "AE x in lborel. x \ \ \ f x = g x" shows "(f has_integral x) \ = (g has_integral x) \" proof - from ae obtain N where N: "N \ sets borel" "emeasure lborel N = 0" "{x. \ (x \ \ \ f x = g x)} \ N" by (auto elim!: AE_E) then have not_N: "AE x in lborel. x \ N" by (simp add: AE_iff_measurable) show ?thesis proof (rule has_integral_spike_eq[symmetric]) show "\x. x\\ - N \ f x = g x" using N(3) by auto show "negligible N" unfolding negligible_def proof (intro allI) fix a b :: "'a" let ?F = "\x::'a. if x \ cbox a b then indicator N x else 0 :: real" have "integrable lborel ?F = integrable lborel (\x::'a. 0::real)" using not_N N(1) by (intro integrable_cong_AE) auto moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)" using not_N N(1) by (intro integral_cong_AE) auto ultimately have "(?F has_integral 0) UNIV" using has_integral_integral_real[of ?F] by simp then show "(indicator N has_integral (0::real)) (cbox a b)" unfolding has_integral_restrict_UNIV . qed qed qed lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto finally show ?thesis using nonneg by auto qed lemma has_integral_iff_nn_integral_lebesgue: assumes f: "\x. 0 \ f x" shows "(f has_integral r) UNIV \ (f \ lebesgue \\<^sub>M borel \ integral\<^sup>N lebesgue f = r \ 0 \ r)" (is "?I = ?N") proof assume ?I have "0 \ r" using has_integral_nonneg[OF \?I\] f by auto then show ?N using nn_integral_has_integral_lebesgue[OF f \?I\] has_integral_implies_lebesgue_measurable[OF \?I\] by (auto simp: nn_integral_completion) next assume ?N then obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" by (auto dest: completion_ex_borel_measurable_real) moreover have "(\\<^sup>+ x. ennreal \f' x\ \lborel) = (\\<^sup>+ x. ennreal \f x\ \lborel)" using f' by (intro nn_integral_cong_AE) auto moreover have "((\x. \f' x\) has_integral r) UNIV \ ((\x. \f x\) has_integral r) UNIV" using f' by (intro has_integral_AE) auto moreover note nn_integral_has_integral[of "\x. \f' x\" r] \?N\ ultimately show ?I using f by (auto simp: nn_integral_completion) qed context fixes f::"'a::euclidean_space \ 'b::euclidean_space" begin lemma has_integral_integral_lborel: assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lborel f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma integrable_on_lborel: "integrable lborel f \ f integrable_on UNIV" using has_integral_integral_lborel by auto lemma integral_lborel: "integrable lborel f \ integral UNIV f = (\x. f x \lborel)" using has_integral_integral_lborel by auto end context begin private lemma has_integral_integral_lebesgue_real: fixes f :: "'a::euclidean_space \ real" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto moreover have "(\\<^sup>+ x. ennreal (norm (f x)) \lborel) = (\\<^sup>+ x. ennreal (norm (f' x)) \lborel)" using f' by (intro nn_integral_cong_AE) auto ultimately have "integrable lborel f'" using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE) note has_integral_integral_real[OF this] moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'" using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion) moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'" using f' by (simp add: integral_completion) moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \ (f has_integral integral\<^sup>L lborel f') UNIV" using f' by (intro has_integral_AE) auto ultimately show ?thesis by auto qed lemma has_integral_integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lebesgue f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma has_integral_integral_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" proof - let ?f = "\x. if x \ S then f x else 0" have "integrable lebesgue (\x. indicat_real S x *\<^sub>R f x)" using indicator_scaleR_eq_if [of S _ f] assms by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) then have "integrable lebesgue ?f" using indicator_scaleR_eq_if [of S _ f] assms by auto then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" by (rule has_integral_integral_lebesgue) then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" using has_integral_restrict_UNIV by blast moreover have "S \ space lebesgue \ sets lebesgue" by (simp add: assms) then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" by (simp add: integral_restrict_space indicator_scaleR_eq_if) ultimately show ?thesis by auto qed lemma lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = integral S f" by (metis has_integral_integral_lebesgue_on assms integral_unique) lemma integrable_on_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ f integrable_on UNIV" using has_integral_integral_lebesgue by auto lemma integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ integral UNIV f = (\x. f x \lebesgue)" using has_integral_integral_lebesgue by auto end subsection \Absolute integrability (this is the same as Lebesgue integrability)\ translations "LBINT x. f" == "CONST lebesgue_integral CONST lborel (\x. f)" translations "LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\x. f)" lemma set_integral_reflect: fixes S and f :: "real \ 'a :: {banach, second_countable_topology}" shows "(LBINT x : S. f x) = (LBINT x : {x. - x \ S}. f (- x))" unfolding set_lebesgue_integral_def by (subst lborel_integral_real_affine[where c="-1" and t=0]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma borel_integrable_atLeastAtMost': fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes f: "continuous_on {a..b} f" shows "set_integrable lborel {a..b} f" unfolding set_integrable_def by (intro borel_integrable_compact compact_Icc f) lemma integral_FTC_atLeastAtMost: fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" and F: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" and f: "continuous_on {a .. b} f" shows "integral\<^sup>L lborel (\x. indicator {a .. b} x *\<^sub>R f x) = F b - F a" proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have "(?f has_integral (\x. ?f x \lborel)) UNIV" using borel_integrable_atLeastAtMost'[OF f] unfolding set_integrable_def by (rule has_integral_integral_lborel) moreover have "(f has_integral F b - F a) {a .. b}" by (intro fundamental_theorem_of_calculus ballI assms) auto then have "(?f has_integral F b - F a) {a .. b}" by (subst has_integral_cong[where g=f]) auto then have "(?f has_integral F b - F a) UNIV" by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto ultimately show "integral\<^sup>L lborel ?f = F b - F a" by (rule has_integral_unique) qed lemma set_borel_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "set_integrable lborel S f" shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f" proof - let ?f = "\x. indicator S x *\<^sub>R f x" have "(?f has_integral LINT x : S | lborel. f x) UNIV" using assms has_integral_integral_lborel unfolding set_integrable_def set_lebesgue_integral_def by blast hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" apply (subst has_integral_restrict_UNIV [symmetric]) apply (rule has_integral_eq) by auto thus "f integrable_on S" by (auto simp add: integrable_on_def) with 1 have "(f has_integral (integral S f)) S" by (intro integrable_integral, auto simp add: integrable_on_def) thus "LINT x : S | lborel. f x = integral S f" by (intro has_integral_unique [OF 1]) qed lemma has_integral_set_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "(f has_integral (LINT x:S|lebesgue. f x)) S" using has_integral_integral_lebesgue f by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def if_distrib[of "\x. x *\<^sub>R f _"] cong: if_cong) lemma set_lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f" using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def) lemma lmeasurable_iff_has_integral: "S \ lmeasurable \ ((indicator S) has_integral measure lebesgue S) UNIV" by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI) abbreviation absolutely_integrable_on :: "('a::euclidean_space \ 'b::{banach, second_countable_topology}) \ 'a set \ bool" (infixr "absolutely'_integrable'_on" 46) where "f absolutely_integrable_on s \ set_integrable lebesgue s f" lemma absolutely_integrable_zero [simp]: "(\x. 0) absolutely_integrable_on S" by (simp add: set_integrable_def) lemma absolutely_integrable_on_def: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ f integrable_on S \ (\x. norm (f x)) integrable_on S" proof safe assume f: "f absolutely_integrable_on S" then have nf: "integrable lebesgue (\x. norm (indicator S x *\<^sub>R f x))" using integrable_norm set_integrable_def by blast show "f integrable_on S" by (rule set_lebesgue_integral_eq_integral[OF f]) have "(\x. norm (indicator S x *\<^sub>R f x)) = (\x. if x \ S then norm (f x) else 0)" by auto with integrable_on_lebesgue[OF nf] show "(\x. norm (f x)) integrable_on S" by (simp add: integrable_restrict_UNIV) next assume f: "f integrable_on S" and nf: "(\x. norm (f x)) integrable_on S" show "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule integrableI_bounded) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lebesgue" using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def) show "(\\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \lebesgue) < \" using nf nn_integral_has_integral_lebesgue[of "\x. norm (f x)" _ S] by (auto simp: integrable_on_def nn_integral_completion) qed qed lemma integrable_on_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable (lebesgue_on S) f" and S: "S \ sets lebesgue" shows "f integrable_on S" proof - have "integrable lebesgue (\x. indicator S x *\<^sub>R f x)" using S f inf_top.comm_neutral integrable_restrict_space by blast then show ?thesis using absolutely_integrable_on_def set_integrable_def by blast qed lemma absolutely_integrable_imp_integrable: assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) lemma absolutely_integrable_on_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "content (cbox a b) = 0 \ f absolutely_integrable_on (cbox a b)" by (auto simp: absolutely_integrable_on_def) lemma absolutely_integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f absolutely_integrable_on box a b \ f absolutely_integrable_on cbox a b" by (auto simp: integrable_on_open_interval absolutely_integrable_on_def) lemma absolutely_integrable_restrict_UNIV: "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ f absolutely_integrable_on S" unfolding set_integrable_def by (intro arg_cong2[where f=integrable]) auto lemma absolutely_integrable_onI: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f integrable_on S \ (\x. norm (f x)) integrable_on S \ f absolutely_integrable_on S" unfolding absolutely_integrable_on_def by auto lemma nonnegative_absolutely_integrable_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "\x. x \ A \ 0 \ f x" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_onI [OF f]) (use assms in \simp add: integrable_eq\) lemma absolutely_integrable_on_iff_nonneg: fixes f :: "'a :: euclidean_space \ real" assumes "\x. x \ S \ 0 \ f x" shows "f absolutely_integrable_on S \ f integrable_on S" proof - { assume "f integrable_on S" then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" using \f integrable_on S\ absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast then have "f absolutely_integrable_on S" using absolutely_integrable_restrict_UNIV by blast } then show ?thesis unfolding absolutely_integrable_on_def by auto qed lemma absolutely_integrable_on_scaleR_iff: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S \ c = 0 \ f absolutely_integrable_on S" proof (cases "c=0") case False then show ?thesis unfolding absolutely_integrable_on_def by (simp add: norm_mult) qed auto lemma absolutely_integrable_spike: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on T" and S: "negligible S" "\x. x \ T - S \ g x = f x" shows "g absolutely_integrable_on T" using assms integrable_spike unfolding absolutely_integrable_on_def by metis lemma absolutely_integrable_negligible: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" shows "f absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def integrable_negligible) lemma absolutely_integrable_spike_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" "\x. x \ T - S \ g x = f x" shows "(f absolutely_integrable_on T \ g absolutely_integrable_on T)" using assms by (blast intro: absolutely_integrable_spike sym) lemma absolutely_integrable_spike_set_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f absolutely_integrable_on S \ f absolutely_integrable_on T)" proof - have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ (\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" proof (rule absolutely_integrable_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto with absolutely_integrable_restrict_UNIV show ?thesis by blast qed lemma absolutely_integrable_spike_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f absolutely_integrable_on T" using absolutely_integrable_spike_set_eq f neg by blast lemma absolutely_integrable_reflect[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \ f absolutely_integrable_on cbox a b" apply (auto simp: absolutely_integrable_on_def dest: integrable_reflect [THEN iffD1]) unfolding integrable_on_def by auto lemma absolutely_integrable_reflect_real[simp]: fixes f :: "real \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on {-b .. -a} \ f absolutely_integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) lemma absolutely_integrable_on_subcbox: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; cbox a b \ S\ \ f absolutely_integrable_on cbox a b" by (meson absolutely_integrable_on_def integrable_on_subcbox) lemma absolutely_integrable_on_subinterval: fixes f :: "real \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; {a..b} \ S\ \ f absolutely_integrable_on {a..b}" using absolutely_integrable_on_subcbox by fastforce lemma integrable_subinterval: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" and "{c..d} \ {a..b}" shows "integrable (lebesgue_on {c..d}) f" proof (rule absolutely_integrable_imp_integrable) show "f absolutely_integrable_on {c..d}" proof - have "f integrable_on {c..d}" using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce moreover have "(\x. norm (f x)) integrable_on {c..d}" proof (rule integrable_on_subinterval) show "(\x. norm (f x)) integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) qed (use assms in auto) ultimately show ?thesis by (auto simp: absolutely_integrable_on_def) qed qed auto lemma indefinite_integral_continuous_real: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" shows "continuous_on {a..b} (\x. integral\<^sup>L (lebesgue_on {a..x}) f)" proof - have "f integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) then have "continuous_on {a..b} (\x. integral {a..x} f)" using indefinite_integral_continuous_1 by blast moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \ x" "x \ b" for x proof - have "{a..x} \ {a..b}" using that by auto then have "integrable (lebesgue_on {a..x}) f" using integrable_subinterval assms by blast then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" by (simp add: lebesgue_integral_eq_integral) qed ultimately show ?thesis by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) qed lemma lmeasurable_iff_integrable_on: "S \ lmeasurable \ (\x. 1::real) integrable_on S" by (subst absolutely_integrable_on_iff_nonneg[symmetric]) (simp_all add: lmeasurable_iff_integrable set_integrable_def) lemma lmeasure_integral_UNIV: "S \ lmeasurable \ measure lebesgue S = integral UNIV (indicator S)" by (simp add: lmeasurable_iff_has_integral integral_unique) lemma lmeasure_integral: "S \ lmeasurable \ measure lebesgue S = integral S (\x. 1::real)" by (fastforce simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) lemma integrable_on_const: "S \ lmeasurable \ (\x. c) integrable_on S" unfolding lmeasurable_iff_integrable by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one) lemma integral_indicator: assumes "(S \ T) \ lmeasurable" shows "integral T (indicator S) = measure lebesgue (S \ T)" proof - have "integral UNIV (indicator (S \ T)) = integral UNIV (\a. if a \ S \ T then 1::real else 0)" by (meson indicator_def) moreover have "(indicator (S \ T) has_integral measure lebesgue (S \ T)) UNIV" using assms by (simp add: lmeasurable_iff_has_integral) ultimately have "integral UNIV (\x. if x \ S \ T then 1 else 0) = measure lebesgue (S \ T)" by (metis (no_types) integral_unique) then show ?thesis using integral_restrict_Int [of UNIV "S \ T" "\x. 1::real"] apply (simp add: integral_restrict_Int [symmetric]) by (meson indicator_def) qed lemma measurable_integrable: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ (indicat_real S) integrable_on UNIV" by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def) lemma integrable_on_indicator: fixes S :: "'a::euclidean_space set" shows "indicat_real S integrable_on T \ (S \ T) \ lmeasurable" unfolding measurable_integrable unfolding integrable_restrict_UNIV [of T, symmetric] by (fastforce simp add: indicator_def elim: integrable_eq) lemma assumes \: "\ division_of S" shows lmeasurable_division: "S \ lmeasurable" (is ?l) and content_division: "(\k\\. measure lebesgue k) = measure lebesgue S" (is ?m) proof - { fix d1 d2 assume *: "d1 \ \" "d2 \ \" "d1 \ d2" then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" using division_ofD(4)[OF \] by blast with division_ofD(5)[OF \ *] have "d1 \ sets lborel" "d2 \ sets lborel" "d1 \ d2 \ (cbox a b - box a b) \ (cbox c d - box c d)" by auto moreover have "(cbox a b - box a b) \ (cbox c d - box c d) \ null_sets lborel" by (intro null_sets.Un null_sets_cbox_Diff_box) ultimately have "d1 \ d2 \ null_sets lborel" by (blast intro: null_sets_subset) } then show ?l ?m unfolding division_ofD(6)[OF \, symmetric] using division_ofD(1,4)[OF \] by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) qed lemma has_measure_limit: assumes "S \ lmeasurable" "e > 0" obtains B where "B > 0" "\a b. ball 0 B \ cbox a b \ \measure lebesgue (S \ cbox a b) - measure lebesgue S\ < e" using assms unfolding lmeasurable_iff_has_integral has_integral_alt' by (force simp: integral_indicator integrable_on_indicator) lemma lmeasurable_iff_indicator_has_integral: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ m = measure lebesgue S \ (indicat_real S has_integral m) UNIV" by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable) lemma has_measure_limit_iff: fixes f :: "'n::euclidean_space \ 'a::banach" shows "S \ lmeasurable \ m = measure lebesgue S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (S \ cbox a b) \ lmeasurable \ \measure lebesgue (S \ cbox a b) - m\ < e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox) next assume RHS [rule_format]: ?rhs show ?lhs apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) using RHS by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator) qed subsection\Applications to Negligibility\ lemma negligible_iff_null_sets: "negligible S \ S \ null_sets lebesgue" proof assume "negligible S" then have "(indicator S has_integral (0::real)) UNIV" by (auto simp: negligible) then show "S \ null_sets lebesgue" by (subst (asm) has_integral_iff_nn_integral_lebesgue) (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff) next assume S: "S \ null_sets lebesgue" show "negligible S" unfolding negligible_def proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2] has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1]) fix a b show "(\x. if x \ cbox a b then indicator S x else 0) \ lebesgue \\<^sub>M borel" using S by (auto intro!: measurable_If) then show "(\\<^sup>+ x. ennreal (if x \ cbox a b then indicator S x else 0) \lebesgue) = ennreal 0" using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) qed auto qed corollary eventually_ae_filter_negligible: "eventually P (ae_filter lebesgue) \ (\N. negligible N \ {x. \ P x} \ N)" by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) lemma starlike_negligible: assumes "closed S" and eq1: "\c x. (a + c *\<^sub>R x) \ S \ 0 \ c \ a + x \ S \ c = 1" shows "negligible S" proof - have "negligible ((+) (-a) ` S)" proof (subst negligible_on_intervals, intro allI) fix u v show "negligible ((+) (- a) ` S \ cbox u v)" using \closed S\ eq1 by (auto simp add: negligible_iff_null_sets algebra_simps intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp) (metis add_diff_eq diff_add_cancel scale_right_diff_distrib) qed then show ?thesis by (rule negligible_translation_rev) qed lemma starlike_negligible_strong: assumes "closed S" and star: "\c x. \0 \ c; c < 1; a+x \ S\ \ a + c *\<^sub>R x \ S" shows "negligible S" proof - show ?thesis proof (rule starlike_negligible [OF \closed S\, of a]) fix c x assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" with star have "\ (c < 1)" by auto moreover have "\ (c > 1)" using star [of "1/c" "c *\<^sub>R x"] cx by force ultimately show "c = 1" by arith qed qed lemma negligible_hyperplane: assumes "a \ 0 \ b \ 0" shows "negligible {x. a \ x = b}" proof - obtain x where x: "a \ x \ b" using assms apply auto apply (metis inner_eq_zero_iff inner_zero_right) using inner_zero_right by fastforce show ?thesis apply (rule starlike_negligible [OF closed_hyperplane, of x]) using x apply (auto simp: algebra_simps) done qed lemma negligible_lowdim: fixes S :: "'N :: euclidean_space set" assumes "dim S < DIM('N)" shows "negligible S" proof - obtain a where "a \ 0" and a: "span S \ {x. a \ x = 0}" using lowdim_subset_hyperplane [OF assms] by blast have "negligible (span S)" using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) then show ?thesis using span_base by (blast intro: negligible_subset) qed proposition negligible_convex_frontier: fixes S :: "'N :: euclidean_space set" assumes "convex S" shows "negligible(frontier S)" proof - have nf: "negligible(frontier S)" if "convex S" "0 \ S" for S :: "'N set" proof - obtain B where "B \ S" and indB: "independent B" and spanB: "S \ span B" and cardB: "card B = dim S" by (metis basis_exists) consider "dim S < DIM('N)" | "dim S = DIM('N)" using dim_subset_UNIV le_eq_less_or_eq by auto then show ?thesis proof cases case 1 show ?thesis by (rule negligible_subset [of "closure S"]) (simp_all add: frontier_def negligible_lowdim 1) next case 2 obtain a where a: "a \ interior S" apply (rule interior_simplex_nonempty [OF indB]) apply (simp add: indB independent_finite) apply (simp add: cardB 2) apply (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) done show ?thesis proof (rule starlike_negligible_strong [where a=a]) fix c::real and x have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" by (simp add: algebra_simps) assume "0 \ c" "c < 1" "a + x \ frontier S" then show "a + c *\<^sub>R x \ frontier S" apply (clarsimp simp: frontier_def) apply (subst eq) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"], auto) done qed auto qed qed show ?thesis proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain a where "a \ S" by auto show ?thesis using nf [of "(\x. -a + x) ` S"] by (metis \a \ S\ add.left_inverse assms convex_translation_eq frontier_translation image_eqI negligible_translation_rev) qed qed corollary negligible_sphere: "negligible (sphere a e)" using frontier_cball negligible_convex_frontier convex_cball by (blast intro: negligible_subset) lemma non_negligible_UNIV [simp]: "\ negligible UNIV" unfolding negligible_iff_null_sets by (auto simp: null_sets_def) lemma negligible_interval: "negligible (cbox a b) \ box a b = {}" "negligible (box a b) \ box a b = {}" by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq intro: eq_refl antisym less_imp_le) proposition open_not_negligible: assumes "open S" "S \ {}" shows "\ negligible S" proof assume neg: "negligible S" obtain a where "a \ S" using \S \ {}\ by blast then obtain e where "e > 0" "cball a e \ S" using \open S\ open_contains_cball_eq by blast let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One" have "cbox ?p ?q \ cball a e" proof (clarsimp simp: mem_box dist_norm) fix x assume "\i\Basis. ?p \ i \ x \ i \ x \ i \ ?q \ i" then have ax: "\(a - x) \ i\ \ e / real DIM('a)" if "i \ Basis" for i using that by (auto simp: algebra_simps) have "norm (a - x) \ (\i\Basis. \(a - x) \ i\)" by (rule norm_le_l1) also have "\ \ DIM('a) * (e / real DIM('a))" by (intro sum_bounded_above ax) also have "\ = e" by simp finally show "norm (a - x) \ e" . qed then have "negligible (cbox ?p ?q)" by (meson \cball a e \ S\ neg negligible_subset) with \e > 0\ show False by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff) qed lemma negligible_convex_interior: "convex S \ (negligible S \ interior S = {})" apply safe apply (metis interior_subset negligible_subset open_interior open_not_negligible) apply (metis Diff_empty closure_subset frontier_def negligible_convex_frontier negligible_subset) done lemma measure_eq_0_null_sets: "S \ null_sets M \ measure M S = 0" by (auto simp: measure_def null_sets_def) lemma negligible_imp_measure0: "negligible S \ measure lebesgue S = 0" by (simp add: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_emeasure0: "S \ sets lebesgue \ (negligible S \ emeasure lebesgue S = 0)" by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_measure0: "S \ lmeasurable \ (negligible S \ measure lebesgue S = 0)" apply (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) by (metis completion.null_sets_outer subsetI) lemma negligible_imp_sets: "negligible S \ S \ sets lebesgue" by (simp add: negligible_iff_null_sets null_setsD2) lemma negligible_imp_measurable: "negligible S \ S \ lmeasurable" by (simp add: fmeasurableI_null_sets negligible_iff_null_sets) lemma negligible_iff_measure: "negligible S \ S \ lmeasurable \ measure lebesgue S = 0" by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0) lemma negligible_outer: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T < e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis negligible_iff_measure order_refl) next assume ?rhs then show "negligible S" by (meson completion.null_sets_outer negligible_iff_null_sets) qed lemma negligible_outer_le: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T \ e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl) next assume ?rhs then show "negligible S" by (metis le_less_trans negligible_outer field_lbound_gt_zero) qed lemma negligible_UNIV: "negligible S \ (indicat_real S has_integral 0) UNIV" (is "_=?rhs") proof assume ?rhs then show "negligible S" apply (auto simp: negligible_def has_integral_iff integrable_on_indicator) by (metis negligible integral_unique lmeasure_integral_UNIV negligible_iff_measure0) qed (simp add: negligible) lemma sets_negligible_symdiff: "\S \ sets lebesgue; negligible((S - T) \ (T - S))\ \ T \ sets lebesgue" by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un) lemma lmeasurable_negligible_symdiff: "\S \ lmeasurable; negligible((S - T) \ (T - S))\ \ T \ lmeasurable" using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast lemma measure_Un3_negligible: assumes meas: "S \ lmeasurable" "T \ lmeasurable" "U \ lmeasurable" and neg: "negligible(S \ T)" "negligible(S \ U)" "negligible(T \ U)" and V: "S \ T \ U = V" shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ T) = 0" using neg(1) negligible_imp_measure0 by blast have [simp]: "measure lebesgue (S \ U \ T \ U) = 0" using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast have "measure lebesgue V = measure lebesgue (S \ T \ U)" using V by simp also have "\ = measure lebesgue S + measure lebesgue T + measure lebesgue U" by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2) finally show ?thesis . qed lemma measure_translate_add: assumes meas: "S \ lmeasurable" "T \ lmeasurable" and U: "S \ ((+)a ` T) = U" and neg: "negligible(S \ ((+)a ` T))" shows "measure lebesgue S + measure lebesgue T = measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ (+) a ` T) = 0" using neg negligible_imp_measure0 by blast have "measure lebesgue (S \ ((+)a ` T)) = measure lebesgue S + measure lebesgue T" by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un) then show ?thesis using U by auto qed lemma measure_negligible_symdiff: assumes S: "S \ lmeasurable" and neg: "negligible (S - T \ (T - S))" shows "measure lebesgue T = measure lebesgue S" proof - have "measure lebesgue (S - T) = 0" using neg negligible_Un_eq negligible_imp_measure0 by blast then show ?thesis by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0) qed lemma measure_closure: assumes "bounded S" and neg: "negligible (frontier S)" shows "measure lebesgue (closure S) = measure lebesgue S" proof - have "measure lebesgue (frontier S) = 0" by (metis neg negligible_imp_measure0) then show ?thesis by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier) qed lemma measure_interior: "\bounded S; negligible(frontier S)\ \ measure lebesgue (interior S) = measure lebesgue S" using measure_closure measure_frontier negligible_imp_measure0 by fastforce lemma measurable_Jordan: assumes "bounded S" and neg: "negligible (frontier S)" shows "S \ lmeasurable" proof - have "closure S \ lmeasurable" by (metis lmeasurable_closure \bounded S\) moreover have "interior S \ lmeasurable" by (simp add: lmeasurable_interior \bounded S\) moreover have "interior S \ S" by (simp add: interior_subset) ultimately show ?thesis using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) qed lemma measurable_convex: "\convex S; bounded S\ \ S \ lmeasurable" by (simp add: measurable_Jordan negligible_convex_frontier) +lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)" +proof - + have "ball c r - cball c r \ (cball c r - ball c r) = sphere c r" + by auto + hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)" + using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all + thus ?thesis by simp +qed + subsection\Negligibility of image under non-injective linear map\ lemma negligible_Union_nat: assumes "\n::nat. negligible(S n)" shows "negligible(\n. S n)" proof - have "negligible (\m\k. S m)" for k using assms by blast then have 0: "integral UNIV (indicat_real (\m\k. S m)) = 0" and 1: "(indicat_real (\m\k. S m)) integrable_on UNIV" for k by (auto simp: negligible has_integral_iff) have 2: "\k x. indicat_real (\m\k. S m) x \ (indicat_real (\m\Suc k. S m) x)" by (simp add: indicator_def) have 3: "\x. (\k. indicat_real (\m\k. S m) x) \ (indicat_real (\n. S n) x)" by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually) have 4: "bounded (range (\k. integral UNIV (indicat_real (\m\k. S m))))" by (simp add: 0) have *: "indicat_real (\n. S n) integrable_on UNIV \ (\k. integral UNIV (indicat_real (\m\k. S m))) \ (integral UNIV (indicat_real (\n. S n)))" by (intro monotone_convergence_increasing 1 2 3 4) then have "integral UNIV (indicat_real (\n. S n)) = (0::real)" using LIMSEQ_unique by (auto simp: 0) then show ?thesis using * by (simp add: negligible_UNIV has_integral_iff) qed lemma negligible_linear_singular_image: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\ inj f" shows "negligible (f ` S)" proof - obtain a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms linear_singular_image_hyperplane by blast then show "negligible (f ` S)" by (metis negligible_hyperplane negligible_subset) qed lemma measure_negligible_finite_Union: assumes "finite \" and meas: "\S. S \ \ \ S \ lmeasurable" and djointish: "pairwise (\S T. negligible (S \ T)) \" shows "measure lebesgue (\\) = (\S\\. measure lebesgue S)" using assms proof (induction) case empty then show ?case by auto next case (insert S \) then have "S \ lmeasurable" "\\ \ lmeasurable" "pairwise (\S T. negligible (S \ T)) \" by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI) then show ?case proof (simp add: measure_Un3 insert) have *: "\T. T \ (\) S ` \ \ negligible T" using insert by (force simp: pairwise_def) have "negligible(S \ \\)" unfolding Int_Union by (rule negligible_Union) (simp_all add: * insert.hyps(1)) then show "measure lebesgue (S \ \\) = 0" using negligible_imp_measure0 by blast qed qed lemma measure_negligible_finite_Union_image: assumes "finite S" and meas: "\x. x \ S \ f x \ lmeasurable" and djointish: "pairwise (\x y. negligible (f x \ f y)) S" shows "measure lebesgue (\(f ` S)) = (\x\S. measure lebesgue (f x))" proof - have "measure lebesgue (\(f ` S)) = sum (measure lebesgue) (f ` S)" using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union) also have "\ = sum (measure lebesgue \ f) S" using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \finite S\]) also have "\ = (\x\S. measure lebesgue (f x))" by simp finally show ?thesis . qed subsection \Negligibility of a Lipschitz image of a negligible set\ text\The bound will be eliminated by a sort of onion argument\ lemma locally_Lipschitz_negl_bounded: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "0 < B" "bounded S" "negligible S" and lips: "\x. x \ S \ \T. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" unfolding negligible_iff_null_sets proof (clarsimp simp: completion.null_sets_outer) fix e::real assume "0 < e" have "S \ lmeasurable" using \negligible S\ by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) then have "S \ sets lebesgue" by blast have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)" using \0 < e\ \0 < B\ by (simp add: field_split_simps) obtain T where "open T" "S \ T" "(T - S) \ lmeasurable" "measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)" using sets_lebesgue_outer_open [OF \S \ sets lebesgue\ e22] by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le) then have T: "measure lebesgue T \ e/2 / (2 * B * DIM('M)) ^ DIM('N)" using \negligible S\ by (simp add: measure_Diff_null_set negligible_iff_null_sets) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. norm(y - x) < r \ y \ T \ (y \ S \ norm(f y - f x) \ B * norm(y - x))))" for x proof (cases "x \ S") case True obtain U where "open U" "x \ U" and U: "\y. y \ S \ U \ norm(f y - f x) \ B * norm(y - x)" using lips [OF \x \ S\] by auto have "x \ T \ U" using \S \ T\ \x \ U\ \x \ S\ by auto then obtain \ where "0 < \" "ball x \ \ T \ U" by (metis \open T\ \open U\ openE open_Int) then show ?thesis apply (rule_tac x="min (1/2) \" in exI) apply (simp del: divide_const_simps) apply (intro allI impI conjI) apply (metis dist_commute dist_norm mem_ball subsetCE) by (metis Int_iff subsetCE U dist_norm mem_ball norm_minus_commute) next case False then show ?thesis by (rule_tac x="1/4" in exI) auto qed then obtain R where R12: "\x. 0 < R x \ R x \ 1/2" and RT: "\x y. \x \ S; norm(y - x) < R x\ \ y \ T" and RB: "\x y. \x \ S; y \ S; norm(y - x) < R x\ \ norm(f y - f x) \ B * norm(y - x)" by metis+ then have gaugeR: "gauge (\x. ball x (R x))" by (simp add: gauge_def) obtain c where c: "S \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \ {}" proof - obtain B where B: "\x. x \ S \ norm x \ B" using \bounded S\ bounded_iff by blast show ?thesis apply (rule_tac c = "abs B + 1" in that) using norm_bound_Basis_le Basis_le_norm apply (fastforce simp: box_eq_empty mem_box dest!: B intro: order_trans)+ done qed obtain \ where "countable \" and Dsub: "\\ \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and pw: "pairwise (\A B. interior A \ interior B = {}) \" and Ksub: "\K. K \ \ \ \x \ S \ K. K \ (\x. ball x (R x)) x" and exN: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (2*c) / 2^n" and "S \ \\" using covering_lemma [OF c gaugeR] by force have "\u v z. K = cbox u v \ box u v \ {} \ z \ S \ z \ cbox u v \ cbox u v \ ball z (R z)" if "K \ \" for K proof - obtain u v where "K = cbox u v" using \K \ \\ cbox by blast with that show ?thesis apply (rule_tac x=u in exI) apply (rule_tac x=v in exI) apply (metis Int_iff interior_cbox cbox Ksub) done qed then obtain uf vf zf where uvz: "\K. K \ \ \ K = cbox (uf K) (vf K) \ box (uf K) (vf K) \ {} \ zf K \ S \ zf K \ cbox (uf K) (vf K) \ cbox (uf K) (vf K) \ ball (zf K) (R (zf K))" by metis define prj1 where "prj1 \ \x::'M. x \ (SOME i. i \ Basis)" define fbx where "fbx \ \D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N) (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)" have vu_pos: "0 < prj1 (vf X - uf X)" if "X \ \" for X using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left) have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \ i" if "X \ \" "i \ Basis" for X i proof - have "cbox (uf X) (vf X) \ \" using uvz \X \ \\ by auto with exN obtain n where "\i. i \ Basis \ vf X \ i - uf X \ i = (2*c) / 2^n" by blast then show ?thesis by (simp add: \i \ Basis\ SOME_Basis inner_diff prj1_def) qed have countbl: "countable (fbx ` \)" using \countable \\ by blast have "(\k\fbx`\'. measure lebesgue k) \ e/2" if "\' \ \" "finite \'" for \' proof - have BM_ge0: "0 \ B * (DIM('M) * prj1 (vf X - uf X))" if "X \ \'" for X using \0 < B\ \\' \ \\ that vu_pos by fastforce have "{} \ \'" using cbox \\' \ \\ interior_empty by blast have "(\k\fbx`\'. measure lebesgue k) \ sum (measure lebesgue o fbx) \'" by (rule sum_image_le [OF \finite \'\]) (force simp: fbx_def) also have "\ \ (\X\\'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)" proof (rule sum_mono) fix X assume "X \ \'" then have "X \ \" using \\' \ \\ by blast then have ufvf: "cbox (uf X) (vf X) = X" using uvz by blast have "prj1 (vf X - uf X) ^ DIM('M) = (\i::'M \ Basis. prj1 (vf X - uf X))" by (rule prod_constant [symmetric]) also have "\ = (\i\Basis. vf X \ i - uf X \ i)" apply (rule prod.cong [OF refl]) by (simp add: \X \ \\ inner_diff_left prj1_idem) finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\i\Basis. vf X \ i - uf X \ i)" . have "uf X \ cbox (uf X) (vf X)" "vf X \ cbox (uf X) (vf X)" using uvz [OF \X \ \\] by (force simp: mem_box)+ moreover have "cbox (uf X) (vf X) \ ball (zf X) (1/2)" by (meson R12 order_trans subset_ball uvz [OF \X \ \\]) ultimately have "uf X \ ball (zf X) (1/2)" "vf X \ ball (zf X) (1/2)" by auto then have "dist (vf X) (uf X) \ 1" unfolding mem_ball by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict) then have 1: "prj1 (vf X - uf X) \ 1" unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce have 0: "0 \ prj1 (vf X - uf X)" using \X \ \\ prj1_def vu_pos by fastforce have "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))" apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \X \ \'\) apply (simp add: power_mult_distrib \0 < B\ prj1_eq [symmetric]) using MleN 0 1 uvz \X \ \\ apply (fastforce simp add: box_ne_empty power_decreasing) done also have "\ = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" by (subst (3) ufvf[symmetric]) simp finally show "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" . qed also have "\ = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \'" by (simp add: sum_distrib_left) also have "\ \ e/2" proof - have div: "\' division_of \\'" apply (auto simp: \finite \'\ \{} \ \'\ division_of_def) using cbox that apply blast using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def apply force+ done have le_meaT: "measure lebesgue (\\') \ measure lebesgue T" proof (rule measure_mono_fmeasurable) show "(\\') \ sets lebesgue" using div lmeasurable_division by auto have "\\' \ \\" using \\' \ \\ by blast also have "... \ T" proof (clarify) fix x D assume "x \ D" "D \ \" show "x \ T" using Ksub [OF \D \ \\] by (metis \x \ D\ Int_iff dist_norm mem_ball norm_minus_commute subsetD RT) qed finally show "\\' \ T" . show "T \ lmeasurable" using \S \ lmeasurable\ \S \ T\ \T - S \ lmeasurable\ fmeasurable_Diff_D by blast qed have "sum (measure lebesgue) \' = sum content \'" using \\' \ \\ cbox by (force intro: sum.cong) then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \' = (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\\')" using content_division [OF div] by auto also have "\ \ (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T" apply (rule mult_left_mono [OF le_meaT]) using \0 < B\ apply (simp add: algebra_simps) done also have "\ \ e/2" using T \0 < B\ by (simp add: field_simps) finally show ?thesis . qed finally show ?thesis . qed then have e2: "sum (measure lebesgue) \ \ e/2" if "\ \ fbx ` \" "finite \" for \ by (metis finite_subset_image that) show "\W\lmeasurable. f ` S \ W \ measure lebesgue W < e" proof (intro bexI conjI) have "\X\\. f y \ fbx X" if "y \ S" for y proof - obtain X where "y \ X" "X \ \" using \S \ \\\ \y \ S\ by auto then have y: "y \ ball(zf X) (R(zf X))" using uvz by fastforce have conj_le_eq: "z - b \ y \ y \ z + b \ abs(y - z) \ b" for z y b::real by auto have yin: "y \ cbox (uf X) (vf X)" and zin: "(zf X) \ cbox (uf X) (vf X)" using uvz \X \ \\ \y \ X\ by auto have "norm (y - zf X) \ (\i\Basis. \(y - zf X) \ i\)" by (rule norm_le_l1) also have "\ \ real DIM('M) * prj1 (vf X - uf X)" proof (rule sum_bounded_above) fix j::'M assume j: "j \ Basis" show "\(y - zf X) \ j\ \ prj1 (vf X - uf X)" using yin zin j by (fastforce simp add: mem_box prj1_idem [OF \X \ \\ j] inner_diff_left) qed finally have nole: "norm (y - zf X) \ DIM('M) * prj1 (vf X - uf X)" by simp have fle: "\f y \ i - f(zf X) \ i\ \ B * DIM('M) * prj1 (vf X - uf X)" if "i \ Basis" for i proof - have "\f y \ i - f (zf X) \ i\ = \(f y - f (zf X)) \ i\" by (simp add: algebra_simps) also have "\ \ norm (f y - f (zf X))" by (simp add: Basis_le_norm that) also have "\ \ B * norm(y - zf X)" by (metis uvz RB \X \ \\ dist_commute dist_norm mem_ball \y \ S\ y) also have "\ \ B * real DIM('M) * prj1 (vf X - uf X)" using \0 < B\ by (simp add: nole) finally show ?thesis . qed show ?thesis by (rule_tac x=X in bexI) (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \X \ \\) qed then show "f ` S \ (\D\\. fbx D)" by auto next have 1: "\D. D \ \ \ fbx D \ lmeasurable" by (auto simp: fbx_def) have 2: "I' \ \ \ finite I' \ measure lebesgue (\D\I'. fbx D) \ e/2" for I' by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def) show "(\D\\. fbx D) \ lmeasurable" by (intro fmeasurable_UN_bound[OF \countable \\ 1 2]) have "measure lebesgue (\D\\. fbx D) \ e/2" by (intro measure_UN_bound[OF \countable \\ 1 2]) then show "measure lebesgue (\D\\. fbx D) < e" using \0 < e\ by linarith qed qed proposition negligible_locally_Lipschitz_image: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and lips: "\x. x \ S \ \T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" proof - let ?S = "\n. ({x \ S. norm x \ n \ (\T. open T \ x \ T \ (\y\S \ T. norm (f y - f x) \ (real n + 1) * norm (y - x)))})" have negfn: "f ` ?S n \ null_sets lebesgue" for n::nat unfolding negligible_iff_null_sets[symmetric] apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded) by (auto simp: MleN bounded_iff intro: negligible_subset [OF \negligible S\]) have "S = (\n. ?S n)" proof (intro set_eqI iffI) fix x assume "x \ S" with lips obtain T B where T: "open T" "x \ T" and B: "\y. y \ S \ T \ norm(f y - f x) \ B * norm(y - x)" by metis+ have no: "norm (f y - f x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" if "y \ S \ T" for y proof - have "B * norm(y - x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans) then show ?thesis using B order_trans that by blast qed have "x \ ?S (nat (ceiling (max B (norm x))))" apply (simp add: \x \ S \, rule) using real_nat_ceiling_ge max.bounded_iff apply blast using T no apply (force simp: algebra_simps) done then show "x \ (\n. ?S n)" by force qed auto then show ?thesis by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn) qed corollary negligible_differentiable_image_negligible: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "\T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" if "x \ S" for x proof - obtain f' where "linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x)" using diff_f \x \ S\ by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt) obtain B where "B > 0" and B: "\x. norm (f' x) \ B * norm x" using linear_bounded_pos \linear f'\ by blast obtain d where "d>0" and d: "\y. \y \ S; norm (y - x) < d\ \ norm (f y - f x - f' (y - x)) \ norm (y - x)" using f' [of 1] by (force simp:) have *: "norm (f y - f x) \ (B + 1) * norm (y - x)" if "y \ S" "norm (y - x) < d" for y proof - have "norm (f y - f x) -B * norm (y - x) \ norm (f y - f x) - norm (f' (y - x))" by (simp add: B) also have "\ \ norm (f y - f x - f' (y - x))" by (rule norm_triangle_ineq2) also have "... \ norm (y - x)" by (rule d [OF that]) finally show ?thesis by (simp add: algebra_simps) qed show ?thesis apply (rule_tac x="ball x d" in exI) apply (rule_tac x="B+1" in exI) using \d>0\ apply (auto simp: dist_norm norm_minus_commute intro!: *) done qed with negligible_locally_Lipschitz_image assms show ?thesis by metis qed corollary negligible_differentiable_image_lowdim: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "x \ DIM('M) \ x \ DIM('N)" for x using MlessN by linarith obtain lift :: "'M * real \ 'N" and drop :: "'N \ 'M * real" and j :: 'N where "linear lift" "linear drop" and dropl [simp]: "\z. drop (lift z) = z" and "j \ Basis" and j: "\x. lift(x,0) \ j = 0" using lowerdim_embeddings [OF MlessN] by metis have "negligible {x. x\j = 0}" by (metis \j \ Basis\ negligible_standard_hyperplane) then have neg0S: "negligible ((\x. lift (x, 0)) ` S)" apply (rule negligible_subset) by (simp add: image_subsetI j) have diff_f': "f \ fst \ drop differentiable_on (\x. lift (x, 0)) ` S" using diff_f apply (clarsimp simp add: differentiable_on_def) apply (intro differentiable_chain_within linear_imp_differentiable [OF \linear drop\] linear_imp_differentiable [OF fst_linear]) apply (force simp: image_comp o_def) done have "f = (f o fst o drop o (\x. lift (x, 0)))" by (simp add: o_def) then show ?thesis apply (rule ssubst) apply (subst image_comp [symmetric]) apply (metis negligible_differentiable_image_negligible order_refl diff_f' neg0S) done qed subsection\Measurability of countable unions and intersections of various kinds.\ lemma assumes S: "\n. S n \ lmeasurable" and leB: "\n. measure lebesgue (S n) \ B" and nest: "\n. S n \ S(Suc n)" shows measurable_nested_Union: "(\n. S n) \ lmeasurable" and measure_nested_Union: "(\n. measure lebesgue (S n)) \ measure lebesgue (\n. S n)" (is ?Lim) proof - have 1: "\n. (indicat_real (S n)) integrable_on UNIV" using S measurable_integrable by blast have 2: "\n x::'a. indicat_real (S n) x \ (indicat_real (S (Suc n)) x)" by (simp add: indicator_leI nest rev_subsetD) have 3: "\x. (\n. indicat_real (S n) x) \ (indicat_real (\(S ` UNIV)) x)" apply (rule tendsto_eventually) apply (simp add: indicator_def) by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) have 4: "bounded (range (\n. integral UNIV (indicat_real (S n))))" using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) have "(\n. S n) \ lmeasurable \ ?Lim" apply (simp add: lmeasure_integral_UNIV S cong: conj_cong) apply (simp add: measurable_integrable) apply (rule monotone_convergence_increasing [OF 1 2 3 4]) done then show "(\n. S n) \ lmeasurable" "?Lim" by auto qed lemma assumes S: "\n. S n \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and leB: "\n. (\k\n. measure lebesgue (S k)) \ B" shows measurable_countable_negligible_Union: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - have 1: "\ (S ` {..n}) \ lmeasurable" for n using S by blast have 2: "measure lebesgue (\ (S ` {..n})) \ B" for n proof - have "measure lebesgue (\ (S ` {..n})) \ (\k\n. measure lebesgue (S k))" by (simp add: S fmeasurableD measure_UNION_le) with leB show ?thesis using order_trans by blast qed have 3: "\n. \ (S ` {..n}) \ \ (S ` {..Suc n})" by (simp add: SUP_subset_mono) have eqS: "(\n. S n) = (\n. \ (S ` {..n}))" using atLeastAtMost_iff by blast also have "(\n. \ (S ` {..n})) \ lmeasurable" by (intro measurable_nested_Union [OF 1 2] 3) finally show "(\n. S n) \ lmeasurable" . have eqm: "(\i\n. measure lebesgue (S i)) = measure lebesgue (\ (S ` {..n}))" for n using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono) have "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. \ (S ` {..n}))" by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3) then show ?Sums by (simp add: eqS) qed lemma negligible_countable_Union [intro]: assumes "countable \" and meas: "\S. S \ \ \ negligible S" shows "negligible (\\)" proof (cases "\ = {}") case False then show ?thesis by (metis from_nat_into range_from_nat_into assms negligible_Union_nat) qed simp lemma assumes S: "\n. (S n) \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and bo: "bounded (\n. S n)" shows measurable_countable_negligible_Union_bounded: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union_bounded: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - obtain a b where ab: "(\n. S n) \ cbox a b" using bo bounded_subset_cbox_symmetric by metis then have B: "(\k\n. measure lebesgue (S k)) \ measure lebesgue (cbox a b)" for n proof - have "(\k\n. measure lebesgue (S k)) = measure lebesgue (\ (S ` {..n}))" using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish by (metis S finite_atMost subset_UNIV) also have "\ \ measure lebesgue (cbox a b)" apply (rule measure_mono_fmeasurable) using ab S by force+ finally show ?thesis . qed show "(\n. S n) \ lmeasurable" by (rule measurable_countable_negligible_Union [OF S djointish B]) show ?Sums by (rule measure_countable_negligible_Union [OF S djointish B]) qed lemma measure_countable_Union_approachable: assumes "countable \" "e > 0" and measD: "\d. d \ \ \ d \ lmeasurable" and B: "\D'. \D' \ \; finite D'\ \ measure lebesgue (\D') \ B" obtains D' where "D' \ \" "finite D'" "measure lebesgue (\\) - e < measure lebesgue (\D')" proof (cases "\ = {}") case True then show ?thesis by (simp add: \e > 0\ that) next case False let ?S = "\n. \k \ n. from_nat_into \ k" have "(\n. measure lebesgue (?S n)) \ measure lebesgue (\n. ?S n)" proof (rule measure_nested_Union) show "?S n \ lmeasurable" for n by (simp add: False fmeasurable.finite_UN from_nat_into measD) show "measure lebesgue (?S n) \ B" for n by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI) show "?S n \ ?S (Suc n)" for n by force qed then obtain N where N: "\n. n \ N \ dist (measure lebesgue (?S n)) (measure lebesgue (\n. ?S n)) < e" using metric_LIMSEQ_D \e > 0\ by blast show ?thesis proof show "from_nat_into \ ` {..N} \ \" by (auto simp: False from_nat_into) have eq: "(\n. \k\n. from_nat_into \ k) = (\\)" using \countable \\ False by (auto intro: from_nat_into dest: from_nat_into_surj [OF \countable \\]) show "measure lebesgue (\\) - e < measure lebesgue (\ (from_nat_into \ ` {..N}))" using N [OF order_refl] by (auto simp: eq algebra_simps dist_norm) qed auto qed subsection\Negligibility is a local property\ lemma locally_negligible_alt: "negligible S \ (\x \ S. \U. openin (top_of_set S) U \ x \ U \ negligible U)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs using openin_subtopology_self by blast next assume ?rhs then obtain U where ope: "\x. x \ S \ openin (top_of_set S) (U x)" and cov: "\x. x \ S \ x \ U x" and neg: "\x. x \ S \ negligible (U x)" by metis obtain \ where "\ \ U ` S" "countable \" and eq: "\\ = \(U ` S)" using ope by (force intro: Lindelof_openin [of "U ` S" S]) then have "negligible (\\)" by (metis imageE neg negligible_countable_Union subset_eq) with eq have "negligible (\(U ` S))" by metis moreover have "S \ \(U ` S)" using cov by blast ultimately show "negligible S" using negligible_subset by blast qed lemma locally_negligible: "locally negligible S \ negligible S" unfolding locally_def apply safe apply (meson negligible_subset openin_subtopology_self locally_negligible_alt) by (meson negligible_subset openin_imp_subset order_refl) subsection\Integral bounds\ lemma set_integral_norm_bound: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "set_integrable M k f \ norm (LINT x:k|M. f x) \ LINT x:k|M. norm (f x)" using integral_norm_bound[of M "\x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def) lemma set_integral_finite_UN_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" and ae: "\i j. i \ I \ j \ I \ AE x in M. (x \ A i \ x \ A j) \ i = j" and [measurable]: "\i. i \ I \ A i \ sets M" and f: "\i. i \ I \ set_integrable M (A i) f" shows "LINT x:(\i\I. A i)|M. f x = (\i\I. LINT x:A i|M. f x)" using \finite I\ order_refl[of I] proof (induction I rule: finite_subset_induct') case (insert i I') have "AE x in M. (\j\I'. x \ A i \ x \ A j)" proof (intro AE_ball_countable[THEN iffD2] ballI) fix j assume "j \ I'" with \I' \ I\ \i \ I'\ have "i \ j" "j \ I" by auto then show "AE x in M. x \ A i \ x \ A j" using ae[of i j] \i \ I\ by auto qed (use \finite I'\ in \rule countable_finite\) then have "AE x\A i in M. \xa\I'. x \ A xa " by auto with insert.hyps insert.IH[symmetric] show ?case by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) qed (simp add: set_lebesgue_integral_def) lemma set_integrable_norm: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M k f" shows "set_integrable M k (\x. norm (f x))" using integrable_norm f by (force simp add: set_integrable_def) lemma absolutely_integrable_bounded_variation: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on UNIV" obtains B where "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" proof (rule that[of "integral UNIV (\x. norm (f x))"]; safe) fix d :: "'a set set" assume d: "d division_of \d" have *: "k \ d \ f absolutely_integrable_on k" for k using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto note d' = division_ofD[OF d] have "(\k\d. norm (integral k f)) = (\k\d. norm (LINT x:k|lebesgue. f x))" by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *) also have "\ \ (\k\d. LINT x:k|lebesgue. norm (f x))" by (intro sum_mono set_integral_norm_bound *) also have "\ = (\k\d. integral k (\x. norm (f x)))" by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *) also have "\ \ integral (\d) (\x. norm (f x))" using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def by (subst integral_combine_division_topdown[OF _ d]) auto also have "\ \ integral UNIV (\x. norm (f x))" using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def by (intro integral_subset_le) auto finally show "(\k\d. norm (integral k f)) \ integral UNIV (\x. norm (f x))" . qed lemma absdiff_norm_less: assumes "sum (\x. norm (f x - g x)) s < e" and "finite s" shows "\sum (\x. norm(f x)) s - sum (\x. norm(g x)) s\ < e" unfolding sum_subtractf[symmetric] apply (rule le_less_trans[OF sum_abs]) apply (rule le_less_trans[OF _ assms(1)]) apply (rule sum_mono) apply (rule norm_triangle_ineq3) done proposition bounded_variation_absolutely_integrable_interval: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes f: "f integrable_on cbox a b" and *: "\d. d division_of (cbox a b) \ sum (\K. norm(integral K f)) d \ B" shows "f absolutely_integrable_on cbox a b" proof - let ?f = "\d. \K\d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x\?D. ?f x" have *: "\\. gauge \ \ (\p. p tagged_division_of cbox a b \ \ fine p \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e)" if e: "e > 0" for e proof - have "?S - e/2 < ?S" using \e > 0\ by simp then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] have "\e>0. \i\d. x \ i \ ball x e \ i = {}" for x proof - have "\d'>0. \x'\\{i \ d. x \ i}. d' \ dist x x'" proof (rule separate_point_closed) show "closed (\{i \ d. x \ i})" using d' by force show "x \ \{i \ d. x \ i}" by auto qed then show ?thesis by force qed then obtain k where k: "\x. 0 < k x" "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" by metis have "e/2 > 0" using e by auto with Henstock_lemma[OF f] obtain \ where g: "gauge \" "\p. \p tagged_partial_division_of cbox a b; \ fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (metis (no_types, lifting)) let ?g = "\x. \ x \ ball x (k x)" show ?thesis proof (intro exI conjI allI impI) show "gauge ?g" using g(1) k(1) by (auto simp: gauge_def) next fix p assume "p tagged_division_of (cbox a b) \ ?g fine p" then have p: "p tagged_division_of cbox a b" "\ fine p" "(\x. ball x (k x)) fine p" by (auto simp: fine_Int) note p' = tagged_division_ofD[OF p(1)] define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" have gp': "\ fine p'" using p(2) by (auto simp: p'_def fine_def) have p'': "p' tagged_division_of (cbox a b)" proof (rule tagged_division_ofI) show "finite p'" proof (rule finite_subset) show "p' \ (\(k, x, l). (x, k \ l)) ` (d \ p)" by (force simp: p'_def image_iff) show "finite ((\(k, x, l). (x, k \ l)) ` (d \ p))" by (simp add: d'(1) p'(1)) qed next fix x K assume "(x, K) \ p'" then have "\i l. x \ i \ i \ d \ (x, l) \ p \ K = i \ l" unfolding p'_def by auto then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" by blast show "x \ K" and "K \ cbox a b" using p'(2-3)[OF il(3)] il by auto show "\a b. K = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) next fix x1 K1 assume "(x1, K1) \ p'" then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ K1 = i \ l" unfolding p'_def by auto then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "K1 = i1 \ l1" by blast fix x2 K2 assume "(x2,K2) \ p'" then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ K2 = i \ l" unfolding p'_def by auto then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "K2 = i2 \ l2" by blast assume "(x1, K1) \ (x2, K2)" then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) then show "interior K1 \ interior K2 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by blast have "y \ \{K. \x. (x, K) \ p'}" if y: "y \ cbox a b" for y proof - obtain x l where xl: "(x, l) \ p" "y \ l" using y unfolding p'(6)[symmetric] by auto obtain i where i: "i \ d" "y \ i" using y unfolding d'(6)[symmetric] by auto have "x \ i" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show ?thesis unfolding p'_def by (rule_tac X="i \ l" in UnionI) (use i xl in auto) qed show "\{K. \x. (x, K) \ p'} = cbox a b" proof show "\{k. \x. (x, k) \ p'} \ cbox a b" using * by auto next show "cbox a b \ \{k. \x. (x, k) \ p'}" proof fix y assume y: "y \ cbox a b" obtain x L where xl: "(x, L) \ p" "y \ L" using y unfolding p'(6)[symmetric] by auto obtain I where i: "I \ d" "y \ I" using y unfolding d'(6)[symmetric] by auto have "x \ I" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show "y \ \{k. \x. (x, k) \ p'}" apply (rule_tac X="I \ L" in UnionI) using i xl by (auto simp: p'_def) qed qed qed then have sum_less_e2: "(\(x,K) \ p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" using g(2) gp' tagged_division_of_def by blast have "(x, I \ L) \ p'" if x: "(x, L) \ p" "I \ d" and y: "y \ I" "y \ L" for x I L y proof - have "x \ I" using fineD[OF p(3) that(1)] k(2)[OF \I \ d\] y by auto with x have "(\i l. x \ i \ i \ d \ (x, l) \ p \ I \ L = i \ l)" by blast then have "(x, I \ L) \ p'" by (simp add: p'_def) with y show ?thesis by auto qed moreover have "\y i l. (x, K) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" if xK: "(x,K) \ p'" for x K proof - obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" using xK unfolding p'_def by auto then show ?thesis using p'(2) by fastforce qed ultimately have p'alt: "p' = {(x, I \ L) | x I L. (x,L) \ p \ I \ d \ I \ L \ {}}" by auto have sum_p': "(\(x,K) \ p'. norm (integral K f)) = (\k\snd ` p'. norm (integral k f))" apply (subst sum.over_tagged_division_lemma[OF p'',of "\k. norm (integral k f)"]) apply (auto intro: integral_null simp: content_eq_0_interior) done have snd_p_div: "snd ` p division_of cbox a b" by (rule division_of_tagged_division[OF p(1)]) note snd_p = division_ofD[OF snd_p_div] have fin_d_sndp: "finite (d \ snd ` p)" by (simp add: d'(1) snd_p(1)) have *: "\sni sni' sf sf'. \\sf' - sni'\ < e/2; ?S - e/2 < sni; sni' \ ?S; sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" by arith show "norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e" unfolding real_norm_def proof (rule *) show "\(\(x,K)\p'. norm (content K *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) show "(\(x,k) \ p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) show "(\k\d. norm (integral k f)) \ (\(x,k) \ p'. norm (integral k f))" proof - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` (d \ snd ` p)" by auto have "(\K\d. norm (integral K f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" proof (rule sum_mono) fix K assume k: "K \ d" from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" have uvab: "cbox u v \ cbox a b" using d(1) k uv by blast have "d' division_of cbox u v" unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) moreover then have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" by (simp add: sum_norm_le) ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" apply (subst integral_combine_division_topdown[of _ _ d']) apply (auto simp: uv intro: integrable_on_subcbox[OF assms(1) uvab]) done also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof - have *: "norm (integral I f) = 0" if "I \ {cbox u v \ l |l. l \ snd ` p}" "I \ {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" for I using that by auto show ?thesis apply (rule sum.mono_neutral_left) apply (simp add: snd_p(1)) unfolding d'_def uv using * by auto qed also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" proof - have *: "norm (integral (K \ l) f) = 0" if "l \ snd ` p" "y \ snd ` p" "l \ y" "K \ l = K \ y" for l y proof - have "interior (K \ l) \ interior (l \ y)" by (metis Int_lower2 interior_mono le_inf_iff that(4)) then have "interior (K \ l) = {}" by (simp add: snd_p(5) that) moreover from d'(4)[OF k] snd_p(4)[OF that(1)] obtain u1 v1 u2 v2 where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis ultimately show ?thesis using that integral_null unfolding uv Int_interval content_eq_0_interior by (metis (mono_tags, lifting) norm_eq_zero) qed show ?thesis unfolding Setcompr_eq_image apply (rule sum.reindex_nontrivial [unfolded o_def]) apply (rule finite_imageI) apply (rule p') using * by auto qed finally show "norm (integral K f) \ (\l\snd ` p. norm (integral (K \ l) f))" . qed also have "\ = (\(i,l) \ d \ snd ` p. norm (integral (i\l) f))" by (simp add: sum.cartesian_product) also have "\ = (\x \ d \ snd ` p. norm (integral (case_prod (\) x) f))" by (force simp: split_def intro!: sum.cong) also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" proof - have eq0: " (integral (l1 \ k1) f) = 0" if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" for l1 l2 k1 k2 j1 j2 proof - obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" by (simp add: that(1)) ultimately have "interior(l1 \ k1) = {}" by auto then show ?thesis unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed show ?thesis unfolding * apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def]) apply clarsimp by (metis eq0 fst_conv snd_conv) qed also have "\ = (\(x,k) \ p'. norm (integral k f))" proof - have 0: "integral (ia \ snd (a, b)) f = 0" if "ia \ snd (a, b) \ snd ` p'" "ia \ d" "(a, b) \ p" for ia a b proof - have "ia \ b = {}" using that unfolding p'alt image_iff Bex_def not_ex apply (erule_tac x="(a, ia \ b)" in allE) apply auto done then show ?thesis by auto qed have 1: "\i l. snd (a, b) = i \ l \ i \ d \ l \ snd ` p" if "(a, b) \ p'" for a b using that apply (clarsimp simp: p'_def image_iff) by (metis (no_types, hide_lams) snd_conv) show ?thesis unfolding sum_p' apply (rule sum.mono_neutral_right) apply (metis * finite_imageI[OF fin_d_sndp]) using 0 1 by auto qed finally show ?thesis . qed show "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" by force have fin_pd: "finite (p \ d)" using finite_cartesian_product[OF p'(1) d'(1)] by metis have "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ ?S. \content k\ * norm (f x))" unfolding norm_scaleR apply (rule sum.mono_neutral_left) apply (subst *) apply (rule finite_imageI [OF fin_pd]) unfolding p'alt apply auto by fastforce also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" proof - have "\content (l1 \ k1)\ * norm (f x1) = 0" if "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" "x1 = x2" "l1 \ k1 = l2 \ k2" "x1 \ x2 \ l1 \ l2 \ k1 \ k2" for x1 l1 k1 x2 l2 k2 proof - obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" by (meson \(x1, l1) \ p\ \k1 \ d\ d(1) division_ofD(4) p'(4)) have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" apply (rule disjE) using that p'(5) d'(5) by auto moreover have "interior (l1 \ k1) = interior (l2 \ k2)" unfolding that .. ultimately have "interior (l1 \ k1) = {}" by auto then show "\content (l1 \ k1)\ * norm (f x1) = 0" unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed then show ?thesis unfolding * apply (subst sum.reindex_nontrivial [OF fin_pd]) unfolding split_paired_all o_def split_def prod.inject apply force+ done qed also have "\ = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - have sumeq: "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" if "(x, l) \ p" for x l proof - note xl = p'(2-4)[OF that] then obtain u v where uv: "l = cbox u v" by blast have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" by (simp add: Int_commute uv) also have "\ = sum content {k \ cbox u v| k. k \ d}" proof - have eq0: "content (k \ cbox u v) = 0" if "k \ d" "y \ d" "k \ y" and eq: "k \ cbox u v = y \ cbox u v" for k y proof - from d'(4)[OF that(1)] d'(4)[OF that(2)] obtain \ \ where \: "k \ cbox u v = cbox \ \" by (meson Int_interval) have "{} = interior ((k \ y) \ cbox u v)" by (simp add: d'(5) that) also have "\ = interior (y \ (k \ cbox u v))" by auto also have "\ = interior (k \ cbox u v)" unfolding eq by auto finally show ?thesis unfolding \ content_eq_0_interior .. qed then show ?thesis unfolding Setcompr_eq_image apply (rule sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) by auto qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" apply (rule sum.mono_neutral_right) unfolding Setcompr_eq_image apply (rule finite_imageI [OF \finite d\]) apply (fastforce simp: inf.commute)+ done finally show "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" unfolding sum_distrib_right[symmetric] real_scaleR_def apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) using xl(2)[unfolded uv] unfolding uv apply auto done qed show ?thesis by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d') qed finally show ?thesis . qed qed (rule d) qed qed then show ?thesis using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] by blast qed lemma bounded_variation_absolutely_integrable: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" and "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" shows "f absolutely_integrable_on UNIV" proof (rule absolutely_integrable_onI, fact) let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (\d)}" have D_1: "?D \ {}" by (rule elementary_interval) auto have D_2: "bdd_above (?f`?D)" by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp note D = D_1 D_2 let ?S = "SUP d\?D. ?f d" have "\a b. f integrable_on cbox a b" using assms(1) integrable_on_subcbox by blast then have f_int: "\a b. f absolutely_integrable_on cbox a b" apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) using assms(2) apply blast done have "((\x. norm (f x)) has_integral ?S) UNIV" apply (subst has_integral_alt') apply safe proof goal_cases case (1 a b) show ?case using f_int[of a b] unfolding absolutely_integrable_on_def by auto next case prems: (2 e) have "\y\sum (\k. norm (integral k f)) ` {d. d division_of \d}. \ y \ ?S - e" proof (rule ccontr) assume "\ ?thesis" then have "?S \ ?S - e" by (intro cSUP_least[OF D(1)]) auto then show False using prems by auto qed then obtain d K where ddiv: "d division_of \d" and "K = (\k\d. norm (integral k f))" "Sup (sum (\k. norm (integral k f)) ` {d. d division_of \ d}) - e < K" by (auto simp add: image_iff not_le) then have d: "Sup (sum (\k. norm (integral k f)) ` {d. d division_of \ d}) - e < (\k\d. norm (integral k f))" by auto note d'=division_ofD[OF ddiv] have "bounded (\d)" by (rule elementary_bounded,fact) from this[unfolded bounded_pos] obtain K where K: "0 < K" "\x\\d. norm x \ K" by auto show ?case proof (intro conjI impI allI exI) fix a b :: 'n assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. \?S - e < s1; s1 \ s; s < ?S + e\ \ \s - ?S\ < e" by arith show "norm (integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" unfolding real_norm_def proof (rule * [OF d]) have "(\k\d. norm (integral k f)) \ sum (\k. integral k (\x. norm (f x))) d" proof (intro sum_mono) fix k assume "k \ d" with d'(4) f_int show "norm (integral k f) \ integral k (\x. norm (f x))" by (force simp: absolutely_integrable_on_def integral_norm_bound_integral) qed also have "\ = integral (\d) (\x. norm (f x))" apply (rule integral_combine_division_bottomup[OF ddiv, symmetric]) using absolutely_integrable_on_def d'(4) f_int by blast also have "\ \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" proof - have "\d \ cbox a b" using K(2) ab by fastforce then show ?thesis using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def by (auto intro!: integral_subset_le) qed finally show "(\k\d. norm (integral k f)) \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" . next have "e/2>0" using \e > 0\ by auto moreover have f: "f integrable_on cbox a b" "(\x. norm (f x)) integrable_on cbox a b" using f_int by (auto simp: absolutely_integrable_on_def) ultimately obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of (cbox a b); d1 fine p\ \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e/2" unfolding has_integral_integral has_integral by meson obtain d2 where "gauge d2" and d2: "\p. \p tagged_partial_division_of (cbox a b); d2 fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (blast intro: Henstock_lemma [OF f(1) \e/2>0\]) obtain p where p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_Int [OF \gauge d1\ \gauge d2\], of a b]) (auto simp add: fine_Int) have *: "\sf sf' si di. \sf' = sf; si \ ?S; \sf - si\ < e/2; \sf' - di\ < e/2\ \ di < ?S + e" by arith have "integral (cbox a b) (\x. norm (f x)) < ?S + e" proof (rule *) show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e/2" unfolding split_def apply (rule absdiff_norm_less) using d2[of p] p(1,3) apply (auto simp: tagged_division_of_def split_def) done show "\(\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e/2" using d1[OF p(1,2)] by (simp only: real_norm_def) show "(\(x,k) \ p. content k *\<^sub>R norm (f x)) = (\(x,k) \ p. norm (content k *\<^sub>R f x))" by (auto simp: split_paired_all sum.cong [OF refl]) show "(\(x,k) \ p. norm (integral k f)) \ ?S" using partial_division_of_tagged_division[of p "cbox a b"] p(1) apply (subst sum.over_tagged_division_lemma[OF p(1)]) apply (auto simp: content_eq_0_interior tagged_partial_division_of_def intro!: cSUP_upper2 D) done qed then show "integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" by simp qed qed (insert K, auto) qed then show "(\x. norm (f x)) integrable_on UNIV" by blast qed subsection\Outer and inner approximation of measurable sets by well-behaved sets.\ proposition measurable_outer_intervals_bounded: assumes "S \ lmeasurable" "S \ cbox a b" "e > 0" obtains \ where "countable \" "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i)/2^n" "\K. \K \ \; box a b \ {}\ \ interior K \ {}" "S \ \\" "\\ \ lmeasurable" "measure lebesgue (\\) \ measure lebesgue S + e" proof (cases "box a b = {}") case True show ?thesis proof (cases "cbox a b = {}") case True with assms have [simp]: "S = {}" by auto show ?thesis proof show "countable {}" by simp qed (use \e > 0\ in auto) next case False show ?thesis proof show "countable {cbox a b}" by simp show "\u v. cbox u v \ {cbox a b} \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i)/2 ^ n" using False by (force simp: eq_cbox intro: exI [where x=0]) show "measure lebesgue (\{cbox a b}) \ measure lebesgue S + e" using assms by (simp add: sum_content.box_empty_imp [OF True]) qed (use assms \cbox a b \ {}\ in auto) qed next case False let ?\ = "measure lebesgue" have "S \ cbox a b \ lmeasurable" using \S \ lmeasurable\ by blast then have indS_int: "(indicator S has_integral (?\ S)) (cbox a b)" by (metis integral_indicator \S \ cbox a b\ has_integral_integrable_integral inf.orderE integrable_on_indicator) with \e > 0\ obtain \ where "gauge \" and \: "\\. \\ tagged_division_of (cbox a b); \ fine \\ \ norm ((\(x,K)\\. content(K) *\<^sub>R indicator S x) - ?\ S) < e" by (force simp: has_integral) have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) obtain \ where \: "countable \" "\\ \ cbox a b" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ \ x" and close: "\u v. cbox u v \ \ \ \n. \i \ Basis. v\i - u\i = (b\i - a\i)/2^n" and covers: "S \ \\" using covering_lemma [of S a b \] \gauge \\ \box a b \ {}\ assms by force show ?thesis proof show "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" by (meson Sup_le_iff \(2) cbox interior_empty) have negl_int: "negligible(K \ L)" if "K \ \" "L \ \" "K \ L" for K L proof - have "interior K \ interior L = {}" using djointish pairwiseD that by fastforce moreover obtain u v x y where "K = cbox u v" "L = cbox x y" using cbox \K \ \\ \L \ \\ by blast ultimately show ?thesis by (simp add: Int_interval box_Int_box negligible_interval(1)) qed have fincase: "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" if "finite \" "\ \ \" for \ proof - obtain t where t: "\K. K \ \ \ t K \ S \ K \ K \ \(t K)" using covered \\ \ \\ subsetD by metis have "\K \ \. \L \ \. K \ L \ interior K \ interior L = {}" using that djointish by (simp add: pairwise_def) (metis subsetD) with cbox that \ have \div: "\ division_of (\\)" by (fastforce simp: division_of_def dest: cbox) then have 1: "\\ \ lmeasurable" by blast have norme: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x,K)\p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" by (auto simp: lmeasure_integral_UNIV assms inteq dest: \) have "\x K y L. (x,K) \ (\K. (t K,K)) ` \ \ (y,L) \ (\K. (t K,K)) ` \ \ (x,K) \ (y,L) \ interior K \ interior L = {}" using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) with that \ have tagged: "(\K. (t K, K)) ` \ tagged_partial_division_of cbox a b" by (auto simp: tagged_partial_division_of_def dest: t cbox) have fine: "\ fine (\K. (t K, K)) ` \" using t by (auto simp: fine_def) have *: "y \ ?\ S \ \x - y\ \ e \ x \ ?\ S + e" for x y by arith have "?\ (\\) \ ?\ S + e" proof (rule *) have "(\K\\. ?\ (K \ S)) = ?\ (\C\\. C \ S)" apply (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) using \div \S \ lmeasurable\ apply blast unfolding pairwise_def by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) also have "\ = ?\ (\\ \ S)" by simp also have "\ \ ?\ S" by (simp add: "1" \S \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Int) finally show "(\K\\. ?\ (K \ S)) \ ?\ S" . next have "?\ (\\) = sum ?\ \" by (metis \div content_division) also have "\ = (\K\\. content K)" using \div by (force intro: sum.cong) also have "\ = (\x\\. content x * indicator S (t x))" using t by auto finally have eq1: "?\ (\\) = (\x\\. content x * indicator S (t x))" . have eq2: "(\K\\. ?\ (K \ S)) = (\K\\. integral K (indicator S))" apply (rule sum.cong [OF refl]) by (metis integral_indicator \div \S \ lmeasurable\ division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) have "\\(x,K)\(\K. (t K, K)) ` \. content K * indicator S x - integral K (indicator S)\ \ e" using Henstock_lemma_part1 [of "indicator S::'a\real", OF _ \e > 0\ \gauge \\ _ tagged fine] indS_int norme by auto then show "\?\ (\\) - (\K\\. ?\ (K \ S))\ \ e" by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) qed with 1 show ?thesis by blast qed have "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" proof (cases "finite \") case True with fincase show ?thesis by blast next case False let ?T = "from_nat_into \" have T: "bij_betw ?T UNIV \" by (simp add: False \(1) bij_betw_from_nat_into) have TM: "\n. ?T n \ lmeasurable" by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) have TN: "\m n. m \ n \ negligible (?T m \ ?T n)" by (simp add: False \(1) from_nat_into infinite_imp_nonempty negl_int) have TB: "(\k\n. ?\ (?T k)) \ ?\ S + e" for n proof - have "(\k\n. ?\ (?T k)) = ?\ (\ (?T ` {..n}))" by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) also have "?\ (\ (?T ` {..n})) \ ?\ S + e" using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def) finally show ?thesis . qed have "\\ \ lmeasurable" by (metis lmeasurable_compact T \(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) moreover have "?\ (\x. from_nat_into \ x) \ ?\ S + e" proof (rule measure_countable_Union_le [OF TM]) show "?\ (\x\n. from_nat_into \ x) \ ?\ S + e" for n by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) qed ultimately show ?thesis by (metis T bij_betw_def) qed then show "\\ \ lmeasurable" "measure lebesgue (\\) \ ?\ S + e" by blast+ qed (use \ cbox djointish close covers in auto) qed subsection\Transformation of measure by linear maps\ +lemma emeasure_lebesgue_ball_conv_unit_ball: + fixes c :: "'a :: euclidean_space" + assumes "r \ 0" + shows "emeasure lebesgue (ball c r) = + ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" +proof (cases "r = 0") + case False + with assms have r: "r > 0" by auto + have "emeasure lebesgue ((\x. c + x) ` (\x. r *\<^sub>R x) ` ball (0 :: 'a) 1) = + r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)" + unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms + by (simp add: add_ac) + also have "(\x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r" + using r by (subst ball_scale) auto + also have "(\x. c + x) ` \ = ball c r" + by (subst image_add_ball) (simp_all add: algebra_simps) + finally show ?thesis by simp +qed auto + +lemma content_ball_conv_unit_ball: + fixes c :: "'a :: euclidean_space" + assumes "r \ 0" + shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)" +proof - + have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)" + using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto + also have "\ = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" + using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto + also have "\ = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))" + using emeasure_lborel_ball_finite[of "0::'a" 1] assms + by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult') + finally show ?thesis + using assms by (subst (asm) ennreal_inj) auto +qed + lemma measurable_linear_image_interval: "linear f \ f ` (cbox a b) \ lmeasurable" by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact) proposition measure_linear_sufficient: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" and S: "S \ lmeasurable" and im: "\a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)" shows "f ` S \ lmeasurable \ m * measure lebesgue S = measure lebesgue (f ` S)" using le_less_linear [of 0 m] proof assume "m < 0" then show ?thesis using im [of 0 One] by auto next assume "m \ 0" let ?\ = "measure lebesgue" show ?thesis proof (cases "inj f") case False then have "?\ (f ` S) = 0" using \linear f\ negligible_imp_measure0 negligible_linear_singular_image by blast then have "m * ?\ (cbox 0 (One)) = 0" by (metis False \linear f\ cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel) then show ?thesis using \linear f\ negligible_linear_singular_image negligible_imp_measure0 False by (auto simp: lmeasurable_iff_has_integral negligible_UNIV) next case True then obtain h where "linear h" and hf: "\x. h (f x) = x" and fh: "\x. f (h x) = x" using \linear f\ linear_injective_isomorphism by blast have fBS: "(f ` S) \ lmeasurable \ m * ?\ S = ?\ (f ` S)" if "bounded S" "S \ lmeasurable" for S proof - obtain a b where "S \ cbox a b" using \bounded S\ bounded_subset_cbox_symmetric by metis have fUD: "(f ` \\) \ lmeasurable \ ?\ (f ` \\) = (m * ?\ (\\))" if "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intint: "pairwise (\A B. interior A \ interior B = {}) \" for \ proof - have conv: "\K. K \ \ \ convex K" using cbox convex_box(1) by blast have neg: "negligible (g ` K \ g ` L)" if "linear g" "K \ \" "L \ \" "K \ L" for K L and g :: "'n\'n" proof (cases "inj g") case True have "negligible (frontier(g ` K \ g ` L) \ interior(g ` K \ g ` L))" proof (rule negligible_Un) show "negligible (frontier (g ` K \ g ` L))" by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) next have "\p N. pairwise p N = (\Na. (Na::'n set) \ N \ (\Nb. Nb \ N \ Na \ Nb \ p Na Nb))" by (metis pairwise_def) then have "interior K \ interior L = {}" using intint that(2) that(3) that(4) by presburger then show "negligible (interior (g ` K \ g ` L))" by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) qed moreover have "g ` K \ g ` L \ frontier (g ` K \ g ` L) \ interior (g ` K \ g ` L)" apply (auto simp: frontier_def) using closure_subset contra_subsetD by fastforce+ ultimately show ?thesis by (rule negligible_subset) next case False then show ?thesis by (simp add: negligible_Int negligible_linear_singular_image \linear g\) qed have negf: "negligible ((f ` K) \ (f ` L))" and negid: "negligible (K \ L)" if "K \ \" "L \ \" "K \ L" for K L using neg [OF \linear f\] neg [OF linear_id] that by auto show ?thesis proof (cases "finite \") case True then have "?\ (\x\\. f ` x) = (\x\\. ?\ (f ` x))" using \linear f\ cbox measurable_linear_image_interval negf by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def]) also have "\ = (\k\\. m * ?\ k)" by (metis (no_types, lifting) cbox im sum.cong) also have "\ = m * ?\ (\\)" unfolding sum_distrib_left [symmetric] by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid) finally show ?thesis by (metis True \linear f\ cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval) next case False with \countable \\ obtain X :: "nat \ 'n set" where S: "bij_betw X UNIV \" using bij_betw_from_nat_into by blast then have eq: "(\\) = (\n. X n)" "(f ` \\) = (\n. f ` X n)" by (auto simp: bij_betw_def) have meas: "\K. K \ \ \ K \ lmeasurable" using cbox by blast with S have 1: "\n. X n \ lmeasurable" by (auto simp: bij_betw_def) have 2: "pairwise (\m n. negligible (X m \ X n)) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have 3: "bounded (\n. X n)" using S unfolding bij_betw_def by blast have "(\n. X n) \ lmeasurable" by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3]) with S have f1: "\n. f ` (X n) \ lmeasurable" unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI) have f2: "pairwise (\m n. negligible (f ` (X m) \ f ` (X n))) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have f3: "bounded (\n. f ` X n)" using S unfolding bij_betw_def by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition) have "(\n. ?\ (X n)) sums ?\ (\n. X n)" by (rule measure_countable_negligible_Union_bounded [OF 1 2 3]) have meq: "?\ (\n. f ` X n) = m * ?\ (\(X ` UNIV))" proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]]) have m: "\n. ?\ (f ` X n) = (m * ?\ (X n))" using S unfolding bij_betw_def by (metis cbox im rangeI) show "(\n. ?\ (f ` X n)) sums (m * ?\ (\(X ` UNIV)))" unfolding m using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast qed show ?thesis using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq by (auto simp: eq [symmetric]) qed qed show ?thesis unfolding completion.fmeasurable_measure_inner_outer_le proof (intro conjI allI impI) fix e :: real assume "e > 0" have 1: "cbox a b - S \ lmeasurable" by (simp add: fmeasurable.Diff that) have 2: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "cbox a b - S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ (cbox a b - S) + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \m\)"]; use 1 2 pairwise_def in force) have meq: "?\ (cbox a b - S) = ?\ (cbox a b) - ?\ S" by (simp add: measurable_measure_Diff \S \ cbox a b\ fmeasurableD that(2)) show "\T \ lmeasurable. T \ f ` S \ m * ?\ S - e \ ?\ T" proof (intro bexI conjI) show "f ` (cbox a b) - f ` (\\) \ f ` S" using \cbox a b - S \ \\\ by force have "m * ?\ S - e \ m * (?\ S - e / (1 + \m\))" using \m \ 0\ \e > 0\ by (simp add: field_simps) also have "\ \ ?\ (f ` cbox a b) - ?\ (f ` (\\))" using le \m \ 0\ \e > 0\ apply (simp add: im fUD [OF \countable \\ cbox intdisj] right_diff_distrib [symmetric]) apply (rule mult_left_mono; simp add: algebra_simps meq) done also have "\ = ?\ (f ` cbox a b - f ` \\)" apply (rule measurable_measure_Diff [symmetric]) apply (simp add: assms(1) measurable_linear_image_interval) apply (simp add: \countable \\ cbox fUD fmeasurableD intdisj) apply (simp add: Sup_le_iff cbox image_mono) done finally show "m * ?\ S - e \ ?\ (f ` cbox a b - f ` \\)" . show "f ` cbox a b - f ` \\ \ lmeasurable" by (simp add: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) qed next fix e :: real assume "e > 0" have em: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ S + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \m\)"]; use \S \ lmeasurable\ \S \ cbox a b\ em in force) show "\U \ lmeasurable. f ` S \ U \ ?\ U \ m * ?\ S + e" proof (intro bexI conjI) show "f ` S \ f ` (\\)" by (simp add: DD(1) image_mono) have "?\ (f ` \\) \ m * (?\ S + e / (1 + \m\))" using \m \ 0\ le mult_left_mono by (auto simp: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) also have "\ \ m * ?\ S + e" using \m \ 0\ \e > 0\ by (simp add: fUD [OF \countable \\ cbox intdisj] field_simps) finally show "?\ (f ` \\) \ m * ?\ S + e" . show "f ` \\ \ lmeasurable" by (simp add: \countable \\ cbox fUD intdisj) qed qed qed show ?thesis unfolding has_measure_limit_iff proof (intro allI impI) fix e :: real assume "e > 0" obtain B where "B > 0" and B: "\a b. ball 0 B \ cbox a b \ \?\ (S \ cbox a b) - ?\ S\ < e / (1 + \m\)" using has_measure_limit [OF S] \e > 0\ by (metis abs_add_one_gt_zero zero_less_divide_iff) obtain c d::'n where cd: "ball 0 B \ cbox c d" by (metis bounded_subset_cbox_symmetric bounded_ball) with B have less: "\?\ (S \ cbox c d) - ?\ S\ < e / (1 + \m\)" . obtain D where "D > 0" and D: "cbox c d \ ball 0 D" by (metis bounded_cbox bounded_subset_ballD) obtain C where "C > 0" and C: "\x. norm (f x) \ C * norm x" using linear_bounded_pos \linear f\ by blast have "f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" if "ball 0 (D*C) \ cbox a b" for a b proof - have "bounded (S \ h ` cbox a b)" by (simp add: bounded_linear_image linear_linear \linear h\ bounded_Int) moreover have Shab: "S \ h ` cbox a b \ lmeasurable" by (simp add: S \linear h\ fmeasurable.Int measurable_linear_image_interval) moreover have fim: "f ` (S \ h ` (cbox a b)) = (f ` S) \ cbox a b" by (auto simp: hf rev_image_eqI fh) ultimately have 1: "(f ` S) \ cbox a b \ lmeasurable" and 2: "m * ?\ (S \ h ` cbox a b) = ?\ ((f ` S) \ cbox a b)" using fBS [of "S \ (h ` (cbox a b))"] by auto have *: "\\z - m\ < e; z \ w; w \ m\ \ \w - m\ \ e" for w z m and e::real by auto have meas_adiff: "\?\ (S \ h ` cbox a b) - ?\ S\ \ e / (1 + \m\)" proof (rule * [OF less]) show "?\ (S \ cbox c d) \ ?\ (S \ h ` cbox a b)" proof (rule measure_mono_fmeasurable [OF _ _ Shab]) have "f ` ball 0 D \ ball 0 (C * D)" using C \C > 0\ apply (clarsimp simp: algebra_simps) by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono) then have "f ` ball 0 D \ cbox a b" by (metis mult.commute order_trans that) have "ball 0 D \ h ` cbox a b" by (metis \f ` ball 0 D \ cbox a b\ hf image_subset_iff subsetI) then show "S \ cbox c d \ S \ h ` cbox a b" using D by blast next show "S \ cbox c d \ sets lebesgue" using S fmeasurable_cbox by blast qed next show "?\ (S \ h ` cbox a b) \ ?\ S" by (simp add: S Shab fmeasurableD measure_mono_fmeasurable) qed have "\?\ (f ` S \ cbox a b) - m * ?\ S\ \ m * e / (1 + \m\)" proof - have mm: "\m\ = m" by (simp add: \0 \ m\) then have "\?\ S - ?\ (S \ h ` cbox a b)\ * m \ e / (1 + m) * m" by (metis (no_types) \0 \ m\ meas_adiff abs_minus_commute mult_right_mono) moreover have "\r. \r * m\ = m * \r\" by (metis \0 \ m\ abs_mult_pos mult.commute) ultimately show ?thesis apply (simp add: 2 [symmetric]) by (metis (no_types) abs_minus_commute mult.commute right_diff_distrib' mm) qed also have "\ < e" using \e > 0\ by (cases "m \ 0") (simp_all add: field_simps) finally have "\?\ (f ` S \ cbox a b) - m * ?\ S\ < e" . with 1 show ?thesis by auto qed then show "\B>0. \a b. ball 0 B \ cbox a b \ f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" using \C>0\ \D>0\ by (metis mult_zero_left real_mult_less_iff1) qed qed qed subsection\Lemmas about absolute integrability\ lemma absolutely_integrable_linear: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and h :: "'n::euclidean_space \ 'p::euclidean_space" shows "f absolutely_integrable_on s \ bounded_linear h \ (h \ f) absolutely_integrable_on s" using integrable_bounded_linear[of h lebesgue "\x. indicator s x *\<^sub>R f x"] by (simp add: linear_simps[of h] set_integrable_def) lemma absolutely_integrable_sum: fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" assumes "finite T" and "\a. a \ T \ (f a) absolutely_integrable_on S" shows "(\x. sum (\a. f a x) T) absolutely_integrable_on S" using assms by induction auto lemma absolutely_integrable_integrable_bound: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes le: "\x. x\S \ norm (f x) \ g x" and f: "f integrable_on S" and g: "g integrable_on S" shows "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) have "g absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (g x)) integrable_on S" using le norm_ge_zero[of "f _"] by (intro integrable_spike_finite[OF _ _ g, of "{}"]) (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero) qed fact then show "integrable lebesgue (\x. indicat_real S x *\<^sub>R g x)" by (simp add: set_integrable_def) show "(\x. indicat_real S x *\<^sub>R f x) \ borel_measurable lebesgue" using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) qed (use le in \force intro!: always_eventually split: split_indicator\) corollary absolutely_integrable_on_const [simp]: fixes c :: "'a::euclidean_space" assumes "S \ lmeasurable" shows "(\x. c) absolutely_integrable_on S" by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) lemma absolutely_integrable_continuous: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "continuous_on (cbox a b) f \ f absolutely_integrable_on cbox a b" using absolutely_integrable_integrable_bound by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous) lemma absolutely_integrable_continuous_real: fixes f :: "real \ 'b::euclidean_space" shows "continuous_on {a..b} f \ f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_continuous box_real(2)) lemma continuous_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on (cbox a b) f" shows "integrable (lebesgue_on (cbox a b)) f" proof - have "f absolutely_integrable_on cbox a b" by (simp add: absolutely_integrable_continuous assms) then show ?thesis by (simp add: integrable_restrict_space set_integrable_def) qed lemma continuous_imp_integrable_real: fixes f :: "real \ 'b::euclidean_space" assumes "continuous_on {a..b} f" shows "integrable (lebesgue_on {a..b}) f" by (metis assms continuous_imp_integrable interval_cbox) subsection \Componentwise\ proposition absolutely_integrable_componentwise_iff: shows "f absolutely_integrable_on A \ (\b\Basis. (\x. f x \ b) absolutely_integrable_on A)" proof - have *: "(\x. norm (f x)) integrable_on A \ (\b\Basis. (\x. norm (f x \ b)) integrable_on A)" if "f integrable_on A" proof - have 1: "\i. \(\x. norm (f x)) integrable_on A; i \ Basis\ \ (\x. f x \ i) absolutely_integrable_on A" apply (rule absolutely_integrable_integrable_bound [where g = "\x. norm(f x)"]) using Basis_le_norm integrable_component that apply fastforce+ done have 2: "\i\Basis. (\x. \f x \ i\) integrable_on A \ f absolutely_integrable_on A" apply (rule absolutely_integrable_integrable_bound [where g = "\x. \i\Basis. norm (f x \ i)"]) using norm_le_l1 that apply (force intro: integrable_sum)+ done show ?thesis apply auto apply (metis (full_types) absolutely_integrable_on_def set_integrable_abs 1) apply (metis (full_types) absolutely_integrable_on_def 2) done qed show ?thesis unfolding absolutely_integrable_on_def by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong) qed lemma absolutely_integrable_componentwise: shows "(\b. b \ Basis \ (\x. f x \ b) absolutely_integrable_on A) \ f absolutely_integrable_on A" using absolutely_integrable_componentwise_iff by blast lemma absolutely_integrable_component: "f absolutely_integrable_on A \ (\x. f x \ (b :: 'b :: euclidean_space)) absolutely_integrable_on A" by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def) lemma absolutely_integrable_scaleR_left: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S" proof - have "(\x. c *\<^sub>R x) o f absolutely_integrable_on S" apply (rule absolutely_integrable_linear [OF assms]) by (simp add: bounded_linear_scaleR_right) then show ?thesis using assms by blast qed lemma absolutely_integrable_scaleR_right: assumes "f absolutely_integrable_on S" shows "(\x. f x *\<^sub>R c) absolutely_integrable_on S" using assms by blast lemma absolutely_integrable_norm: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(norm o f) absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def o_def) lemma absolutely_integrable_abs: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) absolutely_integrable_on S" (is "?g absolutely_integrable_on S") proof - have eq: "?g = (\x. \i\Basis. ((\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm(\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f) x)" by (simp add: sum.delta) have *: "(\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" if "i \ Basis" for i proof - have "bounded_linear (\y. \j\Basis. if j = i then y *\<^sub>R j else 0)" by (simp add: linear_linear algebra_simps linearI) moreover have "(\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" unfolding o_def apply (rule absolutely_integrable_norm [unfolded o_def]) using assms \i \ Basis\ apply (auto simp: algebra_simps dest: absolutely_integrable_component[where b=i]) done ultimately show ?thesis by (subst comp_assoc) (blast intro: absolutely_integrable_linear) qed show ?thesis apply (rule ssubst [OF eq]) apply (rule absolutely_integrable_sum) apply (force simp: intro!: *)+ done qed lemma abs_absolutely_integrableI_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "(\x. \f x\) integrable_on A" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto lemma abs_absolutely_integrableI: assumes f: "f integrable_on S" and fcomp: "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" shows "f absolutely_integrable_on S" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" if "i \ Basis" for i proof - have "(\x. \f x \ i\) integrable_on S" using assms integrable_component [OF fcomp, where y=i] that by simp then have "(\x. f x \ i) absolutely_integrable_on S" using abs_absolutely_integrableI_1 f integrable_component by blast then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_abs_iff: "f absolutely_integrable_on S \ f integrable_on S \ (\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using absolutely_integrable_abs absolutely_integrable_on_def by blast next assume ?rhs moreover have "(\x. if x \ S then \i\Basis. \f x \ i\ *\<^sub>R i else 0) = (\x. \i\Basis. \(if x \ S then f x else 0) \ i\ *\<^sub>R i)" by force ultimately show ?lhs by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI) qed lemma absolutely_integrable_max: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i + \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i + \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib algebra_simps euclidean_representation) finally show "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" apply (intro set_integral_add absolutely_integrable_scaleR_left assms) using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] apply (simp add: algebra_simps) done ultimately show ?thesis by metis qed corollary absolutely_integrable_max_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. max (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_max [OF assms] by simp lemma absolutely_integrable_min: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i - \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i - \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation) finally show "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" apply (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] apply (simp add: algebra_simps) done ultimately show ?thesis by metis qed corollary absolutely_integrable_min_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. min (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_min [OF assms] by simp lemma nonnegative_absolutely_integrable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ 0 \ f x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" if "i \ Basis" for i proof - have "(\x. f x \ i) integrable_on A" by (simp add: assms(1) integrable_component) then have "(\x. f x \ i) absolutely_integrable_on A" by (metis that comp nonnegative_absolutely_integrable_1) then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_component_ubound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. g x - (g x - f x)) absolutely_integrable_on A" apply (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast by (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma absolutely_integrable_component_lbound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "g absolutely_integrable_on A" proof - have "(\x. f x + (g x - f x)) absolutely_integrable_on A" apply (rule set_integral_add [OF f nonnegative_absolutely_integrable]) using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast by (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f integrable_on S \ (\x. f x $ 1) integrable_on S" by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis) lemma integral_on_1_eq: fixes f :: "'a::euclidean_space \ real^1" shows "integral S f = vec (integral S (\x. f x $ 1))" by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral) lemma absolutely_integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f absolutely_integrable_on S \ (\x. f x $ 1) absolutely_integrable_on S" unfolding absolutely_integrable_on_def by (auto simp: integrable_on_1_iff norm_real) lemma absolutely_integrable_absolutely_integrable_lbound: fixes f :: "'m::euclidean_space \ real" assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S" and *: "\x. x \ S \ g x \ f x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *) lemma absolutely_integrable_absolutely_integrable_ubound: fixes f :: "'m::euclidean_space \ real" assumes fg: "f integrable_on S" "g absolutely_integrable_on S" and *: "\x. x \ S \ f x \ g x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *) lemma has_integral_vec1_I_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) (cbox a b)" shows "((f \ vec) has_integral y) {a$1..b$1}" proof - have "((\x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\x. x $ 1) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real^1. vec ` cbox u v = cbox w z" "content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z. (\x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1" using vec_nth_cbox_1_eq by blast qed (auto simp: continuous_vec assms) then show ?thesis by (simp add: o_def) qed lemma has_integral_vec1_I: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) S" shows "(f \ vec has_integral y) ((\x. x $ 1) ` S)" proof - have *: "\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e" if int: "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)" and B: "ball 0 B \ {a..b}" for e B a b proof - have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force have B': "ball (0::real^1) B \ cbox (vec a) (vec b)" using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff) show ?thesis using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_I_cbox split: if_split_asm) apply (metis vector_one_nth) apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+ apply (blast intro!: *) done qed lemma has_integral_vec1_nth_cbox: fixes f :: "real \ 'a::real_normed_vector" assumes "(f has_integral y) {a..b}" shows "((\x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))" proof - have "((\x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real. (\x. x $ 1) ` cbox u v = cbox w z" "content ((\x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1" unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real" using vec_cbox_1_eq by auto qed (auto simp: continuous_vec assms) then show ?thesis using vec_cbox_1_eq by auto qed lemma has_integral_vec1_D_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) {a$1..b$1}" shows "(f has_integral y) (cbox a b)" by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth) lemma has_integral_vec1_D: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) ((\x. x $ 1) ` S)" shows "(f has_integral y) S" proof - have *: "\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e" if int: "\a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e)" and B: "ball 0 B \ cbox a b" for e B and a b::"real^1" proof - have B': "ball 0 B \ {a$1..b$1}" using B apply (simp add: subset_iff Basis_vec_def cart_eq_inner_axis [symmetric] mem_box) apply (metis (full_types) norm_real vec_component) done have eq: "(\x. if vec x \ S then f (vec x) else 0) = (\x. if x \ S then f x else 0) \ vec" by force have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force show ?thesis using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast) apply (intro conjI impI) apply (metis vector_one_nth) apply (erule thin_rl) apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+ using * apply blast done qed lemma integral_vec1_eq: fixes f :: "real^1 \ 'a::real_normed_vector" shows "integral S f = integral ((\x. x $ 1) ` S) (f \ vec)" using has_integral_vec1_I [of f] has_integral_vec1_D [of f] by (metis has_integral_iff not_integrable_integral) lemma absolutely_integrable_drop: fixes f :: "real^1 \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ (f \ vec) absolutely_integrable_on (\x. x $ 1) ` S" unfolding absolutely_integrable_on_def integrable_on_def proof safe fix y r assume "(f has_integral y) S" "((\x. norm (f x)) has_integral r) S" then show "\y. (f \ vec has_integral y) ((\x. x $ 1) ` S)" "\y. ((\x. norm ((f \ vec) x)) has_integral y) ((\x. x $ 1) ` S)" by (force simp: o_def dest!: has_integral_vec1_I)+ next fix y :: "'b" and r :: "real" assume "(f \ vec has_integral y) ((\x. x $ 1) ` S)" "((\x. norm ((f \ vec) x)) has_integral r) ((\x. x $ 1) ` S)" then show "\y. (f has_integral y) S" "\y. ((\x. norm (f x)) has_integral y) S" by (force simp: o_def intro: has_integral_vec1_D)+ qed subsection \Dominated convergence\ lemma dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. (f k) integrable_on S" and h: "h integrable_on S" and le: "\k x. x \ S \ norm (f k x) \ h x" and conv: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" "(\k. integral S (f k)) \ integral S g" proof - have 3: "h absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (h x)) integrable_on S" proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI) fix x assume "x \ S - {}" then show "norm (h x) = h x" by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def) qed auto qed fact have 2: "set_borel_measurable lebesgue S (f k)" for k unfolding set_borel_measurable_def using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) then have 1: "set_borel_measurable lebesgue S g" unfolding set_borel_measurable_def by (rule borel_measurable_LIMSEQ_metric) (use conv in \auto split: split_indicator\) have 4: "AE x in lebesgue. (\i. indicator S x *\<^sub>R f i x) \ indicator S x *\<^sub>R g x" "AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \ indicator S x *\<^sub>R h x" for k using conv le by (auto intro!: always_eventually split: split_indicator) have g: "g absolutely_integrable_on S" using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def by (rule integrable_dominated_convergence) then show "g integrable_on S" by (auto simp: absolutely_integrable_on_def) have "(\k. (LINT x:S|lebesgue. f k x)) \ (LINT x:S|lebesgue. g x)" unfolding set_borel_measurable_def set_lebesgue_integral_def using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def by (rule integral_dominated_convergence) then show "(\k. integral S (f k)) \ integral S g" using g absolutely_integrable_integrable_bound[OF le f h] by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto qed lemma has_integral_dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes "\k. (f k has_integral y k) S" "h integrable_on S" "\k. \x\S. norm (f k x) \ h x" "\x\S. (\k. f k x) \ g x" and x: "y \ x" shows "(g has_integral x) S" proof - have int_f: "\k. (f k) integrable_on S" using assms by (auto simp: integrable_on_def) have "(g has_integral (integral S g)) S" by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f) moreover have "integral S g = x" proof (rule LIMSEQ_unique) show "(\i. integral S (f i)) \ x" using integral_unique[OF assms(1)] x by simp show "(\i. integral S (f i)) \ integral S g" by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f) qed ultimately show ?thesis by simp qed lemma dominated_convergence_integrable_1: fixes f :: "nat \ 'n::euclidean_space \ real" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" proof - have habs: "h absolutely_integrable_on S" using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast let ?f = "\n x. (min (max (- h x) (f n x)) (h x))" have h0: "h x \ 0" if "x \ S" for x using normg that by force have leh: "norm (?f k x) \ h x" if "x \ S" for k x using h0 that by force have limf: "(\k. ?f k x) \ g x" if "x \ S" for x proof - have "\e y. \f y x - g x\ < e \ \min (max (- h x) (f y x)) (h x) - g x\ < e" using h0 [OF that] normg [OF that] by simp then show ?thesis using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono) qed show ?thesis proof (rule dominated_convergence [of ?f S h g]) have "(\x. - h x) absolutely_integrable_on S" using habs unfolding set_integrable_def by auto then show "?f k integrable_on S" for k by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs) qed (use assms leh limf in auto) qed lemma dominated_convergence_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" using f unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k] proof clarify fix b :: "'m" assume fb [rule_format]: "\k. \b\Basis. (\x. f k x \ b) absolutely_integrable_on S" and b: "b \ Basis" show "(\x. g x \ b) integrable_on S" proof (rule dominated_convergence_integrable_1 [OF fb h]) fix x assume "x \ S" show "norm (g x \ b) \ h x" using norm_nth_le \x \ S\ b normg order.trans by blast show "(\k. f k x \ b) \ g x \ b" using \x \ S\ b lim tendsto_componentwise_iff by fastforce qed (use b in auto) qed lemma dominated_convergence_absolutely_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g absolutely_integrable_on S" proof - have "g integrable_on S" by (rule dominated_convergence_integrable [OF assms]) with assms show ?thesis by (blast intro: absolutely_integrable_integrable_bound [where g=h]) qed proposition integral_countable_UN: fixes f :: "real^'m \ real^'n" assumes f: "f absolutely_integrable_on (\(range s))" and s: "\m. s m \ sets lebesgue" shows "\n. f absolutely_integrable_on (\m\n. s m)" and "(\n. integral (\m\n. s m) f) \ integral (\(s ` UNIV)) f" (is "?F \ ?I") proof - show fU: "f absolutely_integrable_on (\m\n. s m)" for n using assms by (blast intro: set_integrable_subset [OF f]) have fint: "f integrable_on (\ (range s))" using absolutely_integrable_on_def f by blast let ?h = "\x. if x \ \(s ` UNIV) then norm(f x) else 0" have "(\n. integral UNIV (\x. if x \ (\m\n. s m) then f x else 0)) \ integral UNIV (\x. if x \ \(s ` UNIV) then f x else 0)" proof (rule dominated_convergence) show "(\x. if x \ (\m\n. s m) then f x else 0) integrable_on UNIV" for n unfolding integrable_restrict_UNIV using fU absolutely_integrable_on_def by blast show "(\x. if x \ \(s ` UNIV) then norm(f x) else 0) integrable_on UNIV" by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV) show "\x. (\n. if x \ (\m\n. s m) then f x else 0) \ (if x \ \(s ` UNIV) then f x else 0)" by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto then show "?F \ ?I" by (simp only: integral_restrict_UNIV) qed subsection \Fundamental Theorem of Calculus for the Lebesgue integral\ text \ For the positive integral we replace continuity with Borel-measurability. \ lemma fixes f :: "real \ real" assumes [measurable]: "f \ borel_measurable borel" assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" shows nn_integral_FTC_Icc: "(\\<^sup>+x. ennreal (f x) * indicator {a .. b} x \lborel) = F b - F a" (is ?nn) and has_bochner_integral_FTC_Icc_nonneg: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_nonneg: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) and integrable_FTC_Icc_nonneg: "integrable lborel (\x. f x * indicator {a .. b} x)" (is ?int) proof - have *: "(\x. f x * indicator {a..b} x) \ borel_measurable borel" "\x. 0 \ f x * indicator {a..b} x" using f(2) by (auto split: split_indicator) have F_mono: "a \ x \ x \ y \ y \ b\ F x \ F y" for x y using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) have "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] intro: has_field_derivative_subset[OF f(1)] \a \ b\) then have i: "((\x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV" unfolding indicator_def if_distrib[where f="\x. a * x" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) then have nn: "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" by (rule nn_integral_has_integral_lborel[OF *]) then show ?has by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \a \ b\) then show ?eq ?int unfolding has_bochner_integral_iff by auto show ?nn by (subst nn[symmetric]) (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator) qed lemma fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" assumes "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" assumes cont: "continuous_on {a .. b} f" shows has_bochner_integral_FTC_Icc: "has_bochner_integral lborel (\x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has) and integral_FTC_Icc: "(\x. indicator {a .. b} x *\<^sub>R f x \lborel) = F b - F a" (is ?eq) proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have int: "integrable lborel ?f" using borel_integrable_compact[OF _ cont] by auto have "(f has_integral F b - F a) {a..b}" using assms(1,2) by (intro fundamental_theorem_of_calculus) auto moreover have "(f has_integral integral\<^sup>L lborel ?f) {a..b}" using has_integral_integral_lborel[OF int] unfolding indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) ultimately show ?eq by (auto dest: has_integral_unique) then show ?has using int by (auto simp: has_bochner_integral_iff) qed lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ DERIV F x :> f x" assumes cont: "\x. a \ x \ x \ b \ isCont f x" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_field_derivative_iff_has_vector_derivative[symmetric] using deriv by (auto intro: DERIV_subset) have 2: "continuous_on {a .. b} f" using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 2] integral_FTC_Icc[OF \a \ b\ 1 2] by (auto simp: mult.commute) qed lemma nn_integral_FTC_atLeast: fixes f :: "real \ real" assumes f_borel: "f \ borel_measurable borel" assumes f: "\x. a \ x \ DERIV F x :> f x" assumes nonneg: "\x. a \ x \ 0 \ f x" assumes lim: "(F \ T) at_top" shows "(\\<^sup>+x. ennreal (f x) * indicator {a ..} x \lborel) = T - F a" proof - let ?f = "\(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x" let ?fR = "\x. ennreal (f x) * indicator {a ..} x" have F_mono: "a \ x \ x \ y \ F x \ F y" for x y using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) then have F_le_T: "a \ x \ F x \ T" for x by (intro tendsto_lowerbound[OF lim]) (auto simp: eventually_at_top_linorder) have "(SUP i. ?f i x) = ?fR x" for x proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) obtain n where "x - a < real n" using reals_Archimedean2[of "x - a"] .. then have "eventually (\n. ?f n x = ?fR x) sequentially" by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) then show "(\n. ?f n x) \ ?fR x" by (rule tendsto_eventually) qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) then have "integral\<^sup>N lborel ?fR = (\\<^sup>+ x. (SUP i. ?f i x) \lborel)" by simp also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \lborel))" proof (rule nn_integral_monotone_convergence_SUP) show "incseq ?f" using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) show "\i. (?f i) \ borel_measurable lborel" using f_borel by auto qed also have "\ = (SUP i. ennreal (F (a + real i) - F a))" by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto also have "\ = T - F a" proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) have "(\x. F (a + real x)) \ T" apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl]) apply (rule filterlim_real_sequentially) done then show "(\n. ennreal (F (a + real n) - F a)) \ ennreal (T - F a)" by (simp add: F_mono F_le_T tendsto_diff) qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono) finally show ?thesis . qed lemma integral_power: "a \ b \ (\x. x^k * indicator {a..b} x \lborel) = (b^Suc k - a^Suc k) / Suc k" proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x^Suc k / Suc k) x :> x^k" by (intro derivative_eq_intros) auto qed (auto simp: field_simps simp del: of_nat_Suc) subsection \Integration by parts\ lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: DERIV_isCont) have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros) by (auto intro!: DERIV_isCont Bochner_Integration.integral_cong split: split_indicator) thus ?thesis using 0 by auto qed lemma integral_by_parts': fixes f g F G::"real \ real" assumes "a \ b" assumes "!!x. a \x \ x\b \ isCont f x" assumes "!!x. a \x \ x\b \ isCont g x" assumes "!!x. DERIV F x :> f x" assumes "!!x. DERIV G x :> g x" shows "(\x. indicator {a .. b} x *\<^sub>R (F x * g x) \lborel) = F b * G b - F a * G a - \x. indicator {a .. b} x *\<^sub>R (f x * G x) \lborel" using integral_by_parts[OF assms] by (simp add: ac_simps) lemma has_bochner_integral_even_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes even: "\x. f (- x) = f x" shows "has_bochner_integral lborel f (2 *\<^sub>R x)" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator even f) with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by (simp add: scaleR_2) qed lemma has_bochner_integral_odd_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes odd: "\x. f (- x) = - f x" shows "has_bochner_integral lborel f 0" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. - indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator odd f) from has_bochner_integral_minus[OF this] have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) (- x)" by simp with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + - x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by simp qed lemma has_integral_0_closure_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (closure S) f" and nonneg_interior: "\x. x \ S \ 0 \ f x" and pos: "0 < emeasure lborel S" and finite: "emeasure lborel S < \" and regular: "emeasure lborel (closure S) = emeasure lborel S" and opn: "open S" assumes int: "(f has_integral 0) (closure S)" assumes x: "x \ closure S" shows "f x = 0" proof - have zero: "emeasure lborel (frontier S) = 0" using finite closure_subset regular unfolding frontier_def by (subst emeasure_Diff) (auto simp: frontier_def interior_open \open S\ ) have nonneg: "0 \ f x" if "x \ closure S" for x using continuous_ge_on_closure[OF f that nonneg_interior] by simp have "0 = integral (closure S) f" by (blast intro: int sym) also note intl = has_integral_integrable[OF int] have af: "f absolutely_integrable_on (closure S)" using nonneg by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f" by (intro set_lebesgue_integral_eq_integral(2)[symmetric]) also have "\ = 0 \ (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" unfolding set_lebesgue_integral_def proof (rule integral_nonneg_eq_0_iff_AE) show "integrable lebesgue (\x. indicat_real (closure S) x *\<^sub>R f x)" by (metis af set_integrable_def) qed (use nonneg in \auto simp: indicator_def\) also have "\ \ (AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by (auto simp: indicator_def) finally have "(AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by simp moreover have "(AE x in lebesgue. x \ - frontier S)" using zero by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"]) ultimately have ae: "AE x \ S in lebesgue. x \ {x \ closure S. f x = 0}" (is ?th) by eventually_elim (use closure_subset in \auto simp: \) have "closed {0::real}" by simp with continuous_on_closed_vimage[OF closed_closure, of S f] f have "closed (f -` {0} \ closure S)" by blast then have "closed {x \ closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute) with \open S\ have "x \ {x \ closure S. f x = 0}" if "x \ S" for x using ae that by (rule mem_closed_if_AE_lebesgue_open) then have "f x = 0" if "x \ S" for x using that by auto from continuous_constant_on_closure[OF f this \x \ closure S\] show "f x = 0" . qed lemma has_integral_0_cbox_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (cbox a b) f" and nonneg_interior: "\x. x \ box a b \ 0 \ f x" assumes int: "(f has_integral 0) (cbox a b)" assumes ne: "box a b \ {}" assumes x: "x \ cbox a b" shows "f x = 0" proof - have "0 < emeasure lborel (box a b)" using ne x unfolding emeasure_lborel_box_eq by (force intro!: prod_pos simp: mem_box algebra_simps) then show ?thesis using assms by (intro has_integral_0_closure_imp_0[of "box a b" f x]) (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box) qed subsection\Various common equivalent forms of function measurability\ lemma indicator_sum_eq: fixes m::real and f :: "'a \ real" assumes "\m\ \ 2 ^ (2*n)" "m/2^n \ f x" "f x < (m+1)/2^n" "m \ \" shows "(\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) = m/2^n" (is "sum ?f ?S = _") proof - have "sum ?f ?S = sum (\k. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {m}" proof (rule comm_monoid_add_class.sum.mono_neutral_right) show "finite ?S" by (rule finite_abs_int_segment) show "{m} \ {k \ \. \k\ \ 2 ^ (2*n)}" using assms by auto show "\i\{k \ \. \k\ \ 2 ^ (2*n)} - {m}. ?f i = 0" using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps) qed also have "\ = m/2^n" using assms by (auto simp: indicator_def not_less) finally show ?thesis . qed lemma measurable_on_sf_limit_lemma1: fixes f :: "'a::euclidean_space \ real" assumes "\a b. {x \ S. a \ f x \ f x < b} \ sets (lebesgue_on S)" obtains g where "\n. g n \ borel_measurable (lebesgue_on S)" "\n. finite(range (g n))" "\x. (\n. g n x) \ f x" proof show "(\x. sum (\k::real. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {k \ \. \k\ \ 2 ^ (2*n)}) \ borel_measurable (lebesgue_on S)" (is "?g \ _") for n proof - have "\k. \k \ \; \k\ \ 2 ^ (2*n)\ \ Measurable.pred (lebesgue_on S) (\x. k / (2^n) \ f x \ f x < (k+1) / (2^n))" using assms by (force simp: pred_def space_restrict_space) then show ?thesis by (simp add: field_class.field_divide_inverse) qed show "finite (range (?g n))" for n proof - have "range (?g n) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof clarify fix x show "?g n x \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof (cases "\k::real. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ (f x) \ (f x) < (k+1)/2^n") case True then show ?thesis apply clarify by (subst indicator_sum_eq) auto next case False then have "?g n x = 0" by auto then show ?thesis by force qed qed moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (simp add: finite_abs_int_segment) ultimately show ?thesis using finite_subset by blast qed show "(\n. ?g n x) \ f x" for x proof (rule LIMSEQ_I) fix e::real assume "e > 0" obtain N1 where N1: "\f x\ < 2 ^ N1" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by force have "norm (?g n x - f x) < e" if n: "n \ max N1 N2" for n proof - define m where "m \ floor(2^n * (f x))" have "1 \ \2^n\ * e" using n N2 \e > 0\ less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps) then have *: "\x \ y; y < x + 1\ \ abs(x - y) < \2^n\ * e" for x y::real by linarith have "\2^n\ * \m/2^n - f x\ = \2^n * (m/2^n - f x)\" by (simp add: abs_mult) also have "\ = \real_of_int \2^n * f x\ - f x * 2^n\" by (simp add: algebra_simps m_def) also have "\ < \2^n\ * e" by (rule *; simp add: mult.commute) finally have "\2^n\ * \m/2^n - f x\ < \2^n\ * e" . then have me: "\m/2^n - f x\ < e" by simp have "\real_of_int m\ \ 2 ^ (2*n)" proof (cases "f x < 0") case True then have "-\f x\ \ \(2::real) ^ N1\" using N1 le_floor_iff minus_le_iff by fastforce with n True have "\real_of_int\f x\\ \ 2 ^ N1" by linarith also have "\ \ 2^n" using n by (simp add: m_def) finally have "\real_of_int \f x\\ * 2^n \ 2^n * 2^n" by simp moreover have "\real_of_int \2^n * f x\\ \ \real_of_int \f x\\ * 2^n" proof - have "\real_of_int \2^n * f x\\ = - (real_of_int \2^n * f x\)" using True by (simp add: abs_if mult_less_0_iff) also have "\ \ - (real_of_int (\(2::real) ^ n\ * \f x\))" using le_mult_floor_Ints [of "(2::real)^n"] by simp also have "\ \ \real_of_int \f x\\ * 2^n" using True by simp finally show ?thesis . qed ultimately show ?thesis by (metis (no_types, hide_lams) m_def order_trans power2_eq_square power_even_eq) next case False with n N1 have "f x \ 2^n" by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing) moreover have "0 \ m" using False m_def by force ultimately show ?thesis by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult real_mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power) qed then have "?g n x = m/2^n" by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith) then have "norm (?g n x - f x) = norm (m/2^n - f x)" by simp also have "\ < e" by (simp add: me) finally show ?thesis . qed then show "\no. \n\no. norm (?g n x - f x) < e" by blast qed qed lemma borel_measurable_simple_function_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x))" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x)" if f: "\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S)" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite(image (g n) UNIV)) \ (\x. ((\n. g n x) \ f x \ i))" if "i \ Basis" for i proof (rule measurable_on_sf_limit_lemma1 [of S "\x. f x \ i"]) show "{x \ S. a \ f x \ i \ f x \ i < b} \ sets (lebesgue_on S)" for a b proof - have "{x \ S. a \ f x \ i \ f x \ i < b} = {x \ S. f x \ i < b} - {x \ S. a > f x \ i}" by auto also have "\ \ sets (lebesgue_on S)" using f that by blast finally show ?thesis . qed qed blast then obtain g where g: "\i n. i \ Basis \ g i n \ borel_measurable (lebesgue_on S)" "\i n. i \ Basis \ finite(range (g i n))" "\i x. i \ Basis \ ((\n. g i n x) \ f x \ i)" by metis show ?thesis proof (intro conjI allI exI) show "(\x. \i\Basis. g i n x *\<^sub>R i) \ borel_measurable (lebesgue_on S)" for n by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g) show "finite (range (\x. \i\Basis. g i n x *\<^sub>R i))" for n proof - have "range (\x. \i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` PiE Basis (\i. range (g i n))" proof clarify fix x show "(\i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` (\\<^sub>E i\Basis. range (g i n))" by (rule_tac x="\i\Basis. g i n x" in image_eqI) auto qed moreover have "finite(PiE Basis (\i. range (g i n)))" by (simp add: g finite_PiE) ultimately show ?thesis by (metis (mono_tags, lifting) finite_surj) qed show "(\n. \i\Basis. g i n x *\<^sub>R i) \ f x" for x proof - have "(\n. \i\Basis. g i n x *\<^sub>R i) \ (\i\Basis. (f x \ i) *\<^sub>R i)" by (auto intro!: tendsto_sum tendsto_scaleR g) moreover have "(\i\Basis. (f x \ i) *\<^sub>R i) = f x" using euclidean_representation by blast ultimately show ?thesis by metis qed qed qed moreover have "f \ borel_measurable (lebesgue_on S)" if meas_g: "\n. g n \ borel_measurable (lebesgue_on S)" and fin: "\n. finite (range (g n))" and to_f: "\x. (\n. g n x) \ f x" for g by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f]) ultimately show ?thesis using borel_measurable_vimage_halfspace_component_lt by blast qed subsection \Lebesgue sets and continuous images\ proposition lebesgue_regular_inner: assumes "S \ sets lebesgue" obtains K C where "negligible K" "\n::nat. compact(C n)" "S = (\n. C n) \ K" proof - have "\T. closed T \ T \ S \ (S - T) \ lmeasurable \ emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n using sets_lebesgue_inner_closed assms by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) then obtain C where clo: "\n. closed (C n)" and subS: "\n. C n \ S" and mea: "\n. (S - C n) \ lmeasurable" and less: "\n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)" by metis have "\F. (\n::nat. compact(F n)) \ (\n. F n) = C m" for m::nat by (metis clo closed_Union_compact_subsets) then obtain D :: "[nat,nat] \ 'a set" where D: "\m n. compact(D m n)" "\m. (\n. D m n) = C m" by metis let ?C = "from_nat_into (\m. range (D m))" have "countable (\m. range (D m))" by blast have "range (from_nat_into (\m. range (D m))) = (\m. range (D m))" using range_from_nat_into by simp then have CD: "\m n. ?C k = D m n" for k by (metis (mono_tags, lifting) UN_iff rangeE range_eqI) show thesis proof show "negligible (S - (\n. C n))" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" then obtain n where n: "(1/2)^n < e" using real_arch_pow_inv [of e "1/2"] by auto show "\T. S - (\n. C n) \ T \ T \ lmeasurable \ measure lebesgue T \ e" proof (intro exI conjI) show "S - (\n. C n) \ S - C n" by blast show "S - C n \ lmeasurable" by (simp add: mea) show "measure lebesgue (S - C n) \ e" using less [of n] n by (simp add: emeasure_eq_measure2 less_le mea) qed qed show "compact (?C n)" for n using CD D by metis show "S = (\n. ?C n) \ (S - (\n. C n))" (is "_ = ?rhs") proof show "S \ ?rhs" using D by fastforce show "?rhs \ S" using subS D CD by auto (metis Sup_upper range_eqI subsetCE) qed qed qed lemma sets_lebesgue_continuous_image: assumes T: "T \ sets lebesgue" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" and "T \ S" shows "f ` T \ sets lebesgue" proof - obtain K C where "negligible K" and com: "\n::nat. compact(C n)" and Teq: "T = (\n. C n) \ K" using lebesgue_regular_inner [OF T] by metis then have comf: "\n::nat. compact(f ` C n)" by (metis Un_subset_iff Union_upper \T \ S\ compact_continuous_image contf continuous_on_subset rangeI) have "((\n. f ` C n) \ f ` K) \ sets lebesgue" proof (rule sets.Un) have "K \ S" using Teq \T \ S\ by blast show "(\n. f ` C n) \ sets lebesgue" proof (rule sets.countable_Union) show "range (\n. f ` C n) \ sets lebesgue" using borel_compact comf by (auto simp: borel_compact) qed auto show "f ` K \ sets lebesgue" by (simp add: \K \ S\ \negligible K\ negim negligible_imp_sets) qed then show ?thesis by (simp add: Teq image_Un image_Union) qed lemma differentiable_image_in_sets_lebesgue: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f`S \ sets lebesgue" proof (rule sets_lebesgue_continuous_image [OF S]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed auto lemma sets_lebesgue_on_continuous_image: assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" shows "f ` X \ sets (lebesgue_on (f ` S))" proof - have "X \ S" by (metis S X sets.Int_space_eq2 sets_restrict_space_iff) moreover have "f ` S \ sets lebesgue" using S contf negim sets_lebesgue_continuous_image by blast moreover have "f ` X \ sets lebesgue" by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2) ultimately show ?thesis by (auto simp: sets_restrict_space_iff) qed lemma differentiable_image_in_sets_lebesgue_on: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f ` X \ sets (lebesgue_on (f`S))" proof (rule sets_lebesgue_on_continuous_image [OF S X]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed subsection \Affine lemmas\ lemma borel_measurable_affine: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f \ borel_measurable lebesgue" and "c \ 0" shows "(\x. f(t + c *\<^sub>R x)) \ borel_measurable lebesgue" proof - { fix a b have "{x. f x \ cbox a b} \ sets lebesgue" using f cbox_borel lebesgue_measurable_vimage_borel by blast then have "(\x. (x - t) /\<^sub>R c) ` {x. f x \ cbox a b} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "(\x. (x - t) /\<^sub>R c) differentiable_on {x. f x \ cbox a b}" unfolding differentiable_on_def differentiable_def by (rule \c \ 0\ derivative_eq_intros strip exI | simp)+ qed auto moreover have "{x. f(t + c *\<^sub>R x) \ cbox a b} = (\x. (x-t) /\<^sub>R c) ` {x. f x \ cbox a b}" using \c \ 0\ by (auto simp: image_def) ultimately have "{x. f(t + c *\<^sub>R x) \ cbox a b} \ sets lebesgue" by (auto simp: borel_measurable_vimage_closed_interval) } then show ?thesis by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) qed lemma lebesgue_integrable_real_affine: fixes f :: "real \ 'a :: euclidean_space" assumes f: "integrable lebesgue f" and "c \ 0" shows "integrable lebesgue (\x. f(t + c * x))" proof - have "(\x. norm (f x)) \ borel_measurable lebesgue" by (simp add: borel_measurable_integrable f) then show ?thesis using assms borel_measurable_affine [of f c] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) qed lemma lebesgue_integrable_real_affine_iff: fixes f :: "real \ 'a :: euclidean_space" shows "c \ 0 \ integrable lebesgue (\x. f(t + c * x)) \ integrable lebesgue f" using lebesgue_integrable_real_affine[of f c t] lebesgue_integrable_real_affine[of "\x. f(t + c * x)" "1/c" "-t/c"] by (auto simp: field_simps) lemma\<^marker>\tag important\ lebesgue_integral_real_affine: fixes f :: "real \ 'a :: euclidean_space" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lebesgue) = \c\ *\<^sub>R (\x. f(t + c * x) \lebesgue)" proof cases have "(\x. t + c * x) \ lebesgue \\<^sub>M lebesgue" using lebesgue_affine_measurable[where c= "\x::real. c"] \c \ 0\ by simp moreover assume "integrable lebesgue f" ultimately show ?thesis by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) next assume "\ integrable lebesgue f" with c show ?thesis by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) qed lemma has_bochner_integral_lebesgue_real_affine_iff: fixes i :: "'a :: euclidean_space" shows "c \ 0 \ has_bochner_integral lebesgue f i \ has_bochner_integral lebesgue (\x. f(t + c * x)) (i /\<^sub>R \c\)" unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma has_bochner_integral_reflect_real_lemma[intro]: fixes f :: "real \ 'a::euclidean_space" assumes "has_bochner_integral (lebesgue_on {a..b}) f i" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i" proof - have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x by (auto simp: indicator_def) have i: "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" using assms by (auto simp: has_bochner_integral_restrict_space) then have "has_bochner_integral lebesgue (\x. indicator {-b..-a} x *\<^sub>R f(-x)) i" using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\x. indicator {a..b} x *\<^sub>R f x)" i 0] by (auto simp: eq) then show ?thesis by (auto simp: has_bochner_integral_restrict_space) qed lemma has_bochner_integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i \ has_bochner_integral (lebesgue_on {a..b}) f i" by (auto simp: dest: has_bochner_integral_reflect_real_lemma) lemma integrable_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integrable (lebesgue_on {-b..-a}) (\x. f(-x)) \ integrable (lebesgue_on {a..b}) f" by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) lemma integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f" using has_bochner_integral_reflect_real [of b a f] by (metis has_bochner_integral_iff not_integrable_integral_eq) subsection\More results on integrability\ lemma integrable_on_all_intervals_UNIV: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. f integrable_on cbox a b" and normf: "\x. norm(f x) \ g x" and g: "g integrable_on UNIV" shows "f integrable_on UNIV" proof - have intg: "(\a b. g integrable_on cbox a b)" and gle_e: "\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ \integral (cbox a b) g - integral (cbox c d) g\ < e" using g by (auto simp: integrable_alt_subset [of _ UNIV] intf) have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \ \integral (cbox a b) g - integral (cbox c d) g\" if "cbox a b \ cbox c d" for a b c d proof - have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)" using intf that by (simp add: norm_minus_commute integral_setdiff) also have "\ \ integral (cbox c d - cbox a b) g" proof (rule integral_norm_bound_integral [OF _ _ normf]) show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b" by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+ qed also have "\ = integral (cbox c d) g - integral (cbox a b) g" using intg that by (simp add: integral_setdiff) also have "\ \ \integral (cbox a b) g - integral (cbox c d) g\" by simp finally show ?thesis . qed show ?thesis using gle_e apply (simp add: integrable_alt_subset [of _ UNIV] intf) apply (erule imp_forward all_forward ex_forward asm_rl)+ by (meson not_less order_trans le) qed lemma integrable_on_all_intervals_integrable_bound: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b" and normf: "\x. x \ S \ norm(f x) \ g x" and g: "g integrable_on S" shows "f integrable_on S" using integrable_on_all_intervals_UNIV [OF intf, of "(\x. if x \ S then g x else 0)"] by (simp add: g integrable_restrict_UNIV normf) lemma measurable_bounded_lemma: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable lebesgue" and g: "g integrable_on cbox a b" and normf: "\x. x \ cbox a b \ norm(f x) \ g x" shows "f integrable_on cbox a b" proof - have "g absolutely_integrable_on cbox a b" by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf) then have "integrable (lebesgue_on (cbox a b)) g" by (simp add: integrable_restrict_space set_integrable_def) then have "integrable (lebesgue_on (cbox a b)) f" proof (rule Bochner_Integration.integrable_bound) show "AE x in lebesgue_on (cbox a b). norm (f x) \ norm (g x)" by (rule AE_I2) (auto intro: normf order_trans) qed (simp add: f measurable_restrict_space1) then show ?thesis by (simp add: integrable_on_lebesgue_on) qed proposition measurable_bounded_by_integrable_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "g integrable_on S" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "f integrable_on S" proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g]) show "(\x. if x \ S then f x else 0) integrable_on cbox a b" for a b proof (rule measurable_bounded_lemma) show "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" by (simp add: S borel_measurable_if f) show "(\x. if x \ S then g x else 0) integrable_on cbox a b" by (simp add: g integrable_altD(1)) show "norm (if x \ S then f x else 0) \ (if x \ S then g x else 0)" for x using normf by simp qed qed lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" proof - have "f absolutely_integrable_on S" by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) then show ?thesis by (simp add: S integrable_restrict_space set_integrable_def) qed lemma measurable_bounded_by_integrable_imp_integrable_real: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable (lebesgue_on S)" "g integrable_on S" "\x. x \ S \ abs(f x) \ g x" "S \ sets lebesgue" shows "f integrable_on S" using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp subsection\ Relation between Borel measurability and integrability.\ lemma integrable_imp_measurable_weak: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2) lemma integrable_imp_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" proof - have "(UNIV::'a set) \ sets lborel" by simp then show ?thesis by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl) qed lemma integrable_iff_integrable_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "(\\<^sup>+ x. ennreal (norm (f x)) \lebesgue_on S) < \" shows "integrable (lebesgue_on S) f \ f integrable_on S" using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast lemma absolutely_integrable_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (norm \ f)" (is "?lhs = ?rhs") proof assume L: ?lhs then have "f \ borel_measurable (lebesgue_on S)" by (simp add: absolutely_integrable_on_def integrable_imp_measurable) then show ?rhs using assms set_integrable_norm [of lebesgue S f] L by (simp add: integrable_restrict_space set_integrable_def) next assume ?rhs then show ?lhs using assms integrable_on_lebesgue_on by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) qed lemma absolutely_integrable_measurable_real: fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. \f x\)" by (simp add: absolutely_integrable_measurable assms o_def) lemma absolutely_integrable_measurable_real': fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ (\x. \f x\) integrable_on S" using assms apply (auto simp: absolutely_integrable_measurable integrable_on_lebesgue_on) apply (simp add: integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_lebesgue_integrable) using abs_absolutely_integrableI_1 absolutely_integrable_measurable measurable_bounded_by_integrable_imp_integrable_real by blast lemma absolutely_integrable_imp_borel_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using absolutely_integrable_measurable assms by blast lemma measurable_bounded_by_integrable_imp_absolutely_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and "g integrable_on S" and "\x. x \ S \ norm(f x) \ (g x)" shows "f absolutely_integrable_on S" using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast proposition negligible_differentiable_vimage: fixes f :: "'a \ 'a::euclidean_space" assumes "negligible T" and f': "\x. x \ S \ inj(f' x)" and derf: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "negligible {x \ S. f x \ T}" proof - define U where "U \ \n::nat. {x \ S. \y. y \ S \ norm(y - x) < 1/n \ norm(y - x) \ n * norm(f y - f x)}" have "negligible {x \ U n. f x \ T}" if "n > 0" for n proof (subst locally_negligible_alt, clarify) fix a assume a: "a \ U n" and fa: "f a \ T" define V where "V \ {x. x \ U n \ f x \ T} \ ball a (1 / n / 2)" show "\V. openin (top_of_set {x \ U n. f x \ T}) V \ a \ V \ negligible V" proof (intro exI conjI) have noxy: "norm(x - y) \ n * norm(f x - f y)" if "x \ V" "y \ V" for x y using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm by (meson norm_triangle_half_r) then have "inj_on f V" by (force simp: inj_on_def) then obtain g where g: "\x. x \ V \ g(f x) = x" by (metis inv_into_f_f) have "\T' B. open T' \ f x \ T' \ (\y\f ` V \ T \ T'. norm (g y - g (f x)) \ B * norm (y - f x))" if "f x \ T" "x \ V" for x apply (rule_tac x = "ball (f x) 1" in exI) using that noxy by (force simp: g) then have "negligible (g ` (f ` V \ T))" by (force simp: \negligible T\ negligible_Int intro!: negligible_locally_Lipschitz_image) moreover have "V \ g ` (f ` V \ T)" by (force simp: g image_iff V_def) ultimately show "negligible V" by (rule negligible_subset) qed (use a fa V_def that in auto) qed with negligible_countable_Union have "negligible (\n \ {0<..}. {x. x \ U n \ f x \ T})" by auto moreover have "{x \ S. f x \ T} \ (\n \ {0<..}. {x. x \ U n \ f x \ T})" proof clarsimp fix x assume "x \ S" and "f x \ T" then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)" using assms by metis moreover have "linear(f' x)" and eps: "\\. \ > 0 \ \\>0. \y\S. norm (y - x) < \ \ norm (f y - f x - f' x (y - x)) \ \ * norm (y - x)" using der by (auto simp: has_derivative_within_alt linear_linear) ultimately obtain g where "linear g" and g: "g \ f' x = id" using linear_injective_left_inverse by metis then obtain B where "B > 0" and B: "\z. B * norm z \ norm(f' x z)" using linear_invertible_bounded_below_pos \linear (f' x)\ by blast then obtain i where "i \ 0" and i: "1 / real i < B" by (metis (full_types) inverse_eq_divide real_arch_invD) then obtain \ where "\ > 0" and \: "\y. \y\S; norm (y - x) < \\ \ norm (f y - f x - f' x (y - x)) \ (B - 1 / real i) * norm (y - x)" using eps [of "B - 1/i"] by auto then obtain j where "j \ 0" and j: "inverse (real j) < \" using real_arch_inverse by blast have "norm (y - x)/(max i j) \ norm (f y - f x)" if "y \ S" and less: "norm (y - x) < 1 / (max i j)" for y proof - have "1 / real (max i j) < \" using j \j \ 0\ \0 < \\ by (auto simp: field_split_simps max_mult_distrib_left of_nat_max) then have "norm (y - x) < \" using less by linarith with \ \y \ S\ have le: "norm (f y - f x - f' x (y - x)) \ B * norm (y - x) - norm (y - x)/i" by (auto simp: algebra_simps) have *: "\norm(f - f' - y) \ b - c; b \ norm y; d \ c\ \ d \ norm(f - f')" for b c d and y f f'::'a using norm_triangle_ineq3 [of "f - f'" y] by simp show ?thesis apply (rule * [OF le B]) using \i \ 0\ \j \ 0\ by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) qed with \x \ S\ \i \ 0\ \j \ 0\ show "\n\{0<..}. x \ U n" by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj) qed ultimately show ?thesis by (rule negligible_subset) qed lemma absolutely_integrable_Un: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T" shows "f absolutely_integrable_on (S \ T)" proof - have [simp]: "{x. (if x \ A then f x else 0) \ 0} = {x \ A. f x \ 0}" for A by auto let ?ST = "{x \ S. f x \ 0} \ {x \ T. f x \ 0}" have "?ST \ sets lebesgue" proof (rule Sigma_Algebra.sets.Int) have "f integrable_on S" using S absolutely_integrable_on_def by blast then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ S. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto next have "f integrable_on T" using T absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ T. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto qed then have "f absolutely_integrable_on ?ST" by (rule set_integrable_subset [OF S]) auto then have Int: "(\x. if x \ ?ST then f x else 0) absolutely_integrable_on UNIV" using absolutely_integrable_restrict_UNIV by blast have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" "(\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" using S T absolutely_integrable_restrict_UNIV by blast+ then have "(\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)) absolutely_integrable_on UNIV" by (rule set_integral_add) then have "(\x. ((if x \ S then f x else 0) + (if x \ T then f x else 0)) - (if x \ ?ST then f x else 0)) absolutely_integrable_on UNIV" using Int by (rule set_integral_diff) then have "(\x. if x \ S \ T then f x else 0) absolutely_integrable_on UNIV" by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible) then show ?thesis unfolding absolutely_integrable_restrict_UNIV . qed lemma absolutely_integrable_on_combine: fixes f :: "real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..c}" and "f absolutely_integrable_on {c..b}" and "a \ c" and "c \ b" shows "f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4)) lemma uniform_limit_set_lebesgue_integral_at_top: fixes f :: "'a \ real \ 'b::{banach, second_countable_topology}" and g :: "real \ real" assumes bound: "\x y. x \ A \ y \ a \ norm (f x y) \ g y" assumes integrable: "set_integrable M {a..} g" assumes measurable: "\x. x \ A \ set_borel_measurable M {a..} (f x)" assumes "sets borel \ sets M" shows "uniform_limit A (\b x. LINT y:{a..b}|M. f x y) (\x. LINT y:{a..}|M. f x y) at_top" proof (cases "A = {}") case False then obtain x where x: "x \ A" by auto have g_nonneg: "g y \ 0" if "y \ a" for y proof - have "0 \ norm (f x y)" by simp also have "\ \ g y" using bound[OF x that] by simp finally show ?thesis . qed have integrable': "set_integrable M {a..} (\y. f x y)" if "x \ A" for x unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) show "integrable M (\x. indicator {a..} x * g x)" using integrable by (simp add: set_integrable_def) show "(\y. indicat_real {a..} y *\<^sub>R f x y) \ borel_measurable M" using measurable[OF that] by (simp add: set_borel_measurable_def) show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \ norm (indicat_real {a..} y * g y)" using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) qed show ?thesis proof (rule uniform_limitI) fix e :: real assume e: "e > 0" have sets [intro]: "A \ sets M" if "A \ sets borel" for A using that assms by blast have "((\b. LINT y:{a..b}|M. g y) \ (LINT y:{a..}|M. g y)) at_top" by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto with e obtain b0 :: real where b0: "\b\b0. \(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) define b where "b = max a b0" have "a \ b" by (simp add: b_def) from b0 have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: b_def) also have "{a..} = {a..b} \ {b<..}" by (auto simp: b_def) also have "\(LINT y:\|M. g y) - (LINT y:{a..b}|M. g y)\ = \(LINT y:{b<..}|M. g y)\" using \a \ b\ by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "\(LINT y:{b<..}|M. g y)\ = (LINT y:{b<..}|M. g y)" by simp finally have less: "(LINT y:{b<..}|M. g y) < e" . have "eventually (\b. b \ b0) at_top" by (rule eventually_ge_at_top) moreover have "eventually (\b. b \ a) at_top" by (rule eventually_ge_at_top) ultimately show "eventually (\b. \x\A. dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top" proof eventually_elim case (elim b) show ?case proof fix x assume x: "x \ A" have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) = norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))" by (simp add: dist_norm norm_minus_commute) also have "{a..} = {a..b} \ {b<..}" using elim by auto also have "(LINT y:\|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)" using elim x by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) also have "norm \ \ (LINT y:{b<..}|M. norm (f x y))" using elim x by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto also have "\ \ (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] set_integrable_subset[OF integrable]) auto also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "(LINT y:{b<..}|M. g y) = \(LINT y:{b<..}|M. g y)\" by simp also have "\ = \(LINT y:{a..b} \ {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\" using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "{a..b} \ {b<..} = {a..}" using elim by auto also have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" using b0 elim by blast finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" . qed qed qed qed auto subsubsection\Differentiability of inverse function (most basic form)\ proposition has_derivative_inverse_within: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes der_f: "(f has_derivative f') (at a within S)" and cont_g: "continuous (at (f a) within f ` S) g" and "a \ S" "linear g'" and id: "g' \ f' = id" and gf: "\x. x \ S \ g(f x) = x" shows "(g has_derivative g') (at (f a) within f ` S)" proof - have [simp]: "g' (f' x) = x" for x by (simp add: local.id pointfree_idE) have "bounded_linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - a) < d \ norm (f y - f a - f' (y - a)) \ e * norm (y - a)" using der_f by (auto simp: has_derivative_within_alt) obtain C where "C > 0" and C: "\x. norm (g' x) \ C * norm x" using linear_bounded_pos [OF \linear g'\] by metis obtain B k where "B > 0" "k > 0" and Bk: "\x. \x \ S; norm(f x - f a) < k\ \ norm(x - a) \ B * norm(f x - f a)" proof - obtain B where "B > 0" and B: "\x. B * norm x \ norm (f' x)" using linear_inj_bounded_below_pos [of f'] \linear g'\ id der_f has_derivative_linear linear_invertible_bounded_below_pos by blast then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ B / 2 * norm (y - a)" using f' [of "B/2"] by auto then obtain e where "e > 0" and e: "\x. \x \ S; norm (f x - f a) < e\ \ norm (g (f x) - g (f a)) < d" using cont_g by (auto simp: continuous_within_eps_delta dist_norm) show thesis proof show "2/B > 0" using \B > 0\ by simp show "norm (x - a) \ 2 / B * norm (f x - f a)" if "x \ S" "norm (f x - f a) < e" for x proof - have xa: "norm (x - a) < d" using e [OF that] gf by (simp add: \a \ S\ that) have *: "\norm(y - f') \ B / 2 * norm x; B * norm x \ norm f'\ \ norm y \ B / 2 * norm x" for y f'::'b and x::'a using norm_triangle_ineq3 [of y f'] by linarith show ?thesis using * [OF d [OF \x \ S\ xa] B] \B > 0\ by (simp add: field_simps) qed qed (use \e > 0\ in auto) qed show ?thesis unfolding has_derivative_within_alt proof (intro conjI impI allI) show "bounded_linear g'" using \linear g'\ by (simp add: linear_linear) next fix e :: "real" assume "e > 0" then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ e / (B * C) * norm (y - a)" using f' [of "e / (B * C)"] \B > 0\ \C > 0\ by auto have "norm (x - a - g' (f x - f a)) \ e * norm (f x - f a)" if "x \ S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x proof - have "norm (x - a) \ B * norm(f x - f a)" using Bk lt_k \x \ S\ by blast also have "\ < d" by (metis \0 < B\ lt_dB mult.commute pos_less_divide_eq) finally have lt_d: "norm (x - a) < d" . have "norm (x - a - g' (f x - f a)) \ norm(g'(f x - f a - (f' (x - a))))" by (simp add: linear_diff [OF \linear g'\] norm_minus_commute) also have "\ \ C * norm (f x - f a - f' (x - a))" using C by blast also have "\ \ e * norm (f x - f a)" proof - have "norm (f x - f a - f' (x - a)) \ e / (B * C) * norm (x - a)" using d [OF \x \ S\ lt_d] . also have "\ \ (norm (f x - f a) * e) / C" using \B > 0\ \C > 0\ \e > 0\ by (simp add: field_simps Bk lt_k \x \ S\) finally show ?thesis using \C > 0\ by (simp add: field_simps) qed finally show ?thesis . qed then show "\d>0. \y\f ` S. norm (y - f a) < d \ norm (g y - g (f a) - g' (y - f a)) \ e * norm (y - f a)" apply (rule_tac x="min k (d / B)" in exI) using \k > 0\ \B > 0\ \d > 0\ \a \ S\ by (auto simp: gf) qed qed end diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy @@ -1,7627 +1,7665 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) Huge cleanup by LCP *) section \Henstock-Kurzweil Gauge Integration in Many Dimensions\ theory Henstock_Kurzweil_Integration imports Lebesgue_Measure Tagged_Division begin lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ \ norm(y-x) \ e" using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] by (simp add: add_diff_add) lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" by auto lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" by auto lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" by blast lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" by blast (* END MOVE *) subsection \Content (length, area, volume...) of an interval\ abbreviation content :: "'a::euclidean_space set \ real" where "content s \ measure lborel s" lemma content_cbox_cases: "content (cbox a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_cbox_eq inner_diff) lemma content_cbox: "\i\Basis. a\i \ b\i \ content (cbox a b) = (\i\Basis. b\i - a\i)" unfolding content_cbox_cases by simp lemma content_cbox': "cbox a b \ {} \ content (cbox a b) = (\i\Basis. b\i - a\i)" by (simp add: box_ne_empty inner_diff) lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \i\Basis. b\i - a\i)" by (simp add: content_cbox') lemma content_cbox_cart: "cbox a b \ {} \ content(cbox a b) = prod (\i. b$i - a$i) UNIV" by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint) lemma content_cbox_if_cart: "content(cbox a b) = (if cbox a b = {} then 0 else prod (\i. b$i - a$i) UNIV)" by (simp add: content_cbox_cart) lemma content_division_of: assumes "K \ \" "\ division_of S" shows "content K = (\i \ Basis. interval_upperbound K \ i - interval_lowerbound K \ i)" proof - obtain a b where "K = cbox a b" using cbox_division_memE assms by metis then show ?thesis using assms by (force simp: division_of_def content_cbox') qed lemma content_real: "a \ b \ content {a..b} = b - a" by simp lemma abs_eq_content: "\y - x\ = (if x\y then content {x..y} else content {y..x})" by (auto simp: content_real) lemma content_singleton: "content {a} = 0" by simp lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" by simp lemma content_pos_le [iff]: "0 \ content X" by simp corollary\<^marker>\tag unimportant\ content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) lemma content_eq_0: "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_empty [simp]: "content {} = 0" by simp lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" by (simp add: content_real) lemma content_subset: "cbox a b \ cbox c d \ content (cbox a b) \ content (cbox c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" unfolding measure_lborel_cbox_eq Basis_prod_def apply (subst prod.union_disjoint) apply (auto simp: bex_Un ball_Un) apply (subst (1 2) prod.reindex_nontrivial) apply auto done lemma content_cbox_pair_eq0_D: "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" by (simp add: content_Pair) lemma content_cbox_plus: fixes x :: "'a::euclidean_space" shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" by (simp add: algebra_simps content_cbox_if box_eq_empty) lemma content_0_subset: "content(cbox a b) = 0 \ s \ cbox a b \ content s = 0" using emeasure_mono[of s "cbox a b" lborel] by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) +lemma content_ball_pos: + assumes "r > 0" + shows "content (ball c r) > 0" +proof - + from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" + by auto + from ab have "0 < content (box a b)" + by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) + have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" + using ab by (intro emeasure_mono) auto + also have "emeasure lborel (box a b) = ennreal (content (box a b))" + using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto + also have "emeasure lborel (ball c r) = ennreal (content (ball c r))" + using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto + finally show ?thesis + using \content (box a b) > 0\ by simp +qed + +lemma content_cball_pos: + assumes "r > 0" + shows "content (cball c r) > 0" +proof - + from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" + by auto + from ab have "0 < content (box a b)" + by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) + have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" + using ab by (intro emeasure_mono) auto + also have "\ \ emeasure lborel (cball c r)" + by (intro emeasure_mono) auto + also have "emeasure lborel (box a b) = ennreal (content (box a b))" + using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto + also have "emeasure lborel (cball c r) = ennreal (content (cball c r))" + using emeasure_lborel_cball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto + finally show ?thesis + using \content (box a b) > 0\ by simp +qed + lemma content_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" \ \Prove using measure theory\ proof (cases "\i\Basis. a \ i \ b \ i") case True have 1: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" by (simp add: if_distrib prod.delta_remove assms) note simps = interval_split[OF assms] content_cbox_cases have 2: "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove) have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" by (auto simp add: field_simps) moreover have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: prod.cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using True assms by auto ultimately show ?thesis using assms unfolding simps ** 1[of "\i x. b\i - x"] 1[of "\i x. x - a\i"] 2 by auto next case False then have "cbox a b = {}" unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_of_content_0: assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \ d" shows "content K = 0" unfolding forall_in_division[OF assms(2)] by (meson assms content_0_subset division_of_def) lemma sum_content_null: assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" using surj_pair[of y] by blast then obtain c d where k: "K = cbox c d" "K \ cbox a b" by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) have "(\(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using assms(1) content_0_subset k(2) by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed global_interpretation sum_content: operative plus 0 content rewrites "comm_monoid_set.F plus 0 = sum" proof - interpret operative plus 0 content by standard (auto simp add: content_split [symmetric] content_eq_0_interior) show "operative plus 0 content" by standard show "comm_monoid_set.F plus 0 = sum" by (simp add: sum_def) qed lemma additive_content_division: "d division_of (cbox a b) \ sum content d = content (cbox a b)" by (fact sum_content.division) lemma additive_content_tagged_division: "d tagged_division_of (cbox a b) \ sum (\(x,l). content l) d = content (cbox a b)" by (fact sum_content.tagged_division) lemma subadditive_content_division: assumes "\ division_of S" "S \ cbox a b" shows "sum content \ \ content(cbox a b)" proof - have "\ division_of \\" "\\ \ cbox a b" using assms by auto then obtain \' where "\ \ \'" "\' division_of cbox a b" using partial_division_extend_interval by metis then have "sum content \ \ sum content \'" using sum_mono2 by blast also have "... \ content(cbox a b)" by (simp add: \\' division_of cbox a b\ additive_content_division less_eq_real_def) finally show ?thesis . qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma interval_bounds_nz_content [simp]: assumes "content (cbox a b) \ 0" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" by (metis assms content_empty interval_bounds')+ subsection \Gauge integral\ text \Case distinction to define it first on compact intervals first, then use a limit. This is only much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\p. \(x,k)\p. content k *\<^sub>R (if x \ s then f x else 0)) \ z) (division_filter (cbox a b)) \ norm (z - I) < e)))" lemma has_integral_cbox: "(f has_integral I) (cbox a b) \ ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter (cbox a b))" by (auto simp add: has_integral_def) lemma has_integral: "(f has_integral y) (cbox a b) \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) lemma has_integral_real: "(f has_integral y) {a..b::real} \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of {a..b} \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" unfolding box_real[symmetric] by (rule has_integral) lemma has_integralD[dest]: assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains \ where "gauge \" and "\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,k)\\. content k *\<^sub>R f x) - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ (if \a b. i = cbox a b then (f has_integral y) i else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" by (subst has_integral_def) (auto simp add: has_integral_cbox) lemma has_integral_altD: assumes "(f has_integral y) i" and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" and "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" by auto subsection \Basic theorems about integrals\ lemma has_integral_eq_rhs: "(f has_integral j) S \ i = j \ (f has_integral i) S" by (rule forw_subst) lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" proof (rule ccontr) let ?e = "norm (k1 - k2)/2" let ?F = "(\x. if x \ i then f x else 0)" assume "k1 \ k2" then have e: "?e > 0" by auto have nonbox: "\ (\a b. i = cbox a b)" using \k1 \ k2\ assms has_integral_unique_cbox by blast obtain B1 where B1: "0 < B1" "\a b. ball 0 B1 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast obtain B2 where B2: "0 < B2" "\a b. ball 0 B2 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric) obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" using B1(2)[OF ab(1)] by blast obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" using B2(2)[OF ab(2)] by blast have "z = w" using has_integral_unique_cbox[OF w(1) z(1)] by auto then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] by (auto simp add: norm_minus_commute) also have "\ < norm (k1 - k2)/2 + norm (k1 - k2)/2" by (metis add_strict_mono z(2) w(2)) finally show False by auto qed lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" unfolding integral_def by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def apply (erule subst) apply (rule someI_ex) by blast lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" using eventually_division_filter_tagged_division[of "cbox a b"] additive_content_tagged_division[of _ a b] by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) lemma has_integral_const_real [intro]: fixes a b :: real shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" by (metis box_real(2) has_integral_const) lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" by blast lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" by (metis box_real(2) integral_const) lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ cbox a b \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] assms by (subst tendsto_cong[where g="\_. 0"]) (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) lemma has_integral_is_0: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ S \ f x = 0" shows "(f has_integral 0) S" proof (cases "(\a b. S = cbox a b)") case True with assms has_integral_is_0_cbox show ?thesis by blast next case False have *: "(\x. if x \ S then f x else 0) = (\x. 0)" by (auto simp add: assms) show ?thesis using has_integral_is_0_cbox False by (subst has_integral_alt) (force simp add: *) qed lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) S" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) S \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) (cbox a b)" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) (cbox a b)" proof - interpret bounded_linear h using h . show ?thesis unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] by (simp add: sum scaleR split_beta') qed lemma has_integral_linear: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) S" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) S" proof (cases "(\a b. S = cbox a b)") case True with f h has_integral_linear_cbox show ?thesis by blast next case False interpret bounded_linear h using h . from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast let ?S = "\f x. if x \ S then f x else 0" show ?thesis proof (subst has_integral_alt, clarsimp simp: False) fix e :: real assume e: "e > 0" have *: "0 < e/B" using e B(1) by simp obtain M where M: "M > 0" "\a b. ball 0 M \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - y) < e/B" using has_integral_altD[OF f False *] by blast show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e)" proof (rule exI, intro allI conjI impI) show "M > 0" using M by metis next fix a b::'n assume sb: "ball 0 M \ cbox a b" obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" using M(2)[OF sb] by blast have *: "?S(h \ f) = h \ ?S f" using zero by auto show "\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e" apply (rule_tac x="h z" in exI) apply (simp add: * has_integral_linear_cbox[OF z(1) h]) apply (metis B diff le_less_trans pos_less_divide_eq z(2)) done qed qed qed lemma has_integral_scaleR_left: "(f has_integral y) S \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) lemma integrable_on_scaleR_left: assumes "f integrable_on A" shows "(\x. f x *\<^sub>R y) integrable_on A" using assms has_integral_scaleR_left unfolding integrable_on_def by blast lemma has_integral_mult_left: fixes c :: "_ :: real_normed_algebra" shows "(f has_integral y) S \ ((\x. f x * c) has_integral (y * c)) S" using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) lemma has_integral_divide: fixes c :: "_ :: real_normed_div_algebra" shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" unfolding divide_inverse by (simp add: has_integral_mult_left) text\The case analysis eliminates the condition \<^term>\f integrable_on S\ at the cost of the type class constraint \division_ring\\ corollary integral_mult_left [simp]: fixes c:: "'a::{real_normed_algebra,division_ring}" shows "integral S (\x. f x * c) = integral S f * c" proof (cases "f integrable_on S \ c = 0") case True then show ?thesis by (force intro: has_integral_mult_left) next case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) qed corollary integral_mult_right [simp]: fixes c:: "'a::{real_normed_field}" shows "integral S (\x. c * f x) = c * integral S f" by (simp add: mult.commute [of c]) corollary integral_divide [simp]: fixes z :: "'a::real_normed_field" shows "integral S (\x. f x / z) = integral S (\x. f x) / z" using integral_mult_left [of S f "inverse z"] by (simp add: divide_inverse_commute) lemma has_integral_mult_right: fixes c :: "'a :: real_normed_algebra" shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) lemma has_integral_cmul: "(f has_integral k) S \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" unfolding o_def[symmetric] by (metis has_integral_linear bounded_linear_scaleR_right) lemma has_integral_cmult_real: fixes c :: real assumes "c \ 0 \ (f has_integral x) A" shows "((\x. c * f x) has_integral c * x) A" proof (cases "c = 0") case True then show ?thesis by simp next case False from has_integral_cmul[OF assms[OF this], of c] show ?thesis unfolding real_scaleR_def . qed lemma has_integral_neg: "(f has_integral k) S \ ((\x. -(f x)) has_integral -k) S" by (drule_tac c="-1" in has_integral_cmul) auto lemma has_integral_neg_iff: "((\x. - f x) has_integral k) S \ (f has_integral - k) S" using has_integral_neg[of f "- k"] has_integral_neg[of "\x. - f x" k] by auto lemma has_integral_add_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" shows "((\x. f x + g x) has_integral (k + l)) (cbox a b)" using assms unfolding has_integral_cbox by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) lemma has_integral_add: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" shows "((\x. f x + g x) has_integral (k + l)) S" proof (cases "\a b. S = cbox a b") case True with has_integral_add_cbox assms show ?thesis by blast next let ?S = "\f x. if x \ S then f x else 0" case False then show ?thesis proof (subst has_integral_alt, clarsimp, goal_cases) case (1 e) then have e2: "e/2 > 0" by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - k) < e/2" using has_integral_altD[OF f False e2] by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ (cbox a b) \ \z. (?S g has_integral z) (cbox a b) \ norm (z - l) < e/2" using has_integral_altD[OF g False e2] by blast show ?case proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \0 < Bf\) fix a b assume "ball 0 (max Bf Bg) \ cbox a (b::'n)" then have fs: "ball 0 Bf \ cbox a (b::'n)" and gs: "ball 0 Bg \ cbox a (b::'n)" by auto obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" using Bf[OF fs] by blast obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" using Bg[OF gs] by blast have *: "\x. (if x \ S then f x + g x else 0) = (?S f x) + (?S g x)" by auto show "\z. (?S(\x. f x + g x) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" apply (rule_tac x="w + z" in exI) apply (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) apply (auto simp add: field_simps) done qed qed qed lemma has_integral_diff: "(f has_integral k) S \ (g has_integral l) S \ ((\x. f x - g x) has_integral (k - l)) S" using has_integral_add[OF _ has_integral_neg, of f k S g l] by (auto simp: algebra_simps) lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by (rule integral_unique has_integral_0)+ lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" by (rule integral_unique) (metis integrable_integral has_integral_add) lemma integral_cmul [simp]: "integral S (\x. c *\<^sub>R f x) = c *\<^sub>R integral S f" proof (cases "f integrable_on S \ c = 0") case True with has_integral_cmul integrable_integral show ?thesis by fastforce next case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" unfolding real_scaleR_def[symmetric] integral_cmul .. lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" proof (cases "f integrable_on S") case True then show ?thesis by (simp add: has_integral_neg integrable_integral integral_unique) next case False then have "\ (\x. - f x) integrable_on S" using has_integral_neg [of "(\x. - f x)" _ S ] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" by (rule integral_unique) (metis integrable_integral has_integral_diff) lemma integrable_0: "(\x. 0) integrable_on S" unfolding integrable_on_def using has_integral_0 by auto lemma integrable_add: "f integrable_on S \ g integrable_on S \ (\x. f x + g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_add) lemma integrable_cmul: "f integrable_on S \ (\x. c *\<^sub>R f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_cmul) lemma integrable_on_scaleR_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c *\<^sub>R f x) integrable_on S \ f integrable_on S" using integrable_cmul[of "\x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \c \ 0\ by auto lemma integrable_on_cmult_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on S \ f integrable_on S" using integrable_on_scaleR_iff [of c f] assms by simp lemma integrable_on_cmult_left: assumes "f integrable_on S" shows "(\x. of_real c * f x) integrable_on S" using integrable_cmul[of f S "of_real c"] assms by (simp add: scaleR_conv_of_real) lemma integrable_neg: "f integrable_on S \ (\x. -f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_neg) lemma integrable_neg_iff: "(\x. -f(x)) integrable_on S \ f integrable_on S" using integrable_neg by fastforce lemma integrable_diff: "f integrable_on S \ g integrable_on S \ (\x. f x - g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_diff) lemma integrable_linear: "f integrable_on S \ bounded_linear h \ (h \ f) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_linear) lemma integral_linear: "f integrable_on S \ bounded_linear h \ integral S (h \ f) = h (integral S f)" apply (rule has_integral_unique [where i=S and f = "h \ f"]) apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) done lemma integrable_on_cnj_iff: "(\x. cnj (f x)) integrable_on A \ f integrable_on A" using integrable_linear[OF _ bounded_linear_cnj, of f A] integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] by (auto simp: o_def) lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" by (cases "f integrable_on A") (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] o_def integrable_on_cnj_iff not_integrable_integral) lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" shows "integral S (\x. f x \ k) = integral S f \ k" unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. lemma has_integral_sum: assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" using assms(1) subset_refl[of T] proof (induct rule: finite_subset_induct) case empty then show ?case by auto next case (insert x F) with assms show ?case by (simp add: has_integral_add) qed lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ integral S (\x. \a\I. f a x) = (\a\I. integral S (f a))" by (simp add: has_integral_sum integrable_integral integral_unique) lemma integrable_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ (\x. \a\I. f a x) integrable_on S" unfolding integrable_on_def using has_integral_sum[of I] by metis lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" and "(f has_integral k) s" shows "(g has_integral k) s" using has_integral_diff[OF assms(2), of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] using assms(1) by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" unfolding integrable_on_def using has_integral_eq[of s f g] has_integral_eq by blast lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" using has_integral_eq[of s f g] has_integral_eq[of s g f] assms by auto lemma integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide: fixes f :: "_ \ 'b :: real_normed_field" assumes "f integrable_on s" shows "(\x. f x / of_real c) integrable_on s" by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by (rule integral_unique) auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "\n. gauge (\x. \{\ i x |i. i \ {0..n}})" apply (rule gauge_Inter) using \ by auto then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{\ i x |i. i \ {0..n}}) fine p" by (meson fine_division_exists) then obtain p where p: "\z. p z tagged_division_of cbox a b" "\z. (\x. \{\ i x |i. i \ {0..z}}) fine p z" by meson have dp: "\i n. i\n \ \ i fine p n" using p unfolding fine_Inter using atLeastAtMost_iff by blast have "Cauchy (\n. sum (\(x,K). content K *\<^sub>R (f x)) (p n))" proof (rule CauchyI) fix e::real assume "0 < e" then obtain N where "N \ 0" and N: "inverse (real N) < e" using real_arch_inverse[of e] by blast show "\M. \m\M. \n\M. norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" proof (intro exI allI impI) fix m n assume mn: "N \ m" "N \ n" have "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" by (simp add: p(1) dp mn \) also have "... < e" using N \N \ 0\ \0 < e\ by (auto simp: field_simps) finally show "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" . qed qed then obtain y where y: "\no. \n\no. norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) show ?l unfolding integrable_on_def has_integral proof (rule_tac x=y in exI, clarify) fix e :: real assume "e>0" then have e2: "e/2 > 0" by auto then obtain N1::nat where N1: "N1 \ 0" "inverse (real N1) < e/2" using real_arch_inverse by blast obtain N2::nat where N2: "\n. n \ N2 \ norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < e/2" using y[OF e2] by metis show "\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" proof (intro exI conjI allI impI) show "gauge (\ (N1+N2))" using \ by auto show "norm ((\(x,K) \ q. content K *\<^sub>R f x) - y) < e" if "q tagged_division_of cbox a b \ \ (N1+N2) fine q" for q proof (rule norm_triangle_half_r) have "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < 1 / (real (N1+N2) + 1)" by (rule \; simp add: dp p that) also have "... < e/2" using N1 \0 < e\ by (auto simp: field_simps intro: less_le_trans) finally show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < e/2" . show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - y) < e/2" using N2 le_add_same_cancel2 by blast qed qed qed qed subsection \Additivity of integral on abutting intervals\ lemma tagged_division_split_left_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. x \ k \ c}) = {}" by (metis tagged_division_split_left_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed lemma tagged_division_split_right_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. c \ x \ k}) = {}" by (metis tagged_division_split_right_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed proposition has_integral_split: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(f has_integral (i + j)) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "0 < e" then have e: "e/2 > 0" by auto obtain \1 where \1: "gauge \1" and \1norm: "\\. \\ tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - i) < e/2" apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done obtain \2 where \2: "gauge \2" and \2norm: "\\. \\ tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine \\ \ norm ((\(x, k) \ \. content k *\<^sub>R f x) - j) < e/2" apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done let ?\ = "\x. if x\k = c then (\1 x \ \2 x) else ball x \x\k - c\ \ \1 x \ \2 x" have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - (i + j)) < e)" proof (rule_tac x="?\" in exI, safe) fix p assume p: "p tagged_division_of (cbox a b)" and "?\ fine p" have ab_eqp: "cbox a b = \{K. \x. (x, K) \ p}" using p by blast have xk_le_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have xk_ge_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have fin_finite: "finite {(x,f K) | x K. (x,K) \ s \ P x K}" if "finite s" for s and f :: "'a set \ 'a set" and P :: "'a \ 'a set \ bool" proof - from that have "finite ((\(x,K). (x, f K)) ` s)" by auto then show ?thesis by (rule rev_finite_subset) auto qed { fix \ :: "'a set \ 'a set" fix i :: "'a \ 'a set" assume "i \ (\(x, k). (x, \ k)) ` p - {(x, \ k) |x k. (x, k) \ p \ \ k \ {}}" then obtain x K where xk: "i = (x, \ K)" "(x,K) \ p" "(x, \ K) \ {(x, \ K) |x K. (x,K) \ p \ \ K \ {}}" by auto have "content (\ K) = 0" using xk using content_empty by auto then have "(\(x,K). content K *\<^sub>R f x) i = 0" unfolding xk split_conv by auto } note [simp] = this have "finite p" using p by blast let ?M1 = "{(x, K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \1_fine: "\1 fine ?M1" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" proof (rule \1norm [OF tagged_division_ofI \1_fine]) show "finite ?M1" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M1" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_le_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M1" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed moreover let ?M2 = "{(x,K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \2_fine: "\2 fine ?M2" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" proof (rule \2norm [OF tagged_division_ofI \2_fine]) show "finite ?M2" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M2" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_ge_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M2" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed ultimately have "norm (((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" using norm_add_less by blast moreover have "((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" proof - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" by auto have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" by auto have *: "\\ :: 'a set \ 'a set. (\(x,K)\{(x, \ K) |x K. (x,K) \ p \ \ K \ {}}. content K *\<^sub>R f x) = (\(x,K)\(\(x,K). (x, \ K)) ` p. content K *\<^sub>R f x)" by (rule sum.mono_neutral_left) (auto simp: \finite p\) have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto moreover have "\ = (\(x,K) \ p. content (K \ {x. x \ k \ c}) *\<^sub>R f x) + (\(x,K) \ p. content (K \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding * apply (subst (1 2) sum.reindex_nontrivial) apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content simp: cont_eq \finite p\) done moreover have "\x. x \ p \ (\(a,B). content (B \ {a. a \ k \ c}) *\<^sub>R f a) x + (\(a,B). content (B \ {a. c \ a \ k}) *\<^sub>R f a) x = (\(a,B). content B *\<^sub>R f a) x" proof clarify fix a B assume "(a, B) \ p" with p obtain u v where uv: "B = cbox u v" by blast then show "content (B \ {x. x \ k \ c}) *\<^sub>R f a + content (B \ {x. c \ x \ k}) *\<^sub>R f a = content B *\<^sub>R f a" by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) qed ultimately show ?thesis by (auto simp: sum.distrib[symmetric]) qed ultimately show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed subsection \A sort of converse, integrability on subintervals\ lemma has_integral_separate_sides: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes f: "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((sum (\(x,k). content k *\<^sub>R f x) p1 + sum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - obtain \ where d: "gauge \" "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e" using has_integralD[OF f \e > 0\] by metis { fix p1 p2 assume tdiv1: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" and "\ fine p1" note p1=tagged_division_ofD[OF this(1)] assume tdiv2: "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" and "\ fine p2" note p2=tagged_division_ofD[OF this(1)] note tagged_division_Un_interval[OF tdiv1 tdiv2] note p12 = tagged_division_ofD[OF this] this { fix a b assume ab: "(a, b) \ p1 \ p2" have "(a, b) \ p1" using ab by auto obtain u v where uv: "b = cbox u v" using \(a, b) \ p1\ p1(4) by moura have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce moreover have "interior {x::'a. x \ k = c} = {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where x: "x \ interior {x::'a. x\k = c}" by auto then obtain \ where "0 < \" and \: "ball x \ \ {x. x \ k = c}" using mem_interior by metis have x: "x\k = c" using x interior_subset by fastforce have *: "\i. i \ Basis \ \(x - (x + (\/2) *\<^sub>R k)) \ i\ = (if i = k then \/2 else 0)" using \0 < \\ k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (\/2 ) *\<^sub>R k)) \ i\) = (\i\Basis. (if i = k then \/2 else 0))" using "*" by (blast intro: sum.cong) also have "\ < \" by (subst sum.delta) (use \0 < \\ in auto) finally have "x + (\/2) *\<^sub>R k \ ball x \" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) then have "x + (\/2) *\<^sub>R k \ {x. x\k = c}" using \ by auto then show False using \0 < \\ x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior using interior_mono by blast then have "content b *\<^sub>R f a = 0" by auto } then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" by (subst sum.union_inter_neutral) (auto simp: p1 p2) also have "\ < e" using d(2) p12 by (simp add: fine_Un k \\ fine p1\ \\ fine p2\) finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . } then show ?thesis using d(1) that by auto qed lemma integrable_split [intro]: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on cbox a b" and k: "k \ Basis" shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis1) and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis2) proof - obtain y where y: "(f has_integral y) (cbox a b)" using f by blast define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a' b]) fix p assume "p tagged_division_of cbox a' b" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis1 by (simp add: interval_split[OF k] integrable_Cauchy) have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a b']) fix p assume "p tagged_division_of cbox a b'" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis2 by (simp add: interval_split[OF k] integrable_Cauchy) qed lemma operative_integralI: fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lift_option (+)) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" proof - interpret comm_monoid "lift_option plus" "Some (0::'b)" by (rule comm_monoid_lift_option) (rule add.comm_monoid_axioms) show ?thesis proof fix a b c fix k :: 'a assume k: "k \ Basis" show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = lift_option (+) (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" proof (cases "f integrable_on cbox a b") case True with k show ?thesis apply (simp add: integrable_split) apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) apply (auto intro: integrable_integral) done next case False have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" then have "f integrable_on cbox a b" unfolding integrable_on_def apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (rule has_integral_split[OF _ _ k]) apply (auto intro: integrable_integral) done then show False using False by auto qed then show ?thesis using False by auto qed next fix a b :: 'a assume "box a b = {}" then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" using has_integral_null_eq by (auto simp: integrable_on_null content_eq_0_interior) qed qed subsection \Bounds on the norm of Riemann sums and the integral itself\ lemma dsum_bound: assumes "p division_of (cbox a b)" and "norm c \ e" shows "norm (sum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" proof - have sumeq: "(\i\p. \content i\) = sum content p" apply (rule sum.cong) using assms apply simp apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) done have e: "0 \ e" using assms(2) norm_ge_zero order_trans by blast have "norm (sum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" using norm_sum by blast also have "... \ e * (\i\p. \content i\)" by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) also have "... \ e * content (cbox a b)" apply (rule mult_left_mono [OF _ e]) apply (simp add: sumeq) using additive_content_division assms(1) eq_iff apply blast done finally show ?thesis . qed lemma rsum_bound: assumes p: "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis using p unfolding True tagged_division_of_trivial by auto next case False then have e: "e \ 0" by (meson ex_in_conv assms(2) norm_ge_zero order_trans) have sum_le: "sum (content \ snd) p \ content (cbox a b)" unfolding additive_content_tagged_division[OF p, symmetric] split_def by (auto intro: eq_refl) have con: "\xk. xk \ p \ 0 \ content (snd xk)" using tagged_division_ofD(4) [OF p] content_pos_le by force have norm: "\xk. xk \ p \ norm (f (fst xk)) \ e" unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms by (metis prod.collapse subset_eq) have "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" by (rule norm_sum) also have "... \ e * content (cbox a b)" unfolding split_def norm_scaleR apply (rule order_trans[OF sum_mono]) apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) apply (metis norm) unfolding sum_distrib_right[symmetric] using con sum_le apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) done finally show ?thesis . qed lemma rsum_diff_bound: assumes "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p - sum (\(x,k). content k *\<^sub>R g x) p) \ e * content (cbox a b)" apply (rule order_trans[OF _ rsum_bound[OF assms]]) apply (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) done lemma has_integral_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof (rule ccontr) assume "\ ?thesis" then have "norm i - B * content (cbox a b) > 0" by auto with f[unfolded has_integral] obtain \ where "gauge \" and \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, K)\p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" by metis then obtain p where p: "p tagged_division_of cbox a b" and "\ fine p" using fine_division_exists by blast have "\s B. norm s \ B \ \ norm (s - i) < norm i - B" unfolding not_less by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) then show False using \ [OF p \\ fine p\] rsum_bound[OF p] assms by metis qed corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "f integrable_on (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ lemma rsum_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes p: "p tagged_division_of (cbox a b)" and "\x. x \ cbox a b \ (f x)\i \ (g x)\i" shows "(\(x, K)\p. content K *\<^sub>R f x) \ i \ (\(x, K)\p. content K *\<^sub>R g x) \ i" unfolding inner_sum_left proof (rule sum_mono, clarify) fix x K assume ab: "(x, K) \ p" with p obtain u v where K: "K = cbox u v" by blast then show "(content K *\<^sub>R f x) \ i \ (content K *\<^sub>R g x) \ i" by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) qed lemma has_integral_component_le: fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) S" "(g has_integral j) S" and f_le_g: "\x. x \ S \ (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have ik_le_jk: "i\k \ j\k" if f_i: "(f has_integral i) (cbox a b)" and g_j: "(g has_integral j) (cbox a b)" and le: "\x\cbox a b. (f x)\k \ (g x)\k" for a b i and j :: 'b and f g :: "'a \ 'b" proof (rule ccontr) assume "\ ?thesis" then have *: "0 < (i\k - j\k) / 3" by auto obtain \1 where "gauge \1" and \1: "\p. \p tagged_division_of cbox a b; \1 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < (i \ k - j \ k) / 3" using f_i[unfolded has_integral,rule_format,OF *] by fastforce obtain \2 where "gauge \2" and \2: "\p. \p tagged_division_of cbox a b; \2 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R g x) - j) < (i \ k - j \ k) / 3" using g_j[unfolded has_integral,rule_format,OF *] by fastforce obtain p where p: "p tagged_division_of cbox a b" and "\1 fine p" "\2 fine p" using fine_division_exists[OF gauge_Int[OF \gauge \1\ \gauge \2\], of a b] unfolding fine_Int by metis then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" using le_less_trans[OF Basis_le_norm[OF k]] k \1 \2 by metis+ then show False unfolding inner_simps using rsum_component_le[OF p] le by (fastforce simp add: abs_real_def split: if_split_asm) qed show ?thesis proof (cases "\a b. S = cbox a b") case True with ik_le_jk assms show ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ i\k \ j\k" then have ij: "(i\k - j\k) / 3 > 0" by auto obtain B1 where "0 < B1" and B1: "\a b. ball 0 B1 \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast obtain B2 where "0 < B2" and B2: "\a b. ball 0 B2 \ cbox a b \ \z. ((\x. if x \ S then g x else 0) has_integral z) (cbox a b) \ norm (z - j) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_cbox_symmetric[OF this] obtain a b::'a where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (meson Un_subset_iff) then obtain w1 w2 where int_w1: "((\x. if x \ S then f x else 0) has_integral w1) (cbox a b)" and norm_w1: "norm (w1 - i) < (i \ k - j \ k) / 3" and int_w2: "((\x. if x \ S then g x else 0) has_integral w2) (cbox a b)" and norm_w2: "norm (w2 - j) < (i \ k - j \ k) / 3" using B1 B2 by blast have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: if_split_asm) have "\(w1 - i) \ k\ < (i \ k - j \ k) / 3" "\(w2 - j) \ k\ < (i \ k - j \ k) / 3" using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ moreover have "w1\k \ w2\k" using ik_le_jk int_w1 int_w2 f_le_g by auto ultimately show False unfolding inner_simps by(rule *) qed qed qed lemma integral_component_le: fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on S" "g integrable_on S" and "\x. x \ S \ (f x)\k \ (g x)\k" shows "(integral S f)\k \ (integral S g)\k" apply (rule has_integral_component_le) using integrable_integral assms apply auto done lemma has_integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" proof (cases "f integrable_on S") case True show ?thesis apply (rule has_integral_component_nonneg) using assms True apply auto done next case False then show ?thesis by (simp add: not_integrable_integral) qed lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::euclidean_space => 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" apply (rule has_integral_component_lbound) using assms unfolding has_integral_integral apply auto done lemma integral_component_lbound_real: assumes "f integrable_on {a ::real..b}" and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" apply (rule has_integral_component_ubound) using assms unfolding has_integral_integral apply auto done lemma integral_component_ubound_real: fixes f :: "real \ 'a::euclidean_space" assumes "f integrable_on {a..b}" and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ lemma real_arch_invD: "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" by (subst(asm) real_arch_inverse) lemma integrable_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "\e. e > 0 \ \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" shows "f integrable_on cbox a b" proof (cases "content (cbox a b) > 0") case False then show ?thesis using has_integral_null by (simp add: content_lt_nz integrable_on_def) next case True have "1 / (real n + 1) > 0" for n by auto then have "\g. (\x\cbox a b. norm (f x - g x) \ 1 / (real n + 1)) \ g integrable_on cbox a b" for n using assms by blast then obtain g where g_near_f: "\n x. x \ cbox a b \ norm (f x - g n x) \ 1 / (real n + 1)" and int_g: "\n. g n integrable_on cbox a b" by metis then obtain h where h: "\n. (g n has_integral h n) (cbox a b)" unfolding integrable_on_def by metis have "Cauchy h" unfolding Cauchy_def proof clarify fix e :: real assume "e>0" then have "e/4 / content (cbox a b) > 0" using True by (auto simp: field_simps) then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" proof (rule exI [where x=M], clarify) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto then obtain gm gn where "gauge gm" "gauge gn" and gm: "\\. \ tagged_division_of cbox a b \ gm fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g m x) - h m) < e/4" and gn: "\\. \ tagged_division_of cbox a b \ gn fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g n x) - h n) < e/4" using h[unfolded has_integral] by meson then obtain \ where \: "\ tagged_division_of cbox a b" "(\x. gm x \ gn x) fine \" by (metis (full_types) fine_division_exists gauge_Int) have triangle3: "norm (i1 - i2) < e" if no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b proof - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed have finep: "gm fine \" "gn fine \" using fine_Int \ by auto have norm_le: "norm (g n x - g m x) \ 2 / real M" if x: "x \ cbox a b" for x proof - have "norm (f x - g n x) + norm (f x - g m x) \ 1 / (real n + 1) + 1 / (real m + 1)" using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp also have "\ \ 1 / (real M) + 1 / (real M)" apply (rule add_mono) using \M \ 0\ m n by (auto simp: field_split_simps) also have "\ = 2 / real M" by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] by (auto simp: algebra_simps simp add: norm_minus_commute) qed have "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ 2 / real M * content (cbox a b)" by (blast intro: norm_le rsum_diff_bound[OF \(1), where e="2 / real M"]) also have "... \ e/2" using M True by (auto simp: field_simps) finally have le_e2: "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ e/2" . then show "dist (h m) (h n) < e" unfolding dist_norm using gm gn \ finep by (auto intro!: triangle3) qed qed then obtain s where s: "h \ s" using convergent_eq_Cauchy[symmetric] by blast show ?thesis unfolding integrable_on_def has_integral proof (rule_tac x=s in exI, clarify) fix e::real assume e: "0 < e" then have "e/3 > 0" by auto then obtain N1 where N1: "\n\N1. norm (h n - s) < e/3" using LIMSEQ_D [OF s] by metis from e True have "e/3 / content (cbox a b) > 0" by (auto simp: field_simps) then obtain N2 :: nat where "N2 \ 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) obtain g' where "gauge g'" and g': "\\. \ tagged_division_of cbox a b \ g' fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" by (metis h has_integral \e/3 > 0\) have *: "norm (sf - s) < e" if no: "norm (sf - sg) \ e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h proof - have "norm (sf - s) \ norm (sf - sg) + norm (sg - h) + norm (h - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] using norm_triangle_ineq[of "sg - h" " h - s"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed { fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "g' fine \" then have norm_less: "norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" using g' by blast have "content (cbox a b) < e/3 * (of_nat N2)" using \N2 \ 0\ N2 using True by (auto simp: field_split_simps) moreover have "e/3 * of_nat N2 \ e/3 * (of_nat (N1 + N2) + 1)" using \e>0\ by auto ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" by linarith then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \ e/3" unfolding inverse_eq_divide by (auto simp: field_simps) have ne3: "norm (h (N1 + N2) - s) < e/3" using N1 by auto have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - (\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x)) \ 1 / (real (N1 + N2) + 1) * content (cbox a b)" by (blast intro: g_near_f rsum_diff_bound[OF ptag]) then have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e" by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) } then show "\d. gauge d \ (\\. \ tagged_division_of cbox a b \ d fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e)" by (blast intro: g' \gauge g'\) qed qed lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] subsection \Negligible sets\ definition "negligible (s:: 'a::euclidean_space set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsubsection \Negligibility of hyperplane\ lemma content_doublesplit: fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" proof cases assume *: "a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j)" define a' where "a' d = (\j\Basis. (if j = k then max (a\j) (c - d) else a\j) *\<^sub>R j)" for d define b' where "b' d = (\j\Basis. (if j = k then min (b\j) (c + d) else b\j) *\<^sub>R j)" for d have "((\d. \j\Basis. (b' d - a' d) \ j) \ (\j\Basis. (b' 0 - a' 0) \ j)) (at_right 0)" by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) also have "(\j\Basis. (b' 0 - a' 0) \ j) = 0" using k * by (intro prod_zero bexI[OF _ k]) (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) also have "((\d. \j\Basis. (b' d - a' d) \ j) \ 0) (at_right 0) = ((\d. content (cbox a b \ {x. \x\k - c\ \ d})) \ 0) (at_right 0)" proof (intro tendsto_cong eventually_at_rightI) fix d :: real assume d: "d \ {0<..<1}" have "cbox a b \ {x. \x\k - c\ \ d} = cbox (a' d) (b' d)" for d using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) moreover have "j \ Basis \ a' d \ j \ b' d \ j" for j using * d k by (auto simp: a'_def b'_def) ultimately show "(\j\Basis. (b' d - a' d) \ j) = content (cbox a b \ {x. \x\k - c\ \ d})" by simp qed simp finally have "((\d. content (cbox a b \ {x. \x \ k - c\ \ d})) \ 0) (at_right 0)" . from order_tendstoD(2)[OF this \0] obtain d' where "0 < d'" and d': "\y. y > 0 \ y < d' \ content (cbox a b \ {x. \x \ k - c\ \ y}) < e" by (subst (asm) eventually_at_right[of _ 1]) auto show ?thesis by (rule that[of "d'/2"], insert \0 d'[of "d'/2"], auto) next assume *: "\ (a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j))" then have "(\j\Basis. b \ j < a \ j) \ (c < a \ k \ b \ k < c)" by (auto simp: not_le) show thesis proof cases assume "\j\Basis. b \ j < a \ j" then have [simp]: "cbox a b = {}" using box_ne_empty(1)[of a b] by auto show ?thesis by (rule that[of 1]) (simp_all add: \0) next assume "\ (\j\Basis. b \ j < a \ j)" with * have "c < a \ k \ b \ k < c" by auto then show thesis proof assume c: "c < a \ k" moreover have "x \ cbox a b \ c \ x \ k" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (a \ k - c)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(a \ k - c)/2"] show ?thesis by auto next assume c: "b \ k < c" moreover have "x \ cbox a b \ x \ k \ c" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (c - b \ k)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(c - b \ k)/2"] show ?thesis by auto qed qed qed proposition negligible_standard_hyperplane[intro]: fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral proof clarsimp fix a b and e::real assume "e > 0" with k obtain d where "0 < d" and d: "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" by (metis content_doublesplit) let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K * ?i x\ < e)" proof (intro exI, safe) show "gauge (\x. ball x d)" using \0 < d\ by blast next fix \ assume p: "\ tagged_division_of (cbox a b)" "(\x. ball x d) fine \" have "content L = content (L \ {x. \x \ k - c\ \ d})" if "(x, L) \ \" "?i x \ 0" for x L proof - have xk: "x\k = c" using that by (simp add: indicator_def split: if_split_asm) have "L \ {x. \x \ k - c\ \ d}" proof fix y assume y: "y \ L" have "L \ ball x d" using p(2) that(1) by auto then have "norm (x - y) < d" by (simp add: dist_norm subset_iff y) then have "\(x - y) \ k\ < d" using k norm_bound_Basis_lt by blast then show "y \ {x. \x \ k - c\ \ d}" unfolding inner_simps xk by auto qed then show "content L = content (L \ {x. \x \ k - c\ \ d})" by (metis inf.orderE) qed then have *: "(\(x,K)\\. content K * ?i x) = (\(x,K)\\. content (K \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" by (force simp add: split_paired_all intro!: sum.cong [OF refl]) note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * indicator {x. x \ k = c} x) < e" proof - have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}))" by (force simp add: indicator_def intro!: sum_mono) also have "\ < e" proof (subst sum.over_tagged_division_lemma[OF p(1)]) fix u v::'a assume "box u v = {}" then have *: "content (cbox u v) = 0" unfolding content_eq_0_interior by simp have "cbox u v \ {x. \x \ k - c\ \ d} \ cbox u v" by auto then have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] by (rule content_subset) then show "content (cbox u v \ {x. \x \ k - c\ \ d}) = 0" unfolding * interval_doublesplit[OF k] by (blast intro: antisym) next have "(\l\snd ` \. content (l \ {x. \x \ k - c\ \ d})) = sum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}})" proof (subst (2) sum.reindex_nontrivial) fix x y assume "x \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" then obtain x' y' where "(x', x) \ \" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ \" "y \ {x. \x \ k - c\ \ d} \ {}" by (auto) from p'(5)[OF \(x', x) \ \\ \(y', y) \ \\] \x \ y\ have "interior (x \ y) = {}" by auto moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" by (auto intro: interior_mono) ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" by (auto simp: eq) then show "content (x \ {x. \x \ k - c\ \ d}) = 0" using p'(4)[OF \(x', x) \ \\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) qed (insert p'(1), auto intro!: sum.mono_neutral_right) also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" by simp also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto also have "\ < e" using d by simp finally show "(\K\snd ` \. content (K \ {x. \x \ k - c\ \ d})) < e" . qed finally show "(\(x, K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed then show "\\(x, K)\\. content K * ?i x\ < e" unfolding * apply (subst abs_of_nonneg) using measure_nonneg by (force simp add: indicator_def intro: sum_nonneg)+ qed qed corollary negligible_standard_hyperplane_cart: fixes k :: "'a::finite" shows "negligible {x. x$k = (0::real)}" by (simp add: cart_eq_inner_axis negligible_standard_hyperplane) subsubsection \Hence the main theorem about negligible sets\ lemma has_integral_negligible_cbox: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and 0: "\x. x \ S \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "e > 0" then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n by simp then have "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n))" for n using negs [unfolded negligible_def has_integral] by auto then obtain \ where gd: "\n. gauge (\ n)" and \: "\n \. \\ tagged_division_of cbox a b; \ n fine \\ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n)" by metis show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e)" proof (intro exI, safe) show "gauge (\x. \ (nat \norm (f x)\) x)" using gd by (auto simp: gauge_def) show "norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e" if "\ tagged_division_of (cbox a b)" "(\x. \ (nat \norm (f x)\) x) fine \" for \ proof (cases "\ = {}") case True with \0 < e\ show ?thesis by simp next case False obtain N where "Max ((\(x, K). norm (f x)) ` \) \ real N" using real_arch_simple by blast then have N: "\x. x \ (\(x, K). norm (f x)) ` \ \ x \ real N" by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) have "\i. \q. q tagged_division_of (cbox a b) \ (\ i) fine q \ (\(x,K) \ \. K \ (\ i) x \ (x, K) \ q)" by (auto intro: tagged_division_finer[OF that(1) gd]) from choice[OF this] obtain q where q: "\n. q n tagged_division_of cbox a b" "\n. \ n fine q n" "\n x K. \(x, K) \ \; K \ \ n x\ \ (x, K) \ q n" by fastforce have "finite \" using that(1) by blast then have sum_le_inc: "\finite T; \x y. (x,y) \ T \ (0::real) \ g(x,y); \y. y\\ \ \x. (x,y) \ T \ f(y) \ g(x,y)\ \ sum f \ \ sum g T" for f g T by (rule sum_le_included[of \ T g snd f]; force) have "norm (\(x,K) \ \. content K *\<^sub>R f x) \ (\(x,K) \ \. norm (content K *\<^sub>R f x))" unfolding split_def by (rule norm_sum) also have "... \ (\(i, j) \ Sigma {..N + 1} q. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" proof (rule sum_le_inc, safe) show "finite (Sigma {..N+1} q)" by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) next fix x K assume xk: "(x, K) \ \" define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, K). norm (f x)) ` \" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" unfolding n_def by auto then have "n \ {0..N + 1}" using N[OF *] by auto moreover have "K \ \ (nat \norm (f x)\) x" using that(2) xk by auto moreover then have "(x, K) \ q (nat \norm (f x)\)" by (simp add: q(3) xk) moreover then have "(x, K) \ q n" using n_def by blast moreover have "norm (content K *\<^sub>R f x) \ (real n + 1) * (content K * indicator S x)" proof (cases "x \ S") case False then show ?thesis by (simp add: 0) next case True have *: "content K \ 0" using tagged_division_ofD(4)[OF that(1) xk] by auto moreover have "content K * norm (f x) \ content K * (real n + 1)" by (simp add: mult_left_mono nfx(2)) ultimately show ?thesis using nfx True by (auto simp: field_simps) qed ultimately show "\y. (y, x, K) \ (Sigma {..N + 1} q) \ norm (content K *\<^sub>R f x) \ (real y + 1) * (content K *\<^sub>R indicator S x)" by force qed auto also have "... = (\i\N + 1. \j\q i. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" apply (rule sum_Sigma_product [symmetric]) using q(1) apply auto done also have "... \ (\i\N + 1. (real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\)" by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) also have "... \ (\i\N + 1. e/2/2 ^ i)" proof (rule sum_mono) show "(real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\ \ e/2/2 ^ i" if "i \ {..N + 1}" for i using \[of "q i" i] q by (simp add: divide_simps mult.left_commute) qed also have "... = e/2 * (\i\N + 1. (1/2) ^ i)" unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) also have "\ < e/2 * 2" proof (rule mult_strict_left_mono) have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..0 < e\ in auto) finally show ?thesis by auto qed qed qed proposition has_integral_negligible: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and "\x. x \ (T - S) \ f x = 0" shows "(f has_integral 0) T" proof (cases "\a b. T = cbox a b") case True then have "((\x. if x \ T then f x else 0) has_integral 0) T" using assms by (auto intro!: has_integral_negligible_cbox) then show ?thesis by (rule has_integral_eq [rotated]) auto next case False let ?f = "(\x. if x \ T then f x else 0)" have "((\x. if x \ T then f x else 0) has_integral 0) T" apply (auto simp: False has_integral_alt [of ?f]) apply (rule_tac x=1 in exI, auto) apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) done then show ?thesis by (rule_tac f="?f" in has_integral_eq) auto qed lemma assumes "negligible S" shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0" using has_integral_negligible [OF assms] by (auto simp: has_integral_iff) lemma has_integral_spike: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" and fint: "(f has_integral y) T" shows "(g has_integral y) T" proof - have *: "(g has_integral y) (cbox a b)" if "(f has_integral y) (cbox a b)" "\x \ cbox a b - S. g x = f x" for a b f and g:: "'b \ 'a" and y proof - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" using that by (intro has_integral_add has_integral_negligible) (auto intro!: \negligible S\) then show ?thesis by auto qed show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp split: if_split_asm) apply (blast dest: *) apply (erule_tac V = "\a b. T \ cbox a b" in thin_rl) apply (elim all_forward imp_forward ex_forward all_forward conj_forward asm_rl) apply (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) done qed lemma has_integral_spike_eq: assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" shows "(f has_integral y) T \ (g has_integral y) T" using has_integral_spike [OF \negligible S\] gf by metis lemma integrable_spike: assumes "f integrable_on T" "negligible S" "\x. x \ T - S \ g x = f x" shows "g integrable_on T" using assms unfolding integrable_on_def by (blast intro: has_integral_spike) lemma integral_spike: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "integral T f = integral T g" using has_integral_spike_eq[OF assms] by (auto simp: integral_def integrable_on_def) subsection \Some other trivialities about negligible sets\ lemma negligible_subset: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible (s - t)" using assms by (meson Diff_subset negligible_subset) lemma negligible_Int: assumes "negligible s \ negligible t" shows "negligible (s \ t)" using assms negligible_subset by force lemma negligible_Un: assumes "negligible S" and T: "negligible T" shows "negligible (S \ T)" proof - have "(indicat_real (S \ T) has_integral 0) (cbox a b)" if S0: "(indicat_real S has_integral 0) (cbox a b)" and "(indicat_real T has_integral 0) (cbox a b)" for a b proof (subst has_integral_spike_eq[OF T]) show "indicat_real S x = indicat_real (S \ T) x" if "x \ cbox a b - T" for x by (metis Diff_iff Un_iff indicator_def that) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" apply (subst insert_is_Un) unfolding negligible_Un_eq apply auto done lemma negligible_empty[iff]: "negligible {}" using negligible_insert by blast text\Useful in this form for backchaining\ lemma empty_imp_negligible: "S = {} \ negligible S" by simp lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" using assms by (induct s) auto lemma negligible_Union[intro]: assumes "finite \" and "\t. t \ \ \ negligible t" shows "negligible(\\)" using assms by induct auto lemma negligible: "negligible S \ (\T. (indicat_real S has_integral 0) T)" proof (intro iffI allI) fix T assume "negligible S" then show "(indicator S has_integral 0) T" by (meson Diff_iff has_integral_negligible indicator_simps(2)) qed (simp add: negligible_def) subsection \Finite case of the spike theorem is quite commonly needed\ lemma has_integral_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "(f has_integral y) T" shows "(g has_integral y) T" using assms has_integral_spike negligible_finite by blast lemma has_integral_spike_finite_eq: assumes "finite S" and "\x. x \ T - S \ g x = f x" shows "((f has_integral y) T \ (g has_integral y) T)" by (metis assms has_integral_spike_finite) lemma integrable_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "f integrable_on T" shows "g integrable_on T" using assms has_integral_spike_finite by blast lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed corollary has_integral_bound_real: fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and "(f has_integral i) {a..b}" and "\x. x \ {a..b} - S \ norm (f x) \ B" shows "norm i \ B * content {a..b}" by (metis assms box_real(2) has_integral_bound_spike_finite) subsection \In particular, the boundary of an interval is negligible\ lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" have "negligible ?A" by (force simp add: negligible_Union[OF finite_imageI]) moreover have "cbox a b - box a b \ ?A" by (force simp add: mem_box) ultimately show ?thesis by (rule negligible_subset) qed lemma has_integral_spike_interior: assumes f: "(f has_integral y) (cbox a b)" and gf: "\x. x \ box a b \ g x = f x" shows "(g has_integral y) (cbox a b)" apply (rule has_integral_spike[OF negligible_frontier_interval _ f]) using gf by auto lemma has_integral_spike_interior_eq: assumes "\x. x \ box a b \ g x = f x" shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" by (metis assms has_integral_spike_interior) lemma integrable_spike_interior: assumes "\x. x \ box a b \ g x = f x" and "f integrable_on cbox a b" shows "g integrable_on cbox a b" using assms has_integral_spike_interior_eq by blast subsection \Integrability of continuous functions\ lemma operative_approximableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" proof - interpret comm_monoid conj True by standard auto show ?thesis proof (standard, safe) fix a b :: 'b show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if "box a b = {}" for a b apply (rule_tac x=f in exI) using assms that by (auto simp: content_eq_0_interior) { fix c g and k :: 'b assume fg: "\x\cbox a b. norm (f x - g x) \ e" and g: "g integrable_on cbox a b" assume k: "k \ Basis" show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" apply (rule_tac[!] x=g in exI) using fg integrable_split[OF g k] by auto } show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if fg1: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" and g1: "g1 integrable_on cbox a b \ {x. x \ k \ c}" and fg2: "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" and g2: "g2 integrable_on cbox a b \ {x. c \ x \ k}" and k: "k \ Basis" for c k g1 g2 proof - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" proof (intro exI conjI ballI) show "norm (f x - ?g x) \ e" if "x \ cbox a b" for x by (auto simp: that assms fg1 fg2) show "?g integrable_on cbox a b" proof - have "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ with has_integral_split[OF _ _ k] show ?thesis unfolding integrable_on_def by blast qed qed qed qed qed lemma comm_monoid_set_F_and: "comm_monoid_set.F (\) True f s \ (finite s \ (\x\s. f x))" proof - interpret bool: comm_monoid_set "(\)" True proof qed auto show ?thesis by (induction s rule: infinite_finite_induct) auto qed lemma approximable_on_division: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" and d: "d division_of (cbox a b)" and f: "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - interpret operative conj True "\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i" using \0 \ e\ by (rule operative_approximableI) from f local.division [OF d] that show thesis by auto qed lemma integrable_continuous: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit) fix e :: real assume e: "e > 0" then obtain d where "0 < d" and d: "\x x'. \x \ cbox a b; x' \ cbox a b; dist x' x < d\ \ dist (f x') (f x) < e" using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x d) fine p" using fine_division_exists[OF gauge_ball[OF \0 < d\], of a b] . have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" proof (safe, unfold snd_conv) fix x l assume as: "(x, l) \ p" obtain a b where l: "l = cbox a b" using as ptag by blast then have x: "x \ cbox a b" using as ptag by auto show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" apply (rule_tac x="\y. f x" in exI) proof safe show "(\y. f x) integrable_on l" unfolding integrable_on_def l by blast next fix y assume y: "y \ l" then have "y \ ball x d" using as finep by fastforce then show "norm (f y - f x) \ e" using d x y as l by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) qed qed from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by metis qed lemma integrable_continuous_interval: fixes f :: "'b::ordered_euclidean_space \ 'a::banach" assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" by (metis assms integrable_continuous interval_cbox) lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] lemma integrable_continuous_closed_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" using assms by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl) subsection \Specialization of additivity to one dimension\ subsection \A useful lemma allowing us to factor out the content size\ lemma has_integral_factor_content: "(f has_integral i) (cbox a b) \ (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" proof (cases "content (cbox a b) = 0") case True have "\e p. p tagged_division_of cbox a b \ norm ((\(x, k)\p. content k *\<^sub>R f x)) \ e * content (cbox a b)" unfolding sum_content_null[OF True] True by force moreover have "i = 0" if "\e. e > 0 \ \d. gauge d \ (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) \ e * content (cbox a b))" using that [of 1] by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b]) ultimately show ?thesis unfolding has_integral_null_eq[OF True] by (force simp add: ) next case False then have F: "0 < content (cbox a b)" using zero_less_measure_iff by blast let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis apply (subst has_integral) proof safe fix e :: real assume e: "e > 0" { assume "\e>0. ?P e (<)" then show "?P (e * content (cbox a b)) (\)" apply (rule allE [where x="e * content (cbox a b)"]) apply (elim impE ex_forward conj_forward) using F e apply (auto simp add: algebra_simps) done } { assume "\e>0. ?P (e * content (cbox a b)) (\)" then show "?P e (<)" apply (rule allE [where x="e/2 / content (cbox a b)"]) apply (elim impE ex_forward conj_forward) using F e apply (auto simp add: algebra_simps) done } qed qed lemma has_integral_factor_content_real: "(f has_integral i) {a..b::real} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b} ))" unfolding box_real[symmetric] by (rule has_integral_factor_content) subsection \Fundamental theorem of calculus\ lemma interval_bounds_real: fixes q b :: real assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" using assms by auto theorem fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" and vecd: "\x. x \ {a..b} \ (f has_vector_derivative f' x) (at x within {a..b})" shows "(f' has_integral (f b - f a)) {a..b}" unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume "e > 0" then have "\x. \d>0. x \ {a..b} \ (\y\{a..b}. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x))" using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast then obtain d where d: "\x. 0 < d x" "\x y. \x \ {a..b}; y \ {a..b}; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x)" by metis show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" proof (rule exI, safe) show "gauge (\x. ball x (d x))" using d(1) gauge_ball_dependent by blast next fix p assume ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x (d x)) fine p" have ba: "b - a = (\(x,K)\p. Sup K - Inf K)" "f b - f a = (\(x,K)\p. f(Sup K) - f(Inf K))" using additive_tagged_division_1[where f= "\x. x"] additive_tagged_division_1[where f= f] \a \ b\ ptag by auto have "norm (\(x, K) \ p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, k) \ Sup k - Inf k))" proof (rule sum_norm_le,safe) fix x K assume "(x, K) \ p" then have "x \ K" and kab: "K \ cbox a b" using ptag by blast+ then obtain u v where k: "K = cbox u v" and "x \ K" and kab: "K \ cbox a b" using ptag \(x, K) \ p\ by auto have "u \ v" using \x \ K\ unfolding k by auto have ball: "\y\K. y \ ball x (d x)" using finep \(x, K) \ p\ by blast have "u \ K" "v \ K" by (simp_all add: \u \ v\ k) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" by (auto simp add: algebra_simps) also have "... \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" by (rule norm_triangle_ineq4) finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . also have "\ \ e * norm (u - x) + e * norm (v - x)" proof (rule add_mono) show "norm (f u - f x - (u - x) *\<^sub>R f' x) \ e * norm (u - x)" apply (rule d(2)[of x u]) using \x \ K\ kab \u \ K\ ball dist_real_def by (auto simp add:dist_real_def) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" apply (rule d(2)[of x v]) using \x \ K\ kab \v \ K\ ball dist_real_def by (auto simp add:dist_real_def) qed also have "\ \ e * (Sup K - Inf K)" using \x \ K\ by (auto simp: k interval_bounds_real[OF \u \ v\] field_simps) finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (Sup K - Inf K)" using \u \ v\ by (simp add: k) qed with \a \ b\ show "norm ((\(x, K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) qed qed lemma ident_has_integral: fixes a::real assumes "a \ b" shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}" proof - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" apply (rule fundamental_theorem_of_calculus [OF assms]) unfolding power2_eq_square by (rule derivative_eq_intros | simp)+ then show ?thesis by (simp add: field_simps) qed lemma integral_ident [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2)/2 else 0)" by (metis assms ident_has_integral integral_unique) lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) lemma integral_sin [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" proof - have "(sin has_integral (- cos b - - cos a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) subsection \Taylor series expansion\ lemma mvt_integral: fixes f::"'a::real_normed_vector\'b::banach" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" (is ?th1) proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = has_derivative_subset[OF _ subset] has_derivative_in_compose[where f="(\xa. x + xa *\<^sub>R y)" and g = f] note [continuous_intros] = continuous_on_compose2[where f="(\xa. x + xa *\<^sub>R y)"] continuous_on_subset[OF _ subset] have "\t. t \ {0..1} \ ((\t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y) (at t within {0..1})" using assms by (auto simp: has_vector_derivative_def linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] show ?th1 by (auto intro!: integral_unique[symmetric]) qed lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and g0: "Dg 0 = g" and Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" and ivl: "a \ t" "t \ b" shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) has_vector_derivative prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) (at t within {a..b})" using assms proof cases assume p: "p \ 1" define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" by auto let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = (\i\(Suc p'). ?f i - ?f (Suc i))" by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) also note sum_telescope finally have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" unfolding p'[symmetric] by (simp add: assms) thus ?thesis using assms by (auto intro!: derivative_eq_intros has_vector_derivative) qed (auto intro!: derivative_eq_intros has_vector_derivative) lemma fixes f::"real\'a::banach" assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and ivl: "a \ b" defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" shows Taylor_has_integral: "(i has_integral f b - (\iR Df i a)) {a..b}" and Taylor_integral: "f b = (\iR Df i a) + integral {a..b} i" and Taylor_integrable: "i integrable_on {a..b}" proof goal_cases case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s define Dg where [abs_def]: "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm) { fix m assume "p > Suc m" hence "p - Suc m = Suc (p - Suc (Suc m))" by auto hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" by auto } note fact_eq = this have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" unfolding Dg_def by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] have deriv: "\t. a \ t \ t \ b \ (?sum has_vector_derivative g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" by auto from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] have "(i has_integral ?sum b - ?sum a) {a..b}" using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" and "{.. {i. p = Suc i} = {p - 1}" for p' using \p > 0\ by (auto simp: power_mult_distrib[symmetric]) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) finally show c: ?case . case 2 show ?case using c integral_unique by (metis (lifting) add.commute diff_eq_eq integral_unique) case 3 show ?case using c by force qed subsection \Only need trivial subintervals if the interval itself is trivial\ proposition division_of_nontrivial: fixes \ :: "'a::euclidean_space set set" assumes sdiv: "\ division_of (cbox a b)" and cont0: "content (cbox a b) \ 0" shows "{k. k \ \ \ content k \ 0} division_of (cbox a b)" using sdiv proof (induction "card \" arbitrary: \ rule: less_induct) case less note \ = division_ofD[OF less.prems] { presume *: "{k \ \. content k \ 0} \ \ \ ?case" then show ?case using less.prems by fastforce } assume noteq: "{k \ \. content k \ 0} \ \" then obtain K c d where "K \ \" and contk: "content K = 0" and keq: "K = cbox c d" using \(4) by blast then have "card \ > 0" unfolding card_gt_0_iff using less by auto then have card: "card (\ - {K}) < card \" using less \K \ \\ by (simp add: \(1)) have closed: "closed (\(\ - {K}))" using less.prems by auto have "x islimpt \(\ - {K})" if "x \ K" for x unfolding islimpt_approachable proof (intro allI impI) fix e::real assume "e > 0" obtain i where i: "c\i = d\i" "i\Basis" using contk \(3) [OF \K \ \\] unfolding box_ne_empty keq by (meson content_eq_0 dual_order.antisym) then have xi: "x\i = d\i" using \x \ K\ unfolding keq mem_box by (metis antisym) define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i)/2 then c\i + min e (b\i - c\i)/2 else c\i - min e (c\i - a\i)/2 else x\j) *\<^sub>R j)" show "\x'\\(\ - {K}). x' \ x \ dist x' x < e" proof (intro bexI conjI) have "d \ cbox c d" using \(3)[OF \K \ \\] by (simp add: box_ne_empty(1) keq mem_box(2)) then have "d \ cbox a b" using \(2)[OF \K \ \\] by (auto simp: keq) then have di: "a \ i \ d \ i \ d \ i \ b \ i" using \i \ Basis\ mem_box(2) by blast then have xyi: "y\i \ x\i" unfolding y_def i xi using \e > 0\ cont0 \i \ Basis\ by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) then show "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto have "norm (y-x) \ (\b\Basis. \(y - x) \ b\)" by (rule norm_le_l1) also have "... = \(y - x) \ i\ + (\b \ Basis - {i}. \(y - x) \ b\)" by (meson finite_Basis i(2) sum.remove) also have "... < e + sum (\i. 0) Basis" proof (rule add_less_le_mono) show "\(y-x) \ i\ < e" using di \e > 0\ y_def i xi by (auto simp: inner_simps) show "(\i\Basis - {i}. \(y-x) \ i\) \ (\i\Basis. 0)" unfolding y_def by (auto simp: inner_simps) qed finally have "norm (y-x) < e + sum (\i. 0) Basis" . then show "dist y x < e" unfolding dist_norm by auto have "y \ K" unfolding keq mem_box using i(1) i(2) xi xyi by fastforce moreover have "y \ \\" using subsetD[OF \(2)[OF \K \ \\] \x \ K\] \e > 0\ di i by (auto simp: \ mem_box y_def field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(\ - {K})" by auto qed qed then have "K \ \(\ - {K})" using closed closed_limpt by blast then have "\(\ - {K}) = cbox a b" unfolding \(6)[symmetric] by auto then have "\ - {K} division_of cbox a b" by (metis Diff_subset less.prems division_of_subset \(6)) then have "{ka \ \ - {K}. content ka \ 0} division_of (cbox a b)" using card less.hyps by blast moreover have "{ka \ \ - {K}. content ka \ 0} = {K \ \. content K \ 0}" using contk by auto ultimately show ?case by auto qed subsection \Integrability on subintervals\ lemma operative_integrableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. f integrable_on i)" proof - interpret comm_monoid conj True by standard auto have 1: "\a b. box a b = {} \ f integrable_on cbox a b" by (simp add: content_eq_0_interior integrable_on_null) have 2: "\a b c k. \k \ Basis; f integrable_on cbox a b \ {x. x \ k \ c}; f integrable_on cbox a b \ {x. c \ x \ k}\ \ f integrable_on cbox a b" unfolding integrable_on_def by (auto intro!: has_integral_split) show ?thesis apply standard using 1 2 by blast+ qed lemma integrable_subinterval: fixes f :: "'b::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and cd: "cbox c d \ cbox a b" shows "f integrable_on cbox c d" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis proof (cases "cbox c d = {}") case True then show ?thesis using division [symmetric] f by (auto simp: comm_monoid_set_F_and) next case False then show ?thesis by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1) qed qed lemma integrable_subinterval_real: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "{c..d} \ {a..b}" shows "f integrable_on {c..d}" by (metis assms box_real(2) integrable_subinterval) subsection \Combining adjacent intervals in 1 dimension\ lemma has_integral_combine: fixes a b c :: real and j :: "'a::banach" assumes "a \ c" and "c \ b" and ac: "(f has_integral i) {a..c}" and cb: "(f has_integral j) {c..b}" shows "(f has_integral (i + j)) {a..b}" proof - interpret operative_real "lift_option plus" "Some 0" "\i. if f integrable_on i then Some (integral i f) else None" using operative_integralI by (rule operative_realI) from \a \ c\ \c \ b\ ac cb coalesce_less_eq have *: "lift_option (+) (if f integrable_on {a..c} then Some (integral {a..c} f) else None) (if f integrable_on {c..b} then Some (integral {c..b} f) else None) = (if f integrable_on {a..b} then Some (integral {a..b} f) else None)" by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed lemma integral_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and ab: "f integrable_on {a..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" proof - have "(f has_integral integral {a..c} f) {a..c}" using ab \c \ b\ integrable_subinterval_real by fastforce moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}" using \a \ c\ \c \ b\ has_integral_combine by blast then show ?thesis by (simp add: has_integral_integrable_integral) qed lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by (auto intro!: has_integral_combine) lemma integral_minus_sets: fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {c .. max a b} \ integral {c .. a} f - integral {c .. b} f = (if a \ b then - integral {a .. b} f else integral {b .. a} f)" using integral_combine[of c a b f] integral_combine[of c b a f] by (auto simp: algebra_simps max_def) lemma integral_minus_sets': fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {min a b .. c} \ integral {a .. c} f - integral {b .. c} f = (if a \ b then integral {a .. b} f else - integral {b .. a} f)" using integral_combine[of b a c f] integral_combine[of a b c f] by (auto simp: algebra_simps min_def) subsection \Reduce integrability to "local" integrability\ lemma integrable_on_little_subintervals: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v" shows "f integrable_on cbox a b" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) have "\x. \d>0. x\cbox a b \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v)" using assms by (metis zero_less_one) then obtain d where d: "\x. 0 < d x" "\x u v. \x \ cbox a b; x \ cbox u v; cbox u v \ ball x (d x); cbox u v \ cbox a b\ \ f integrable_on cbox u v" by metis obtain p where p: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast then have sndp: "snd ` p division_of cbox a b" by (metis division_of_tagged_division) have "f integrable_on k" if "(x, k) \ p" for x k using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto then show ?thesis unfolding division [symmetric, OF sndp] comm_monoid_set_F_and by auto qed subsection \Second FTC or existence of antiderivative\ lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def by blast lemma integral_has_vector_derivative_continuous_at: fixes f :: "real \ 'a::banach" assumes f: "f integrable_on {a..b}" and x: "x \ {a..b} - S" and "finite S" and fx: "continuous (at x within ({a..b} - S)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - S; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" if y: "y \ {a..b} - S" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" x y] False apply simp done have "\xa. y - x < d \ (\x'. a \ x' \ x' \ b \ x' \ S \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ xa \ S \ a \ x \ x \ S \ y \ b \ y \ S \ x \ xa \ xa \ y \ norm (f xa - f x) \ e" using assms by auto show ?thesis using False apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) using yx False d x y \e>0\ assms by (auto simp: Idiff fux_int) next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" y x] True apply simp done have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" using True apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) using yx True d x y \e>0\ assms by (auto simp: Idiff fux_int) then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - S. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})" using integral_has_vector_derivative[of a b g t] assms by (auto simp: has_field_derivative_iff_has_vector_derivative) lemma antiderivative_continuous: fixes q b :: real assumes "continuous_on {a..b} f" obtains g where "\x. x \ {a..b} \ (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" using integral_has_vector_derivative[OF assms] by auto subsection \Combined fundamental theorem of calculus\ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" proof - obtain g where g: "\x. x \ {a..b} \ (g has_vector_derivative f x) (at x within {a..b})" using antiderivative_continuous[OF assms] by metis have "(f has_integral g v - g u) {u..v}" if "u \ {a..b}" "v \ {a..b}" "u \ v" for u v proof - have "\x. x \ cbox u v \ (g has_vector_derivative f x) (at x within cbox u v)" by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g has_vector_derivative_within_subset subsetCE that(1,2)) then show ?thesis by (metis box_real(2) that(3) fundamental_theorem_of_calculus) qed then show ?thesis using that by blast qed subsection \General "twiddling" for interval-to-interval function image\ lemma has_integral_twiddle: assumes "0 < r" and hg: "\x. h(g x) = x" and gh: "\x. g(h x) = x" and contg: "\x. continuous (at x) g" and g: "\u v. \w z. g ` cbox u v = cbox w z" and h: "\u v. \w z. h ` cbox u v = cbox w z" and r: "\u v. content(g ` cbox u v) = r * content (cbox u v)" and intfi: "(f has_integral i) (cbox a b)" shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof (cases "cbox a b = {}") case True then show ?thesis using intfi by auto next case False obtain w z where wz: "h ` cbox a b = cbox w z" using h by blast have inj: "inj g" "inj h" using hg gh injI by metis+ from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast have "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" if "e > 0" for e proof - have "e * r > 0" using that \0 < r\ by simp with intfi[unfolded has_integral] obtain d where "gauge d" and d: "\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e * r" by metis define d' where "d' x = g -` d (g x)" for x show ?thesis proof (rule_tac x=d' in exI, safe) show "gauge d'" using \gauge d\ continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) next fix p assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" note p = tagged_division_ofD[OF ptag] have gab: "g y \ cbox a b" if "y \ K" "(x, K) \ p" for x y K by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) have gimp: "(\(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using ptag by auto show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next fix x k assume xk: "(x, k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto show "\u v. g ` k = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u assume xk': "(x', K') \ p" and u: "u \ interior (g ` k)" "u \ interior (g ` K')" have "interior k \ interior K' \ {}" proof assume "interior k \ interior K' = {}" moreover have "u \ g ` (interior k \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed then have same: "(x, k) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto show "g u \ g ` K'"if "u \ k" for u using that same by auto show "g u \ g ` k"if "u \ K'" for u using that same by auto next fix x assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then obtain X y where "h x \ X" "(y, X) \ p" by blast then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" apply clarsimp by (metis (no_types, lifting) assms(3) image_eqI pair_imageI) qed (use gab in auto) have *: "inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastforce have "(\(x, k)\(\(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\(x, k)\p. content k *\<^sub>R f (g x)) - i" (is "?l = _") using r apply (simp only: algebra_simps add_left_cancel scaleR_right.sum) apply (subst sum.reindex_bij_betw[symmetric, where h="\(x, k). (g x, g ` k)" and S=p]) apply (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) done also have "\ = r *\<^sub>R ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") using \0 < r\ by (auto simp: scaleR_diff_right) finally have eq: "?l = ?r" . show "norm ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using d[OF gimp] \0 < r\ by (auto simp add: eq) qed qed then show ?thesis by (auto simp: h_eq has_integral) qed subsection \Special case of a basic affine transformation\ lemma AE_lborel_inner_neq: assumes k: "k \ Basis" shows "AE x in lborel. x \ k \ c" proof - interpret finite_product_sigma_finite "\_. lborel" Basis proof qed simp have "emeasure lborel {x\space lborel. x \ k = c} = emeasure (\\<^sub>M j::'a\Basis. lborel) (\\<^sub>E j\Basis. if j = k then {c} else UNIV)" using k by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) also have "\ = (\j\Basis. emeasure lborel (if j = k then {c} else UNIV))" by (intro measure_times) auto also have "\ = 0" by (intro prod_zero bexI[OF _ k]) auto finally show ?thesis by (subst AE_iff_measurable[OF _ refl]) auto qed lemma content_image_stretch_interval: fixes m :: "'a::euclidean_space \ real" defines "s f x \ (\k::'a\Basis. (f k * (x\k)) *\<^sub>R k)" shows "content (s m ` cbox a b) = \\k\Basis. m k\ * content (cbox a b)" proof cases have s[measurable]: "s f \ borel \\<^sub>M borel" for f by (auto simp: s_def[abs_def]) assume m: "\k\Basis. m k \ 0" then have s_comp_s: "s (\k. 1 / m k) \ s m = id" "s m \ s (\k. 1 / m k) = id" by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) then have "inv (s (\k. 1 / m k)) = s m" "bij (s (\k. 1 / m k))" by (auto intro: inv_unique_comp o_bij) then have eq: "s m ` cbox a b = s (\k. 1 / m k) -` cbox a b" using bij_vimage_eq_inv_image[OF \bij (s (\k. 1 / m k))\, of "cbox a b"] by auto show ?thesis using m unfolding eq measure_def by (subst lborel_affine_euclidean[where c=m and t=0]) (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) next assume "\ (\k\Basis. m k \ 0)" then obtain k where k: "k \ Basis" "m k = 0" by auto then have [simp]: "(\k\Basis. m k) = 0" by (intro prod_zero) auto have "emeasure lborel {x\space lborel. x \ s m ` cbox a b} = 0" proof (rule emeasure_eq_0_AE) show "AE x in lborel. x \ s m ` cbox a b" using AE_lborel_inner_neq[OF \k\Basis\] proof eventually_elim show "x \ k \ 0 \ x \ s m ` cbox a b " for x using k by (auto simp: s_def[abs_def] cbox_def) qed qed then show ?thesis by (simp add: measure_def) qed lemma interval_image_affinity_interval: "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" unfolding image_affinity_cbox by auto lemma content_image_affinity_cbox: "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "m \ 0") case True with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps mult_left_mono) done moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps mult_left_mono) done moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False by (simp add: image_affinity_cbox content_cbox' prod.distrib[symmetric] inner_diff_left flip: prod_constant) qed qed lemma has_integral_affinity: fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (\m\ ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" apply (rule has_integral_twiddle) using assms apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox) apply (rule zero_less_power) unfolding scaleR_right_distrib apply auto done lemma integrable_affinity: assumes "f integrable_on cbox a b" and "m \ 0" shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using assms unfolding integrable_on_def apply safe apply (drule has_integral_affinity) apply auto done lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral ((1/ \prod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms by auto lemma has_integral_stretch_real: fixes f :: "real \ 'b::real_normed_vector" assumes "(f has_integral i) {a..b}" and "m \ 0" shows "((\x. f (m * x)) has_integral (1 / \m\) *\<^sub>R i) ((\x. x / m) ` {a..b})" using has_integral_stretch [of f i a b "\b. m"] assms by simp lemma integrable_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def by (force dest: has_integral_stretch) lemma vec_lambda_eq_sum: shows "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" apply (simp add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) done lemma has_integral_stretch_cart: fixes m :: "'n::finite \ real" assumes f: "(f has_integral i) (cbox a b)" and m: "\k. m k \ 0" shows "((\x. f(\ k. m k * x$k)) has_integral i /\<^sub>R \prod m UNIV\) ((\x. \ k. x$k / m k) ` (cbox a b))" proof - have *: "\k:: real^'n \ Basis. m (axis_index k) \ 0" using axis_index by (simp add: m) have eqp: "(\k:: real^'n \ Basis. m (axis_index k)) = prod m UNIV" by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def) show ?thesis using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\i x. m i * x"] vec_lambda_eq_sum [where f="\i x. x / m i"] by (simp add: field_simps eqp) qed lemma image_stretch_interval_cart: fixes m :: "'n::finite \ real" shows "(\x. \ k. m k * x$k) ` cbox a b = (if cbox a b = {} then {} else cbox (\ k. min (m k * a$k) (m k * b$k)) (\ k. max (m k * a$k) (m k * b$k)))" proof - have *: "(\k\Basis. min (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. min (m k * a $ k) (m k * b $ k))" "(\k\Basis. max (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. max (m k * a $ k) (m k * b $ k))" apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong) done show ?thesis by (simp add: vec_lambda_eq_sum [where f="\i x. m i * x"] image_stretch_interval eq_cbox *) qed subsection \even more special cases\ lemma uminus_interval_vector[simp]: fixes a b :: "'a::euclidean_space" shows "uminus ` cbox a b = cbox (-b) (-a)" apply safe apply (simp add: mem_box(2)) apply (rule_tac x="-x" in image_eqI) apply (auto simp add: mem_box) done lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) (cbox a b)" shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" using assms unfolding box_real[symmetric] by (rule has_integral_reflect_lemma) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" by (auto dest: has_integral_reflect_lemma) lemma has_integral_reflect_real[simp]: fixes a b::real shows "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" by (metis has_integral_reflect interval_cbox) lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule integrable_reflect) lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a..b::real} f" unfolding box_real[symmetric] by (rule integral_reflect) subsection \Stronger form of FCT; quite a tedious proof\ lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" by (simp add: split_def) theorem fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and contf: "continuous_on {a..b} f" and derf: "\x. x \ {a <..< b} \ (f has_vector_derivative f' x) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" proof (cases "a = b") case True then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) with True show ?thesis by auto next case False with \a \ b\ have ab: "a < b" by arith show ?thesis unfolding has_integral_factor_content_real proof (intro allI impI) fix e :: real assume e: "e > 0" then have eba8: "(e * (b-a)) / 8 > 0" using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" using derf_exp by simp have "\x \ box a b. \d>0. \y. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" (is "\x \ box a b. ?Q x") proof fix x assume x: "x \ box a b" show "?Q x" apply (rule allE [where x="e/2", OF derf_exp [THEN conjunct2, of x]]) using x e by auto qed from this [unfolded bgauge_existence_lemma] obtain d where d: "\x. 0 < d x" "\x y. \x \ box a b; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have "bounded (f ` cbox a b)" using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" unfolding bounded_pos by metis obtain da where "0 < da" and da: "\c. \a \ c; {a..c} \ {a..b}; {a..c} \ ball a da\ \ norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b-a)) / 4" proof - have "continuous (at a within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm (x-a); norm (x-a) < k\ \ norm (f x - f a) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b-a) / 8" proof (cases "f' a = 0") case True with ab e that show ?thesis by auto next case False then show ?thesis apply (rule_tac l="(e * (b-a)) / 8 / norm (f' a)" in that) using ab e apply (auto simp add: field_simps) done qed have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" if "a \ c" "{a..c} \ {a..b}" and bmin: "{a..c} \ ball a (min k l)" for c proof - have minkl: "\a - x\ < min k l" if "x \ {a..c}" for x using bmin dist_real_def that by auto then have lel: "\c - a\ \ \l\" using that by force have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" . next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") case True then show ?thesis using eba8 by auto next case False show ?thesis by (rule k) (use minkl \a \ c\ that False in auto) qed then show "norm (f c - f a) \ e * (b-a) / 8" by simp qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" unfolding content_real[OF \a \ c\] by auto qed then show ?thesis by (rule_tac da="min k l" in that) (auto simp: l \0 < k\) qed obtain db where "0 < db" and db: "\c. \c \ b; {c..b} \ {a..b}; {c..b} \ ball b db\ \ norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b-a)) / 4" proof - have "continuous (at b within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm(x-b); norm(x-b) < k\ \ norm (f b - f x) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \ (e * (b-a)) / 8" proof (cases "f' b = 0") case True thus ?thesis using ab e that by auto next case False then show ?thesis apply (rule_tac l="(e * (b-a)) / 8 / norm (f' b)" in that) using ab e by (auto simp add: field_simps) qed have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" if "c \ b" "{c..b} \ {a..b}" and bmin: "{c..b} \ ball b (min k l)" for c proof - have minkl: "\b - x\ < min k l" if "x \ {c..b}" for x using bmin dist_real_def that by auto then have lel: "\b - c\ \ \l\" using that by force have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((b - c) *\<^sub>R f' b) \ norm (l *\<^sub>R f' b)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((b - c) *\<^sub>R f' b) \ e * (b-a) / 8" . next have "norm (f b - f c) < e * (b-a) / 8" proof (cases "b = c") case True with eba8 show ?thesis by auto next case False show ?thesis by (rule k) (use minkl \c \ b\ that False in auto) qed then show "norm (f b - f c) \ e * (b-a) / 8" by simp qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" unfolding content_real[OF \c \ b\] by auto qed then show ?thesis by (rule_tac db="min k l" in that) (auto simp: l \0 < k\) qed let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" proof (rule exI, safe) show "gauge ?d" using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next fix p assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF ptag] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using ptag fine by auto have le_xz: "\w x y z::real. y \ z/2 \ w - x \ z/2 \ w + y \ x + z" by arith have non: False if xk: "(x,K) \ p" and "x \ a" "x \ b" and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" for x K proof - obtain u v where k: "K = cbox u v" using p(4) xk by blast then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF xk] by auto then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using less[unfolded k box_real interval_bounds_real content_real] by auto then have "x \ box a b" using p(2) p(3) \x \ a\ \x \ b\ xk by fastforce with d have *: "\y. norm (y-x) < d x \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have xd: "norm (u - x) < d x" "norm (v - x) < d x" using fineD[OF fine xk] \x \ a\ \x \ b\ uv by (auto simp add: k subset_eq dist_commute dist_real_def) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) also have "\ \ e/2 * norm (u - x) + e/2 * norm (v - x)" by (metis norm_triangle_le_diff add_mono * xd) also have "\ \ e/2 * norm (v - u)" using p(2)[OF xk] by (auto simp add: field_simps k) also have "\ < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using result by (simp add: \u \ v\) finally have "e * (v - u)/2 < e * (v - u)/2" using uv by auto then show False by auto qed have "norm (\(x, K)\p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\(x, K)\p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" proof - have ge0: "0 \ e * (Sup k - Inf k)" if xkp: "(x, k) \ p \ ?A" for x k proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) with p(2) that uv have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) qed let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" let ?C = "{t \ p. fst t \ {a, b} \ content (snd t) \ 0}" have "norm (\(x, k)\p \ {t. fst t \ {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" proof - have *: "\S f e. sum f S = sum f (p \ ?C) \ norm (sum f (p \ ?C)) \ e \ norm (sum f S) \ e" by auto have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" if "(x,K) \ p \ {t. fst t \ {a, b}} - p \ ?C" for x K proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto then obtain u v where uv: "K = cbox u v" using p(4) by blast then have "u = v" using xk k0 p(2) by force then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed have 2: "norm(\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b-a)/2" proof - have norm_le: "norm (sum f S) \ e" if "\x y. \x \ S; y \ S\ \ x = y" "\x. x \ S \ norm (f x) \ e" "e > 0" for S f and e :: real proof (cases "S = {}") case True with that show ?thesis by auto next case False then obtain x where "x \ S" by auto then have "S = {x}" using that(1) by auto then show ?thesis using \x \ S\ that(2) by auto qed have *: "p \ ?C = ?B a \ ?B b" by blast then have "norm (\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = norm (\(x,K) \ ?B a \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" by simp also have "... = norm ((\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" apply (subst sum.union_disjoint) using p(1) ab e by auto also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(a, K) \ p" "(a, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pa obtain v v' where v: "K = cbox a v" "a \ v" and v': "K' = cbox a v'" "a \ v'" by blast let ?v = "min v v'" have "box a ?v \ K \ K'" unfolding v v' by (auto simp add: mem_box) then have "interior (box a (min v v')) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "(a + ?v)/2 \ box a ?v" using ne0 unfolding v v' content_eq_0 not_le by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(a, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(a, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox a v" and "a \ v" using pa[OF \(a, c) \ p\] by metis then have "a \ {a..v}" "v \ b" using p(3)[OF \(a, c) \ p\] by auto moreover have "{a..v} \ ball a da" using fineD[OF \?d fine p\ \(a, c) \ p\] by (simp add: v split: if_split_asm) ultimately show ?thesis unfolding v interval_bounds_real[OF \a \ v\] box_real using da \a \ v\ by auto qed qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(b, K) \ p" "(b, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pb obtain v v' where v: "K = cbox v b" "v \ b" and v': "K' = cbox v' b" "v' \ b" by blast let ?v = "max v v'" have "box ?v b \ K \ K'" unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have " ((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(b, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(b, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox v b" and "v \ b" using \(b, c) \ p\ pb by blast then have "v \ a""b \ {v.. b}" using p(3)[OF \(b, c) \ p\] by auto moreover have "{v..b} \ ball b db" using fineD[OF \?d fine p\ \(b, c) \ p\] box_real(2) v False by force ultimately show ?thesis using db v by auto qed qed (use ab e in auto) qed also have "... = e * (b-a)/2" by simp finally show "norm (\(x,k)\p \ ?C. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" . qed show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b-a)/2" apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) using 1 2 by (auto simp: split_paired_all) qed also have "... = (\n\p. e * (case n of (x, k) \ Sup k - Inf k))/2" unfolding sum_distrib_left[symmetric] apply (subst additive_tagged_division_1[OF \a \ b\ ptag]) by auto finally have norm_le: "norm (\(x,K)\p \ {t. fst t \ {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, K) \ Sup K - Inf K))/2" . have le2: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2)/2 \ x - s1 \ s2/2" by auto show ?thesis apply (rule le2 [OF sum_nonneg]) using ge0 apply force unfolding sum.union_disjoint[OF pA(2-), symmetric] pA(1)[symmetric] by (metis norm_le) qed note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] have "norm (\(x,K)\p \ ?A \ (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (\(x,K)\p \ ?A \ (p - ?A). Sup K - Inf K)" unfolding sum_distrib_left unfolding sum.union_disjoint[OF pA(2-)] using le_xz norm_triangle_le I II by blast then show "norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" by (simp only: content_real[OF \a \ b\] *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) qed qed qed subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" using \a < x\ \x < b\ has_integral_combine less_imp_le by blast then show ?thesis by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" shows "(f' has_integral (f b - f a)) {a..b}" by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" assumes intf: "f integrable_on {a..b}" and "a < c" "c \ b" "e > 0" obtains d where "d > 0" and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" proof - obtain w where "w > 0" and w: "\t. \c - w < t; t < c\ \ norm (f c) * norm(c - t) < e/3" proof (cases "f c = 0") case False hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t proof - have "norm (c - t) < e/3 / norm (f c)" using that by auto then show "norm (f c) * norm (c - t) < e/3" by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) qed ultimately show ?thesis using that by auto next case True then show ?thesis using \e > 0\ that by auto qed let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" using \e > 0\ by auto have "f integrable_on {a..c}" apply (rule integrable_subinterval_real[OF intf]) using \a < c\ \c \ b\ by auto then obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of {a..c}; d1 fine p\ \ norm (?SUM p - integral {a..c} f) < e/3" using integrable_integral has_integral_real e3 by metis define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using \w > 0\ \gauge d1\ by auto then obtain k where "0 < k" and k: "ball c k \ d c" by (meson gauge_def open_contains_ball) let ?d = "min k (c - a)/2" show thesis proof (intro that[of ?d] allI impI, safe) show "?d > 0" using \0 < k\ \a < c\ by auto next fix t assume t: "c - ?d < t" "t \ c" show "norm (integral ({a..c}) f - integral ({a..t}) f) < e" proof (cases "t < c") case False with \t \ c\ show ?thesis by (simp add: \e > 0\) next case True have "f integrable_on {a..t}" apply (rule integrable_subinterval_real[OF intf]) using \t < c\ \c \ b\ by auto then obtain d2 where d2: "gauge d2" "\p. p tagged_division_of {a..t} \ d2 fine p \ norm (?SUM p - integral {a..t} f) < e/3" using integrable_integral has_integral_real e3 by metis define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using \gauge d1\ \gauge d2\ unfolding d3_def gauge_def by auto then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p" by (metis box_real(2) fine_division_exists) note p' = tagged_division_ofD[OF ptag] have pt: "(x,K)\p \ x \ t" for x K by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) with pfine have "d2 fine p" unfolding fine_def d3_def by fastforce then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3" using d2(2) ptag by auto have eqs: "{a..c} \ {x. x \ t} = {a..t}" "{a..c} \ {x. x \ t} = {t..c}" using t by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" apply (rule tagged_division_Un_interval_real[of _ _ _ 1 "t"]) using \t \ c\ by (auto simp: eqs ptag tagged_division_of_self_real) moreover have "d1 fine p \ {(c, {t..c})}" unfolding fine_def proof safe fix x K y assume "(x,K) \ p" and "y \ K" then show "y \ d1 x" by (metis Int_iff d3_def subsetD fineD pfine) next fix x assume "x \ {t..c}" then have "dist c x < k" using t(1) by (auto simp add: field_simps dist_real_def) with k show "x \ d1 c" unfolding d_def by auto qed ultimately have d1_fin: "norm (?SUM(p \ {(c, {t..c})}) - integral {a..c} f) < e/3" using d1 by metis have SUMEQ: "?SUM(p \ {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p" proof - have "?SUM(p \ {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p" proof (subst sum.union_disjoint) show "p \ {(c, {t..c})} = {}" using \t < c\ pt by force qed (use p'(1) in auto) also have "... = (c - t) *\<^sub>R f c + ?SUM p" using \t \ c\ by auto finally show ?thesis . qed have "c - k < t" using \k>0\ t(1) by (auto simp add: field_simps) moreover have "k \ w" proof (rule ccontr) assume "\ k \ w" then have "c + (k + w) / 2 \ d c" by (auto simp add: field_simps not_le not_less dist_real_def d_def) then have "c + (k + w) / 2 \ ball c k" using k by blast then show False using \0 < w\ \\ k \ w\ dist_real_def by auto qed ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq proof (intro norm_triangle_lt add_strict_mono) show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3" by (metis SUMEQ d1_fin norm_minus_cancel) show "norm (?SUM p - integral {a..t} f) < e/3" using d2_fin by blast show "norm ((c - t) *\<^sub>R f c) < e/3" using w cwt \t < c\ by simp (simp add: field_simps) qed then show ?thesis by simp qed qed qed lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" proof - have intm: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF intm \e>0\] obtain d where "0 < d" and d: "\t. \- c - d < t; t \ -c\ \ norm (integral {- b..- c} (\x. f (-x)) - integral {- b..t} (\x. f (-x))) < e" by metis let ?d = "min d (b - c)" show ?thesis proof (intro that[of "?d"] allI impI) show "0 < ?d" using \0 < d\ \c < b\ by auto fix t :: real assume t: "c \ t \ t < c + ?d" have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" "integral {a..t} f = integral {a..b} f - integral {t..b} f" apply (simp_all only: algebra_simps) using assms t by (auto simp: integral_combine) have "(- c) - d < (- t)" "- t \ - c" using t by auto from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e" by (auto simp add: algebra_simps norm_minus_commute *) qed qed lemma indefinite_integral_continuous_1: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {a..x} f)" proof - have "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" if x: "x \ {a..b}" and "e > 0" for x e :: real proof (cases "a = b") case True with that show ?thesis by force next case False with x have "a < b" by force with x consider "x = a" | "x = b" | "a < x" "x < b" by force then show ?thesis proof cases case 1 show ?thesis apply (rule indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\], force) using \x = a\ apply (force simp: dist_norm algebra_simps) done next case 2 show ?thesis apply (rule indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\], force) using \x = b\ apply (force simp: dist_norm norm_minus_commute algebra_simps) done next case 3 obtain d1 where "0 < d1" and d1: "\t. \x - d1 < t; t \ x\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \a < x\ _ \e > 0\]) obtain d2 where "0 < d2" and d2: "\t. \x \ t; t < x + d2\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \x < b\ \e > 0\]) show ?thesis proof (intro exI ballI conjI impI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" if "y \ {a..b}" "dist y x < min d1 d2" for y proof (cases "y < x") case True with that d1 show ?thesis by (auto simp: dist_commute dist_norm) next case False with that d2 show ?thesis by (auto simp: dist_commute dist_norm) qed qed qed qed then show ?thesis by (auto simp: continuous_on_iff) qed lemma indefinite_integral_continuous_1': fixes f::"real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {x..b} f)" proof - have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \ {a..b}" for x using integral_combine[OF _ _ assms, of x] that by (auto simp: algebra_simps) with _ show ?thesis by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) qed theorem integral_has_vector_derivative': fixes f :: "real \ 'b::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})" proof - have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \ x" "x \ b" for x using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]] by (simp add: algebra_simps) show ?thesis using \x \ _\ * by (rule has_vector_derivative_transform) (auto intro!: derivative_eq_intros assms integral_has_vector_derivative) qed lemma integral_has_real_derivative': assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})" using integral_has_vector_derivative'[OF assms] by (auto simp: has_field_derivative_iff_has_vector_derivative) subsection \This doesn't directly involve integration, but that gives an easy proof\ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" and contf: "continuous_on {a..b} f" and "f a = y" and fder: "\x. x \ {a..b} - k \ (f has_derivative (\h. 0)) (at x within {a..b})" and x: "x \ {a..b}" shows "f x = y" proof - have "a \ b" "a \ x" using assms by auto have "((\x. 0::'a) has_integral f x - f a) {a..x}" proof (rule fundamental_theorem_of_calculus_interior_strong[OF \finite k\ \a \ x\]; clarify?) have "{a..x} \ {a..b}" using x by auto then show "continuous_on {a..x} f" by (rule continuous_on_subset[OF contf]) show "(f has_vector_derivative 0) (at z)" if z: "z \ {a<.. k" for z unfolding has_vector_derivative_def proof (simp add: at_within_open[OF z, symmetric]) show "(f has_derivative (\x. 0)) (at z within {a<..Generalize a bit to any convex set\ lemma has_derivative_zero_unique_strong_convex: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex S" "finite K" and contf: "continuous_on S f" and "c \ S" "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof (cases "x = c") case True with \f c = y\ show ?thesis by blast next case False let ?\ = "\u. (1 - u) *\<^sub>R c + u *\<^sub>R x" have contf': "continuous_on {0 ..1} (f \ ?\)" proof (rule continuous_intros continuous_on_subset[OF contf])+ show "(\u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \ S" using \convex S\ \x \ S\ \c \ S\ by (auto simp add: convex_alt algebra_simps) qed have "t = u" if "?\ t = ?\ u" for t u proof - from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" by (auto simp add: algebra_simps) then show ?thesis using \x \ c\ by auto qed then have eq: "(SOME t. ?\ t = ?\ u) = u" for u by blast then have "(?\ -` K) \ (\z. SOME t. ?\ t = z) ` K" by (clarsimp simp: image_iff) (metis (no_types) eq) then have fin: "finite (?\ -` K)" by (rule finite_surj[OF \finite K\]) have derf': "((\u. f (?\ u)) has_derivative (\h. 0)) (at t within {0..1})" if "t \ {0..1} - {t. ?\ t \ K}" for t proof - have df: "(f has_derivative (\h. 0)) (at (?\ t) within ?\ ` {0..1})" apply (rule has_derivative_within_subset [OF derf]) using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps) have "(f \ ?\ has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" by (rule derivative_eq_intros df | simp)+ then show ?thesis unfolding o_def . qed have "(f \ ?\) 1 = y" apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) unfolding o_def using \f c = y\ derf' by auto then show ?thesis by auto qed text \Also to any open connected set with finite set of exceptions. Could generalize to locally convex set with limpt-free set of exceptions.\ lemma has_derivative_zero_unique_strong_connected: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite K" and contf: "continuous_on S f" and "c \ S" and "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof - have "\e>0. ball x e \ (S \ f -` {f x})" if "x \ S" for x proof - obtain e where "0 < e" and e: "ball x e \ S" using \x \ S\ \open S\ open_contains_ball by blast have "ball x e \ {u \ S. f u \ {f x}}" proof safe fix y assume y: "y \ ball x e" then show "y \ S" using e by auto show "f y = f x" proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \finite K\]) show "continuous_on (ball x e) f" using contf continuous_on_subset e by blast show "(f has_derivative (\h. 0)) (at u within ball x e)" if "u \ ball x e - K" for u by (metis Diff_iff contra_subsetD derf e has_derivative_within_subset that) qed (use y e \0 < e\ in auto) qed then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) then show ?thesis using \x \ S\ \f c = y\ \c \ S\ by auto qed lemma has_derivative_zero_connected_constant: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite k" and "continuous_on S f" and "\x\(S - k). (f has_derivative (\h. 0)) (at x within S)" obtains c where "\x. x \ S \ f(x) = c" proof (cases "S = {}") case True then show ?thesis by (metis empty_iff that) next case False then obtain c where "c \ S" by (metis equals0I) then show ?thesis by (metis has_derivative_zero_unique_strong_connected assms that) qed lemma DERIV_zero_connected_constant: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "connected S" and "open S" and "finite K" and "continuous_on S f" and "\x\(S - K). DERIV f x :> 0" obtains c where "\x. x \ S \ f(x) = c" using has_derivative_zero_connected_constant [OF assms(1-4)] assms by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) subsection \Integrating characteristic function of an interval\ lemma has_integral_restrict_open_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "(f has_integral i) (cbox c d)" and cb: "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof (cases "cbox c d = {}") case True then have "box c d = {}" by (metis bot.extremum_uniqueI box_subset_cbox) then show ?thesis using True intf by auto next case False then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \ p" using cb partial_division_extend_1 by blast define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x interpret operative "lift_option plus" "Some (0 :: 'b)" "\i. if g integrable_on i then Some (integral i g) else None" by (fact operative_integralI) note operat = division [OF pdiv, symmetric] show ?thesis proof (cases "(g has_integral i) (cbox a b)") case True then show ?thesis by (simp add: g_def) next case False have iterate:"F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" proof (intro neutral ballI) fix x assume x: "x \ p - {cbox c d}" then have "x \ p" by auto then obtain u v where uv: "x = cbox u v" using pdiv by blast have "interior x \ interior (cbox c d) = {}" using pdiv inp x by blast then have "(g has_integral 0) x" unfolding uv using has_integral_spike_interior[where f="\x. 0"] by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" by auto qed interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) have intg: "g integrable_on cbox c d" using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" proof (rule has_integral_unique[OF has_integral_spike_interior intf]) show "\x. x \ box c d \ f x = g x" by (auto simp: g_def) show "(g has_integral integral (cbox c d) g) (cbox c d)" by (rule integrable_integral[OF intg]) qed ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) then have "(g has_integral i) (cbox a b)" by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) with False show ?thesis by blast qed qed lemma has_integral_restrict_closed_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) (cbox c d)" and "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis by (rule *[of c d]) (use box_subset_cbox[of c d] in auto) qed lemma has_integral_restrict_closed_subintervals_eq: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") proof (cases "cbox c d = {}") case False let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis proof assume ?l then have "?g integrable_on cbox c d" using assms has_integral_integrable integrable_subinterval by blast then have "f integrable_on cbox c d" by (rule integrable_eq) auto moreover then have "i = integral (cbox c d) f" by (meson \((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)\ assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) ultimately show ?r by auto next assume ?r then show ?l by (rule has_integral_restrict_closed_subinterval[OF _ assms]) qed qed auto text \Hence we can apply the limit process uniformly to all integrals.\ lemma has_integral': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof (cases "\a b. S = cbox a b") case False then show ?thesis by (simp add: has_integral_alt) next case True then obtain a b where S: "S = cbox a b" by blast obtain B where " 0 < B" and B: "\x. x \ cbox a b \ norm x \ B" using bounded_cbox[unfolded bounded_pos] by blast show ?thesis proof safe fix e :: real assume ?l and "e > 0" have "((\x. if x \ S then f x else 0) has_integral i) (cbox c d)" if "ball 0 (B+1) \ cbox c d" for c d unfolding S using B that by (force intro: \?l\[unfolded S] has_integral_restrict_closed_subinterval) then show "?r e" apply (rule_tac x="B+1" in exI) using \B>0\ \e>0\ by force next assume as: "\e>0. ?r e" then obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b)" by (meson zero_less_one) define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" "i \ Basis" for x i using that and Basis_le_norm[OF \i\Basis\, of x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (meson B mem_box(2) subsetI) have "c \ i \ x \ i \ x \ i \ d \ i" if x: "norm (0 - x) < C" and i: "i \ Basis" for x i using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain y where y: "(f has_integral y) (cbox a b)" using c_d has_integral_restrict_closed_subintervals_eq S by blast have "y = i" proof (rule ccontr) assume "y \ i" then have "0 < norm (y - i)" by auto from as[rule_format,OF this] obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z-i) < norm (y-i)" by auto define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" for x i using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (simp add: B mem_box(2) subset_eq) have "c \ i \ x \ i \ x \ i \ d \ i" if "norm (0 - x) < C" and "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast moreover then have "z = y" by (blast intro: has_integral_unique[OF _ y]) ultimately show False by auto qed then show ?l using y by (auto simp: S) qed qed lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" assumes fg: "(f has_integral i) S" "(g has_integral j) S" and le: "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ fg, of 1] le by auto lemma integral_le: fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on S" and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" using has_integral_component_nonneg[of 1 f i S] unfolding o_def using assms by auto lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes f: "f integrable_on S" and 0: "\x. x \ S \ 0 \ f x" shows "0 \ integral S f" by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) text \Hence a general restriction property.\ lemma has_integral_restrict [simp]: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) S" proof - have *: "\x. (if x \ T then if x \ S then f x else 0 else 0) = (if x\S then f x else 0)" using assms by auto show ?thesis apply (subst(2) has_integral') apply (subst has_integral') apply (simp add: *) done qed corollary has_integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) (S \ T)" proof - have "((\x. if x \ T then if x \ S then f x else 0 else 0) has_integral i) UNIV = ((\x. if x \ S \ T then f x else 0) has_integral i) UNIV" by (rule has_integral_cong) auto then show ?thesis using has_integral_restrict_UNIV by fastforce qed lemma integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral T (\x. if x \ S then f x else 0) = integral (S \ T) f" by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) lemma integrable_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(\x. if x \ S then f x else 0) integrable_on T \ f integrable_on (S \ T)" using has_integral_restrict_Int by fastforce lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" proof - have "(\x. if x \ S then f x else 0) = (\x. if x \ T then f x else 0)" using assms by fastforce with f show ?thesis by (simp only: has_integral_restrict_UNIV [symmetric, of f]) qed lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\x. x \ S \ f x = 0" and "S \ t" shows "f integrable_on t" using assms unfolding integrable_on_def by (auto intro:has_integral_on_superset) lemma integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "integral UNIV (\x. if x \ S then f x else 0) = integral S f" by (simp add: integral_restrict_Int) lemma integrable_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and as: "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) then show ?thesis apply (rule has_integral_component_le[OF k]) using as by auto qed subsection\Integrals on set differences\ lemma has_integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes S: "(f has_integral i) S" and T: "(f has_integral j) T" and neg: "negligible (T - S)" shows "(f has_integral (i - j)) (S - T)" proof - show ?thesis unfolding has_integral_restrict_UNIV [symmetric, of f] proof (rule has_integral_spike [OF neg]) have eq: "(\x. (if x \ S then f x else 0) - (if x \ T then f x else 0)) = (\x. if x \ T - S then - f x else if x \ S - T then f x else 0)" by (force simp add: ) have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" using S T has_integral_restrict_UNIV by auto from has_integral_diff [OF this] show "((\x. if x \ T - S then - f x else if x \ S - T then f x else 0) has_integral i-j) UNIV" by (simp add: eq) qed force qed lemma integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)" shows "integral (S - T) f = integral S f - integral T f" by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral) lemma integrable_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)" shows "f integrable_on (S - T)" using has_integral_setdiff [OF assms] by (simp add: has_integral_iff) lemma negligible_setdiff [simp]: "T \ S \ negligible (T - S)" by (metis Diff_eq_empty_iff negligible_empty) lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume ?r show ?l unfolding negligible_def proof safe fix a b show "(indicator s has_integral 0) (cbox a b)" apply (rule has_integral_negligible[OF \?r\[rule_format,of a b]]) unfolding indicator_def apply auto done qed qed (simp add: negligible_Int) lemma negligible_translation: assumes "negligible S" shows "negligible ((+) c ` S)" proof - have inj: "inj ((+) c)" by simp show ?thesis using assms proof (clarsimp simp: negligible_def) fix a b assume "\x y. (indicator S has_integral 0) (cbox x y)" then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) have eq: "indicator ((+) c ` S) = (\x. indicator S (x - c))" by (force simp add: indicator_def) show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" using has_integral_affinity [OF *, of 1 "-c"] cbox_translation [of "c" "-c+a" "-c+b"] by (simp add: eq) (simp add: ac_simps) qed qed lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" shows "negligible S" by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) subsection \More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" apply (rule has_integral_subset_component_le) using assms apply auto done lemma integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x" shows "integral s f \ integral t f" apply (rule has_integral_subset_le) using assms apply auto done lemma has_integral_alt': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume rhs: ?r show ?l proof (subst has_integral', intro allI impI) fix e::real assume "e > 0" from rhs[THEN conjunct2,rule_format,OF this] show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e)" apply (rule ex_forward) using rhs by blast qed next let ?\ = "\e a b. \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e" assume ?l then have lhs: "\B>0. \a b. ball 0 B \ cbox a b \ ?\ e a b" if "e > 0" for e using that has_integral'[of f] by auto let ?f = "\x. if x \ s then f x else 0" show ?r proof (intro conjI allI impI) fix a b :: 'n from lhs[OF zero_less_one] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ ?\ 1 a b" by blast let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) have "?a \ i \ x \ i \ x \ i \ ?b \ i" if "norm (0 - x) < B" "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp add:field_simps) then have "ball 0 B \ cbox ?a ?b" by (auto simp: mem_box dist_norm) then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def using B by blast show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" by (metis (no_types, lifting) has_integral_integrable_integral) qed qed subsection \Continuity of the integral (for a 1-dimensional interval)\ lemma integrable_alt: fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof let ?F = "\x. if x \ s then f x else 0" assume ?l then obtain y where intF: "\a b. ?F integrable_on cbox a b" and y: "\e. 0 < e \ \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e" unfolding integrable_on_def has_integral_alt'[of f] by auto show ?r proof (intro conjI allI impI intF) fix e::real assume "e > 0" then have "e/2 > 0" by auto obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e/2" using \0 < e/2\ y by blast show "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" proof (intro conjI exI impI allI, rule \0 < B\) fix a b c d::'n assume sub: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" using sub by (auto intro: norm_triangle_half_l dest: B) qed qed next let ?F = "\x. if x \ s then f x else 0" assume rhs: ?r let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) ?F)" unfolding Cauchy_def proof (intro allI impI) fix e::real assume "e > 0" with rhs obtain B where "0 < B" and B: "\a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" by blast obtain N where N: "B \ real N" using real_arch_simple by blast have "ball 0 B \ ?cube n" if n: "n \ N" for n proof - have "sum ((*\<^sub>R) (- real n)) Basis \ i \ x \ i \ x \ i \ sum ((*\<^sub>R) (real n)) Basis \ i" if "norm x < B" "i \ Basis" for x i::'n using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) then show ?thesis by (auto simp: mem_box dist_norm) qed then show "\M. \m\M. \n\M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" by (fastforce simp add: dist_norm intro!: B) qed then obtain i where i: "(\n. integral (?cube n) ?F) \ i" using convergent_eq_Cauchy by blast have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - i) < e" if "e > 0" for e proof - have *: "e/2 > 0" using that by auto then obtain N where N: "\n. N \ n \ norm (i - integral (?cube n) ?F) < e/2" using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson obtain B where "0 < B" and B: "\a b c d. \ball 0 B \ cbox a b; ball 0 B \ cbox c d\ \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" using rhs * by meson let ?B = "max (real N) B" show ?thesis proof (intro exI conjI allI impI) show "0 < ?B" using \B > 0\ by auto fix a b :: 'n assume "ball 0 ?B \ cbox a b" moreover obtain n where n: "max (real N) B \ real n" using real_arch_simple by blast moreover have "ball 0 B \ ?cube n" proof fix x :: 'n assume x: "x \ ball 0 B" have "\norm (0 - x) < B; i \ Basis\ \ sum ((*\<^sub>R) (-n)) Basis \ i\ x \ i \ x \ i \ sum ((*\<^sub>R) n) Basis \ i" for i using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) then show "x \ ?cube n" using x by (auto simp: mem_box dist_norm) qed ultimately show "norm (integral (cbox a b) ?F - i) < e" using norm_triangle_half_l [OF B N] by force qed qed then show ?l unfolding integrable_on_def has_integral_alt'[of f] using rhs by blast qed lemma integrable_altD: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto lemma integrable_alt_subset: fixes f :: "'a::euclidean_space \ 'b::banach" shows "f integrable_on S \ (\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm(integral (cbox a b) (\x. if x \ S then f x else 0) - integral (cbox c d) (\x. if x \ S then f x else 0)) < e)" (is "_ = ?rhs") proof - let ?g = "\x. if x \ S then f x else 0" have "f integrable_on S \ (\a b. ?g integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)" by (rule integrable_alt) also have "\ = ?rhs" proof - { fix e :: "real" assume e: "\e. e>0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" and "e > 0" obtain B where "B > 0" and B: "\a b c d. \ball 0 B \ cbox a b; cbox a b \ cbox c d\ \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2" using \e > 0\ e [of "e/2"] by force have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" proof (intro exI allI conjI impI) fix a b c d :: "'a" let ?\ = "\i\Basis. max (a \ i) (c \ i) *\<^sub>R i" let ?\ = "\i\Basis. min (b \ i) (d \ i) *\<^sub>R i" show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" if ball: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" proof - have B': "norm (integral (cbox a b \ cbox c d) ?g - integral (cbox x y) ?g) < e/2" if "cbox a b \ cbox c d \ cbox x y" for x y using B [of ?\ ?\ x y] ball that by (simp add: Int_interval [symmetric]) show ?thesis using B' [of a b] B' [of c d] norm_triangle_half_r by blast qed qed (use \B > 0\ in auto)} then show ?thesis by force qed finally show ?thesis . qed lemma integrable_on_subcbox: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on S" and sub: "cbox a b \ S" shows "f integrable_on cbox a b" proof - have "(\x. if x \ S then f x else 0) integrable_on cbox a b" by (simp add: intf integrable_altD(1)) then show ?thesis by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) qed subsection \A straddling criterion for integrability\ lemma integrable_straddle_interval: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" shows "f integrable_on cbox a b" proof - have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ d fine p1 \ p2 tagged_division_of cbox a b \ d fine p2 \ \(\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)\ < e)" if "e > 0" for e proof - have e: "e/3 > 0" using that by auto then obtain g h i j where ij: "\i - j\ < e/3" and "(g has_integral i) (cbox a b)" and "(h has_integral j) (cbox a b)" and fgh: "\x. x \ cbox a b \ g x \ f x \ f x \ h x" using assms real_norm_def by metis then obtain d1 d2 where "gauge d1" "gauge d2" and d1: "\p. \p tagged_division_of cbox a b; d1 fine p\ \ \(\(x,K)\p. content K *\<^sub>R g x) - i\ < e/3" and d2: "\p. \p tagged_division_of cbox a b; d2 fine p\ \ \(\(x,K) \ p. content K *\<^sub>R h x) - j\ < e/3" by (metis e has_integral real_norm_def) have "\(\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)\ < e" if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 proof - have *: "\g1 g2 h1 h2 f1 f2. \\g2 - i\ < e/3; \g1 - i\ < e/3; \h2 - j\ < e/3; \h1 - j\ < e/3; g1 - h2 \ f1 - f2; f1 - f2 \ h1 - g2\ \ \f1 - f2\ < e" using \e > 0\ ij by arith have 0: "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" unfolding sum_subtractf[symmetric] apply (auto intro!: sum_nonneg) apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ done show ?thesis proof (rule *) show "\(\(x,K) \ p2. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p2 12]) show "\(\(x,K) \ p1. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p1 11]) show "\(\(x,K) \ p2. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p2 22]) show "\(\(x,K) \ p1. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p1 21]) qed (use 0 in auto) qed then show ?thesis by (rule_tac x="\x. d1 x \ d2 x" in exI) (auto simp: fine_Int intro: \gauge d1\ \gauge d2\ d1 d2) qed then show ?thesis by (simp add: integrable_Cauchy) qed lemma integrable_straddle: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) s \ (h has_integral j) s \ \i - j\ < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - let ?fs = "(\x. if x \ s then f x else 0)" have "?fs integrable_on cbox a b" for a b proof (rule integrable_straddle_interval) fix e::real assume "e > 0" then have *: "e/4 > 0" by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/4" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/4" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/4" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson define c where "c = (\i\Basis. min (a\i) (- (max Bg Bh)) *\<^sub>R i)" define d where "d = (\i\Basis. max (b\i) (max Bg Bh) *\<^sub>R i)" have "\norm (0 - x) < Bg; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBg: "ball 0 Bg \ cbox c d" by (auto simp: mem_box dist_norm) have "\norm (0 - x) < Bh; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBh: "ball 0 Bh \ cbox c d" by (auto simp: mem_box dist_norm) have ab_cd: "cbox a b \ cbox c d" by (auto simp: c_def d_def subset_box_imp) have **: "\ch cg ag ah::real. \\ah - ag\ \ \ch - cg\; \cg - i\ < e/4; \ch - j\ < e/4\ \ \ag - ah\ < e" using ij by arith show "\g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. g x \ (if x \ s then f x else 0) \ (if x \ s then f x else 0) \ h x)" proof (intro exI ballI conjI) have eq: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = (if x \ s then f x - g x else (0::real))" by auto have int_hg: "(\x. if x \ s then h x - g x else 0) integrable_on cbox a b" "(\x. if x \ s then h x - g x else 0) integrable_on cbox c d" by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" by (intro integrable_integral int_g int_h)+ then have "integral (cbox a b) ?gs \ integral (cbox a b) ?hs" apply (rule has_integral_le) using fgh by force then have "0 \ integral (cbox a b) ?hs - integral (cbox a b) ?gs" by simp then have "\integral (cbox a b) ?hs - integral (cbox a b) ?gs\ \ \integral (cbox c d) ?hs - integral (cbox c d) ?gs\" apply (simp add: integral_diff [symmetric] int_g int_h) apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ done then show "\integral (cbox a b) ?gs - integral (cbox a b) ?hs\ < e" apply (rule **) apply (rule Bg ballBg Bh ballBh)+ done show "\x. x \ cbox a b \ ?gs x \ ?fs x" "\x. x \ cbox a b \ ?fs x \ ?hs x" using fgh by auto qed qed then have int_f: "?fs integrable_on cbox a b" for a b by simp have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" if "0 < e" for e proof - have *: "e/3 > 0" using that by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/3" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where "Bg > 0" and Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/3" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where "Bh > 0" and Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/3" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson { fix a b c d :: 'n assume as: "ball 0 (max Bg Bh) \ cbox a b" "ball 0 (max Bg Bh) \ cbox c d" have **: "ball 0 Bg \ ball (0::'n) (max Bg Bh)" "ball 0 Bh \ ball (0::'n) (max Bg Bh)" by auto have *: "\ga gc ha hc fa fc. \\ga - i\ < e/3; \gc - i\ < e/3; \ha - j\ < e/3; \hc - j\ < e/3; ga \ fa; fa \ ha; gc \ fc; fc \ hc\ \ \fa - fc\ < e" using ij by arith have "abs (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" proof (rule *) show "\integral (cbox a b) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox c d) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox a b) ?hs - j\ < e/3" using "**" Bh as by blast show "\integral (cbox c d) ?hs - j\ < e/3" using "**" Bh as by blast qed (use int_f int_g int_h fgh in \simp_all add: integral_le\) } then show ?thesis apply (rule_tac x="max Bg Bh" in exI) using \Bg > 0\ by auto qed then show ?thesis unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) qed subsection \Adding integrals over several sets\ lemma has_integral_Un: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" "(f has_integral j) T" and neg: "negligible (S \ T)" shows "(f has_integral (i + j)) (S \ T)" unfolding has_integral_restrict_UNIV[symmetric, of f] proof (rule has_integral_spike[OF neg]) let ?f = "\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)" show "(?f has_integral i + j) UNIV" by (simp add: f has_integral_add) qed auto lemma integral_Un [simp]: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" "f integrable_on T" "negligible (S \ T)" shows "integral (S \ T) f = integral S f + integral T f" by (simp add: has_integral_Un assms integrable_integral integral_unique) lemma integrable_Un: fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" shows "f integrable_on (A \ B)" proof - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" by (auto simp: integrable_on_def) from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) qed lemma integrable_Un': fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" shows "f integrable_on C" using integrable_Un[of A B f] assms by simp lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "finite \" and int: "\S. S \ \ \ (f has_integral (i S)) S" and neg: "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - let ?\ = "((\(a,b). a \ b) ` {(a,b). a \ \ \ b \ {y. y \ \ \ a \ y}})" have "((\x. if x \ \\ then f x else 0) has_integral sum i \) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (\ \ \)" by (simp add: \) moreover have "{(a, b). a \ \ \ b \ {y \ \. a \ y}} \ \ \ \" by auto ultimately show "finite ?\" by (blast intro: finite_subset[of _ "\ \ \"]) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ \\ then f x else 0) = (\A\\. if x \ A then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix S assume "S \ \" "x \ S" moreover then have "\b\\. x \ b \ b = S" using that by blast ultimately show "f x = (\A\\. if x \ A then f x else 0)" by (simp add: sum.delta[OF \]) qed next show "((\x. \A\\. if x \ A then f x else 0) has_integral (\A\\. i A)) UNIV" apply (rule has_integral_sum [OF \]) using int by (simp add: has_integral_restrict_UNIV) qed then show ?thesis using has_integral_restrict_UNIV by blast qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" apply (rule integral_unique) by (meson assms has_integral_combine_division has_integral_integrable_integral) lemma has_integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and \: "\ division_of K" and "K \ S" shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L proof - have "L \ S" using \K \ S\ \ that by blast then show "f integrable_on L" using that by (metis (no_types) f \ division_ofD(4) integrable_on_subcbox) qed then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed lemma integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\ division_of S" shows "integral S f = sum (\i. integral i f) \" apply (rule integral_unique) apply (rule has_integral_combine_division_topdown) using assms apply auto done lemma integrable_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of S" and f: "\i. i \ \ \ f integrable_on i" shows "f integrable_on S" using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \]) lemma integrable_on_subdivision: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of i" and f: "f integrable_on S" and "i \ S" shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i proof - have "i \ S" using assms that by auto then show "f integrable_on i" using that by (metis (no_types) \ f division_ofD(4) integrable_on_subcbox) qed then show ?thesis using \ integrable_combine_division by blast qed subsection \Also tagged divisions\ lemma has_integral_combine_tagged_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of S" and "\(x,k) \ p. (f has_integral (i k)) k" shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" using assms(2) apply (intro has_integral_combine_division) apply (auto simp: has_integral_integral[symmetric] intro: division_of_tagged_division[OF assms(1)]) apply auto done also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) finally show ?thesis using assms by (auto simp add: has_integral_iff intro!: sum.cong) qed lemma integral_combine_tagged_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of (cbox a b)" and "\(x,k)\p. f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" apply (rule integral_unique) apply (rule has_integral_combine_tagged_division[OF assms(1)]) using assms(2) apply auto done lemma has_integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and p: "p tagged_division_of (cbox a b)" shows "(f has_integral (sum (\(x,K). integral K f) p)) (cbox a b)" proof - have "(f has_integral integral K f) K" if "(x,K) \ p" for x K by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) then show ?thesis by (metis assms case_prodI2 has_integral_integrable_integral integral_combine_tagged_division_bottomup) qed lemma integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on cbox a b" and "p tagged_division_of (cbox a b)" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" apply (rule integral_unique [OF has_integral_combine_tagged_division_topdown]) using assms apply auto done subsection \Henstock's lemma\ lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on cbox a b" and "e > 0" and "gauge d" and less_e: "\p. \p tagged_division_of (cbox a b); d fine p\ \ norm (sum (\(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?lhs \ e") proof (rule field_le_epsilon) fix k :: real assume "k > 0" let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" note p' = tagged_partial_division_ofD[OF p(1)] have "\(snd ` p) \ cbox a b" using p'(3) by fastforce then obtain q where q: "snd ` p \ q" and qdiv: "q division_of cbox a b" by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) note q' = division_ofD[OF qdiv] define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have "finite r" using q' unfolding r_def by auto have "\p. p tagged_division_of i \ d fine p \ norm (?SUM p - integral i f) < k / (real (card r) + 1)" if "i\r" for i proof - have gt0: "k / (real (card r) + 1) > 0" using \k > 0\ by simp have i: "i \ q" using that unfolding r_def by auto then obtain u v where uv: "i = cbox u v" using q'(4) by blast then have "cbox u v \ cbox a b" using i q'(2) by auto then have "f integrable_on cbox u v" by (rule integrable_subinterval[OF intf]) with integrable_integral[OF this, unfolded has_integral[of f]] obtain dd where "gauge dd" and dd: "\\. \\ tagged_division_of cbox u v; dd fine \\ \ norm (?SUM \ - integral (cbox u v) f) < k / (real (card r) + 1)" using gt0 by auto with gauge_Int[OF \gauge d\ \gauge dd\] obtain qq where qq: "qq tagged_division_of cbox u v" "(\x. d x \ dd x) fine qq" using fine_division_exists by blast with dd[of qq] show ?thesis by (auto simp: fine_Int uv) qed then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ d fine qq i \ norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" by metis let ?p = "p \ \(qq ` r)" have "norm (?SUM ?p - integral (cbox a b) f) < e" proof (rule less_e) show "d fine ?p" by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "finite (snd ` p)" by (simp add: p'(1)) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" apply (subst Int_commute) apply (rule Int_interior_Union_intervals) using r_def q'(5) q(1) apply auto by (simp add: p'(4)) qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm (?SUM p + (?SUM (\(qq ` r))) - integral (cbox a b) f) < e" proof (subst sum.union_inter_neutral[symmetric, OF \finite p\], safe) show "content L *\<^sub>R f x = 0" if "(x, L) \ p" "(x, L) \ qq K" "K \ r" for x K L proof - obtain u v where uv: "L = cbox u v" using \(x,L) \ p\ p'(4) by blast have "L \ K" using qq[OF that(3)] tagged_division_ofD(3) \(x,L) \ qq K\ by metis have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" using that(1,3) q(1) unfolding r_def by auto with q'(5) have "interior L = {}" using interior_mono[OF \L \ K\] by blast then show "content L *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) qed moreover have "content M *\<^sub>R f x = 0" if x: "(x,M) \ qq K" "(x,M) \ qq L" and KL: "qq K \ qq L" and r: "K \ r" "L \ r" for x M K L proof - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] obtain u v where uv: "M = cbox u v" using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) have "interior M = {}" by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) then show "content M *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) using qq by (force simp: split_paired_all)+ moreover have "content M *\<^sub>R f x = 0" if "K \ r" "L \ r" "K \ L" "qq K = qq L" "(x, M) \ qq K" for K L x M using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) ultimately have less_e: "norm (?SUM p + sum (?SUM \ qq) r - integral (cbox a b) f) < e" apply (subst (asm) sum.reindex_nontrivial [OF \finite r\]) apply (auto simp: split_paired_all sum.neutral) done have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a proof - from that show ?thesis using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] unfolding that(3)[symmetric] norm_minus_cancel by (auto simp add: algebra_simps) qed have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. also have "\ \ e + k" proof (rule norm_le[OF less_e]) have lessk: "k * real (card r) / (1 + real (card r)) < k" using \k>0\ by (auto simp add: field_simps) have "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) \ (\x\r. k / (real (card r) + 1))" unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) also have "... < k" by (simp add: lessk add.commute mult.commute) finally show "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) < k" . next from q(1) have [simp]: "snd ` p \ q = q" by auto have "integral l f = 0" if inp: "(x, l) \ p" "(y, m) \ p" and ne: "(x, l) \ (y, m)" and "l = m" for x l y m proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast have "content (cbox u v) = 0" unfolding content_eq_0_interior using that p(1) uv by auto then show ?thesis using uv by blast qed then have "(\(x, k)\p. integral k f) = (\k\snd ` p. integral k f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) unfolding split_paired_all split_def by auto then show "(\(x, k)\p. integral k f) + (\k\r. integral k f) = integral (cbox a b) f" unfolding integral_combine_division_topdown[OF intf qdiv] r_def using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] by simp qed finally show "?lhs \ e + k" . qed lemma Henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" and less_e: "\\. \\ tagged_division_of (cbox a b); d fine \\ \ norm (sum (\(x,k). content k *\<^sub>R f x) \ - integral (cbox a b) f) < e" and tag: "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "sum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" proof - have "finite p" using tag by blast then show ?thesis unfolding split_def proof (rule sum_norm_allsubsets_bound) fix Q assume Q: "Q \ p" then have fine: "d fine Q" by (simp add: \d fine p\ fine_subset) show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) using Q tag tagged_partial_division_subset apply (force simp add: fine)+ done qed qed lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" obtains \ where "gauge \" and "\p. \p tagged_partial_division_of (cbox a b); \ fine p\ \ sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e/(2 * (real DIM('n) + 1)) > 0" using \e > 0\ by simp with integrable_integral[OF intf, unfolded has_integral] obtain \ where "gauge \" and \: "\\. \\ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - integral (cbox a b) f) < e/(2 * (real DIM('n) + 1))" by metis show thesis proof (rule that [OF \gauge \\]) fix p assume p: "p tagged_partial_division_of cbox a b" "\ fine p" have "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) \ 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))" using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally show "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) < e" . qed qed subsection \Monotone convergence (bounded interval first)\ lemma bounded_increasing_convergent: fixes f :: "nat \ real" shows "\bounded (range f); \n. f n \ f (Suc n)\ \ \l. f \ l" using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" and le: "\k x. x \ cbox a b \ (f k x) \ f (Suc k) x" and fg: "\x. x \ cbox a b \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral (cbox a b) (f k)))" shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" proof (cases "content (cbox a b) = 0") case True then show ?thesis by auto next case False have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - have "\\<^sub>F j in sequentially. f k x \ f j x" apply (rule eventually_sequentiallyI [of k]) using le x apply (force intro: transitive_stepwise_le) done then show "f k x \ g x" using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast qed have int_inc: "\n. integral (cbox a b) (f n) \ integral (cbox a b) (f (Suc n))" by (metis integral_le intf le) then obtain i where i: "(\k. integral (cbox a b) (f k)) \ i" using bounded_increasing_convergent bou by blast have "\k. \\<^sub>F x in sequentially. integral (cbox a b) (f k) \ integral (cbox a b) (f x)" unfolding eventually_sequentially by (force intro: transitive_stepwise_le int_inc) then have i': "\k. (integral(cbox a b) (f k)) \ i" using tendsto_le [OF trivial_limit_sequentially i] by blast have "(g has_integral i) (cbox a b)" unfolding has_integral real_norm_def proof clarify fix e::real assume e: "e > 0" have "\k. (\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" using intf e by (auto simp: has_integral_integral has_integral) then obtain c where c: "\x. gauge (c x)" "\x \. \\ tagged_division_of cbox a b; c x fine \\ \ abs ((\(u,K)\\. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) < e/2 ^ (x + 2)" by metis have "\r. \k\r. 0 \ i - (integral (cbox a b) (f k)) \ i - (integral (cbox a b) (f k)) < e/4" proof - have "e/4 > 0" using e by auto show ?thesis using LIMSEQ_D [OF i \e/4 > 0\] i' by auto qed then obtain r where r: "\k. r \ k \ 0 \ i - integral (cbox a b) (f k)" "\k. r \ k \ i - integral (cbox a b) (f k) < e/4" by metis have "\n\r. \k\n. 0 \ (g x) - (f k x) \ (g x) - (f k x) < e/(4 * content(cbox a b))" if "x \ cbox a b" for x proof - have "e/(4 * content (cbox a b)) > 0" by (simp add: False content_lt_nz e) with fg that LIMSEQ_D obtain N where "\n\N. norm (f n x - g x) < e/(4 * content (cbox a b))" by metis then show "\n\r. \k\n. 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" apply (rule_tac x="N + r" in exI) using fg1[OF that] apply (auto simp add: field_simps) done qed then obtain m where r_le_m: "\x. x \ cbox a b \ r \ m x" and m: "\x k. \x \ cbox a b; m x \ k\ \ 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" by metis define d where "d x = c (m x) x" for x show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R g x) - i) < e)" proof (rule exI, safe) show "gauge d" using c(1) unfolding gauge_def d_def by auto next fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "d fine \" note p'=tagged_division_ofD[OF ptag] obtain s where s: "\x. x \ \ \ m (fst x) \ s" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) have *: "\a - d\ < e" if "\a - b\ \ e/4" "\b - c\ < e/2" "\c - d\ < e/4" for a b c d using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] norm_triangle_lt[of "a - b + (b - c)" "c - d" e] by (auto simp add: algebra_simps) show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) have "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ (\i\\. \(case i of (x, K) \ content K *\<^sub>R g x) - (case i of (x, K) \ content K *\<^sub>R f (m x) x)\)" by (metis (mono_tags) sum_subtractf sum_abs) also have "... \ (\(x, k)\\. content k * (e/(4 * content (cbox a b))))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xk: "(x,K) \ \" with ptag have x: "x \ cbox a b" by blast then have "abs (content K * (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) then show "\content K * g x - content K * f (m x) x\ \ content K * e/(4 * content (cbox a b))" by (simp add: algebra_simps) qed also have "... = (e/(4 * content (cbox a b))) * (\(x, k)\\. content k)" by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) also have "... \ e/4" by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) finally show "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ e/4" . next have "norm ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))" apply (subst sum.group) using s by (auto simp: sum_subtractf split_def p'(1)) also have "\ < e/2" proof - have "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) \ (\i = 0..s. e/2 ^ (i + 2))" proof (rule sum_norm_le) fix t assume "t \ {0..s}" have "norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))" by (force intro!: sum.cong arg_cong[where f=norm]) also have "... \ e/2 ^ (t + 2)" proof (rule Henstock_lemma_part1 [OF intf]) show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" apply (rule tagged_partial_division_subset[of \]) using ptag by (auto simp: tagged_division_of_def) show "c t fine {xk \ \. m (fst xk) = t}" using \d fine \\ by (auto simp: fine_def d_def) qed (use c e in auto) finally show "norm (\(x,K)\{xk \ \. m (fst xk) = t}. content K *\<^sub>R f (m x) x - integral K (f (m x))) \ e/2 ^ (t + 2)" . qed also have "... = (e/2/2) * (\i = 0..s. (1/2) ^ i)" by (simp add: sum_distrib_left field_simps) also have "\ < e/2" by (simp add: sum_gp mult_strict_left_mono[OF _ e]) finally show "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . qed finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next have comb: "integral (cbox a b) (f y) = (\(x, k)\\. integral k (f y))" for y using integral_combine_tagged_division_topdown[OF intf ptag] by metis have f_le: "\y m n. \y \ cbox a b; n\m\ \ f m y \ f n y" using le by (auto intro: transitive_stepwise_le) have "(\(x, k)\\. integral k (f r)) \ (\(x, K)\\. integral K (f (m x)))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f r) \ integral K (f (m x))" proof (rule integral_le) show "f r integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f r y \ f (m x) y" if "y \ K" for y using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto qed qed moreover have "(\(x, K)\\. integral K (f (m x))) \ (\(x, k)\\. integral k (f s))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f (m x)) \ integral K (f s)" proof (rule integral_le) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f s integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) y \ f s y" if "y \ K" for y using that s xK f_le p'(3) by fastforce qed qed moreover have "0 \ i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" using r by auto ultimately show "\(\(x,K)\\. integral K (f (m x))) - i\ < e/4" using comb i'[of s] by auto qed qed qed with i integral_unique show ?thesis by blast qed lemma monotone_convergence_increasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes int_f: "\k. (f k) integrable_on S" and "\k x. x \ S \ (f k x) \ (f (Suc k) x)" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral S (f k)))" shows "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" proof - have lem: "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" if f0: "\k x. x \ S \ 0 \ f k x" and int_f: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f k x \ f (Suc k) x" and lim: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" for f :: "nat \ 'n::euclidean_space \ real" and g S proof - have fg: "(f k x) \ (g x)" if "x \ S" for x k apply (rule tendsto_lowerbound [OF lim [OF that]]) apply (rule eventually_sequentiallyI [of k]) using le by (force intro: transitive_stepwise_le that)+ obtain i where i: "(\k. integral S (f k)) \ i" using bounded_increasing_convergent [OF bou] le int_f integral_le by blast have i': "(integral S (f k)) \ i" for k proof - have "\k. \x. x \ S \ \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) then show ?thesis using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] by (meson int_f integral_le) qed let ?f = "(\k x. if x \ S then f k x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have int: "?f k integrable_on cbox a b" for a b k by (simp add: int_f integrable_altD(1)) have int': "\k a b. f k integrable_on cbox a b \ S" using int by (simp add: Int_commute integrable_restrict_Int) have g: "?g integrable_on cbox a b \ (\k. integral (cbox a b) (?f k)) \ integral (cbox a b) ?g" for a b proof (rule monotone_convergence_interval) have "norm (integral (cbox a b) (?f k)) \ norm (integral S (f k))" for k proof - have "0 \ integral (cbox a b) (?f k)" by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') moreover have "0 \ integral S (f k)" by (simp add: integral_nonneg f0 int_f) moreover have "integral (S \ cbox a b) (f k) \ integral S (f k)" by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) ultimately show ?thesis by (simp add: integral_restrict_Int) qed moreover obtain B where "\x. x \ range (\k. integral S (f k)) \ norm x \ B" using bou unfolding bounded_iff by blast ultimately show "bounded (range (\k. integral (cbox a b) (?f k)))" unfolding bounded_iff by (blast intro: order_trans) qed (use int le lim in auto) moreover have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?g - i) < e" if "0 < e" for e proof - have "e/4>0" using that by auto with LIMSEQ_D [OF i] obtain N where N: "\n. n \ N \ norm (integral S (f n) - i) < e/4" by metis with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" by (meson \0 < e/4\) have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \ cbox a b" for a b proof - obtain M where M: "\n. n \ M \ abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" using \e > 0\ g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) have *: "\\ \ g. \\\ - i\ < e/2; \\ - g\ < e/2; \ \ \; \ \ i\ \ \g - i\ < e" unfolding real_inner_1_right by arith show "norm (integral (cbox a b) ?g - i) < e" unfolding real_norm_def proof (rule *) show "\integral (cbox a b) (?f N) - i\ < e/2" proof (rule abs_triangle_half_l) show "\integral (cbox a b) (?f N) - integral S (f N)\ < e/2/2" using B[OF ab] by simp show "abs (i - integral S (f N)) < e/2/2" using N by (simp add: abs_minus_commute) qed show "\integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\ < e/2" by (metis le_add1 M[of "M + N"]) show "integral (cbox a b) (?f N) \ integral (cbox a b) (?f (M + N))" proof (intro ballI integral_le[OF int int]) fix x assume "x \ cbox a b" have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n apply (rule transitive_stepwise_le [OF \n \ m\ order_refl]) using dual_order.trans apply blast by (simp add: le \x \ S\) then show "(?f N)x \ (?f (M+N))x" by auto qed have "integral (cbox a b \ S) (f (M + N)) \ integral S (f (M + N))" by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) then have "integral (cbox a b) (?f (M + N)) \ integral S (f (M + N))" by (metis (no_types) inf_commute integral_restrict_Int) also have "... \ i" using i'[of "M + N"] by auto finally show "integral (cbox a b) (?f (M + N)) \ i" . qed qed then show ?thesis using \0 < B\ by blast qed ultimately have "(g has_integral i) S" unfolding has_integral_alt' by auto then show ?thesis using has_integral_integrable_integral i integral_unique by metis qed have sub: "\k. integral S (\x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" by (simp add: integral_diff int_f) have *: "\x m n. x \ S \ n\m \ f m x \ f n x" using assms(2) by (force intro: transitive_stepwise_le) have gf: "(\x. g x - f 0 x) integrable_on S \ ((\k. integral S (\x. f (Suc k) x - f 0 x)) \ integral S (\x. g x - f 0 x)) sequentially" proof (rule lem) show "\k. (\x. f (Suc k) x - f 0 x) integrable_on S" by (simp add: integrable_diff int_f) show "(\k. f (Suc k) x - f 0 x) \ g x - f 0 x" if "x \ S" for x proof - have "(\n. f (Suc n) x) \ g x" using LIMSEQ_ignore_initial_segment[OF fg[OF \x \ S\], of 1] by simp then show ?thesis by (simp add: tendsto_diff) qed show "bounded (range (\k. integral S (\x. f (Suc k) x - f 0 x)))" proof - obtain B where B: "\k. norm (integral S (f k)) \ B" using bou by (auto simp: bounded_iff) then have "norm (integral S (\x. f (Suc k) x - f 0 x)) \ B + norm (integral S (f 0))" for k unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) then show ?thesis unfolding bounded_iff by blast qed qed (use * in auto) then have "(\x. integral S (\xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) \ integral S (\x. g x - f 0 x) + integral S (f 0)" by (auto simp add: tendsto_add) moreover have "(\x. g x - f 0 x + f 0 x) integrable_on S" using gf integrable_add int_f [of 0] by metis ultimately show ?thesis by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) qed lemma has_integral_monotone_convergence_increasing: fixes f :: "nat \ 'a::euclidean_space \ real" assumes f: "\k. (f k has_integral x k) s" assumes "\k x. x \ s \ f k x \ f (Suc k) x" assumes "\x. x \ s \ (\k. f k x) \ g x" assumes "x \ x'" shows "(g has_integral x') s" proof - have x_eq: "x = (\i. integral s (f i))" by (simp add: integral_unique[OF f]) then have x: "range(\k. integral s (f k)) = range x" by auto have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" proof (intro monotone_convergence_increasing allI ballI assms) show "bounded (range(\k. integral s (f k)))" using x convergent_imp_bounded assms by metis qed (use f in auto) then have "integral s g = x'" by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) with * show ?thesis by (simp add: has_integral_integral) qed lemma monotone_convergence_decreasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f (Suc k) x \ f k x" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" shows "g integrable_on S \ (\k. integral S (f k)) \ integral S g" proof - have *: "range(\k. integral S (\x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\k. integral S (f k)))" by force have "(\x. - g x) integrable_on S \ (\k. integral S (\x. - f k x)) \ integral S (\x. - g x)" proof (rule monotone_convergence_increasing) show "\k. (\x. - f k x) integrable_on S" by (blast intro: integrable_neg intf) show "\k x. x \ S \ - f k x \ - f (Suc k) x" by (simp add: le) show "\x. x \ S \ (\k. - f k x) \ - g x" by (simp add: fg tendsto_minus) show "bounded (range(\k. integral S (\x. - f k x)))" using "*" bou bounded_scaling by auto qed then show ?thesis by (force dest: integrable_neg tendsto_minus) qed lemma integral_norm_bound_integral: fixes f :: "'n::euclidean_space \ 'a::banach" assumes int_f: "f integrable_on S" and int_g: "g integrable_on S" and le_g: "\x. x \ S \ norm (f x) \ g x" shows "norm (integral S f) \ integral S g" proof - have norm: "norm \ \ y + e" if "norm \ \ x" and "\x - y\ < e/2" and "norm (\ - \) < e/2" for e x y and \ \ :: 'a proof - have "norm (\ - \) < e/2" by (metis norm_minus_commute that(3)) moreover have "x \ y + e/2" using that(2) by linarith ultimately show ?thesis using that(1) le_less_trans[OF norm_triangle_sub[of \ \]] by (auto simp: less_imp_le) qed have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" if f: "f integrable_on cbox a b" and g: "g integrable_on cbox a b" and nle: "\x. x \ cbox a b \ norm (f x) \ g x" for f :: "'n \ 'a" and g a b proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" by meson moreover from integrable_integral[OF g,unfolded has_integral[of g]] e obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" by meson ultimately have "gauge (\x. \ x \ \ x)" using gauge_Int by blast with fine_division_exists obtain \ where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis have "\ fine \" "\ fine \" using fine_Int p(2) by blast+ show "norm (integral (cbox a b) f) \ integral (cbox a b) g + e" proof (rule norm) have "norm (content K *\<^sub>R f x) \ content K *\<^sub>R g x" if "(x, K) \ \" for x K proof- have K: "x \ K" "K \ cbox a b" using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" using \(x, K) \ \\ p(1) by blast moreover have "content K * norm (f x) \ content K * g x" by (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" apply (rule integral_norm_bound_integral[OF f integrable_linear[OF g]]) apply (simp add: bounded_linear_inner_left) apply (metis fg o_def) done then show ?thesis unfolding o_def integral_component_eq[OF g] . qed lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "(f has_integral i) S" and g: "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" using integral_norm_bound_integral_component[of f S g k] unfolding integral_unique[OF f] integral_unique[OF g] using assms by auto lemma uniformly_convergent_improper_integral: fixes f :: "'b \ real \ 'a :: {banach}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "\y x. y \ A \ x \ a \ norm (f y x) \ g x" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) case (1 \) from G have "Cauchy G" by (auto intro!: convergent_Cauchy) with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \" if "m \ M" "n \ M" for m n by (force simp: Cauchy_def) define M' where "M' = max (nat \a\) M" show ?case proof (rule exI[of _ M'], safe, goal_cases) case (1 x m n) have M': "M' \ a" "M' \ M" unfolding M'_def by linarith+ have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}" using 1 M' by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: DERIV_subset[OF deriv]) have int_f: "f x integrable_on {a'..real n}" if "a' \ a" for a' using that 1 by (cases "a' \ real n") (auto intro: integrable) have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) = norm (integral {a..real n} (f x) - integral {a..real m} (f x))" by (simp add: dist_norm norm_minus_commute) also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = integral {a..real n} (f x)" using M' and 1 by (intro integral_combine int_f) auto hence "integral {a..real n} (f x) - integral {a..real m} (f x) = integral {real m..real n} (f x)" by (simp add: algebra_simps) also have "norm \ \ integral {real m..real n} g" using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto also from int_g have "integral {real m..real n} g = G (real n) - G (real m)" by (simp add: has_integral_iff) also have "\ \ dist (G m) (G n)" by (simp add: dist_norm) also from 1 and M' have "\ < \" by (intro M) auto finally show ?case . qed qed lemma uniformly_convergent_improper_integral': fixes f :: "'b \ real \ 'a :: {banach, real_normed_algebra}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof - from le obtain a'' where le: "\y x. y \ A \ x \ a'' \ norm (f y x) \ g x" by (auto simp: eventually_at_top_linorder) define a' where "a' = max a a''" have "uniformly_convergent_on A (\b x. integral {a'..real b} (f x))" proof (rule uniformly_convergent_improper_integral) fix t assume t: "t \ a'" hence "(G has_field_derivative g t) (at t within {a..})" by (intro deriv) (auto simp: a'_def) moreover have "{a'..} \ {a..}" unfolding a'_def by auto ultimately show "(G has_field_derivative g t) (at t within {a'..})" by (rule DERIV_subset) qed (insert le, auto simp: a'_def intro: integrable G) hence "uniformly_convergent_on A (\b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" (is ?P) by (intro uniformly_convergent_add) auto also have "eventually (\x. \y\A. integral {a..a'} (f y) + integral {a'..x} (f y) = integral {a..x} (f y)) sequentially" by (intro eventually_mono [OF eventually_ge_at_top[of "nat \a'\"]] ballI integral_combine) (auto simp: a'_def intro: integrable) hence "?P \ ?thesis" by (intro uniformly_convergent_cong) simp_all finally show ?thesis . qed subsection \differentiation under the integral sign\ lemma integral_continuous_on_param: fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" shows "continuous_on U (\x. integral (cbox a b) (f x))" proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding continuous_on_def proof (safe intro!: tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] obtain X0 where X0: "x \ X0" "open X0" and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" unfolding split_beta fst_conv snd_conv dist_norm by metis have "\\<^sub>F y in at x within U. y \ X0 \ U" using X0(1) X0(2) eventually_at_topological by auto then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" proof eventually_elim case (elim y) have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = norm (integral (cbox a b) (\t. f y t - f x t))" using elim \x \ U\ unfolding dist_norm by (subst integral_diff) (auto intro!: integrable_continuous continuous_intros) also have "\ \ e * content (cbox a b)" using elim \x \ U\ by (intro integrable_bound) (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] integrable_continuous continuous_intros) also have "\ < e'" using \0 < e'\ \e > 0\ by (auto simp: e_def field_split_simps) finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . qed qed qed (auto intro!: continuous_on_const) lemma leibniz_rule: fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" shows "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" unfolding split_beta fst_conv snd_conv by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover have "\\<^sub>F x in at x0 within U. x \ X0" using \open X0\ \x0 \ X0\ eventually_at_topological by blast moreover have "\\<^sub>F x in at x0 within U. x \ x0" by (auto simp: eventually_at_filter) moreover have "\\<^sub>F x in at x0 within U. x \ U" by (auto simp: eventually_at_filter) ultimately show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" proof eventually_elim case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, auto intro!: integrable_diff integrable_f2 continuous_intros intro: integrable_continuous)+ also { fix t assume t: "t \ (cbox a b)" have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) from t have deriv: "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" if "y \ X0 \ U" for y unfolding has_vector_derivative_def[symmetric] using that \x \ X0\ by (intro has_derivative_within_subset[OF fx]) auto have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) from differentiable_bound_linearization[OF seg deriv this] X0 have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" by (intro integral_norm_bound_integral) (auto intro!: continuous_intros integrable_diff integrable_f2 intro: integrable_continuous) also have "\ = content (cbox a b) * e * norm (x - x0)" by simp also have "\ < e' * norm (x - x0)" using \e' > 0\ apply (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) apply (simp add: divide_simps e_def not_less) done finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case by (auto simp: divide_simps) qed qed (rule blinfun.bounded_linear_right) qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) lemma has_vector_derivative_eq_has_derivative_blinfun: "(f has_vector_derivative f') (at x within U) \ (f has_derivative blinfun_scaleR_left f') (at x within U)" by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: fixes f::"real \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_scaleR_left (fx x t)"]) using fx cont_fx apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) done qed lemma has_field_derivative_eq_has_derivative_blinfun: "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]) using fx cont_fx apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) done qed lemma leibniz_rule_field_differentiable: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "x0 \ U" "convex U" shows "(\x. integral (cbox a b) (f x)) field_differentiable at x0 within U" using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) subsection \Exchange uniform limit and integral\ lemma uniform_limit_integral_cbox: fixes f::"'a \ 'b::euclidean_space \ 'c::banach" assumes u: "uniform_limit (cbox a b) f g F" assumes c: "\n. continuous_on (cbox a b) (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) (cbox a b)" "(g has_integral J) (cbox a b)" "(I \ J) F" proof - have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" by atomize_elim (auto simp: integrable_on_def intro!: choice) moreover have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover have "(I \ J) F" proof cases assume "content (cbox a b) = 0" hence "I = (\_. 0)" "J = 0" by (auto intro!: has_integral_unique I J) thus ?thesis by simp next assume content_nonzero: "content (cbox a b) \ 0" show ?thesis proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' = e/2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) then show "\\<^sub>F n in F. dist (I n) J < e" proof eventually_elim case (elim n) have "I n = integral (cbox a b) (f n)" "J = integral (cbox a b) g" using I[of n] J by (simp_all add: integral_unique) then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" by (simp add: integral_diff dist_norm) also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed qed ultimately show ?thesis .. qed lemma uniform_limit_integral: fixes f::"'a \ 'b::ordered_euclidean_space \ 'c::banach" assumes u: "uniform_limit {a..b} f g F" assumes c: "\n. continuous_on {a..b} (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) {a..b}" "(g has_integral J) {a..b}" "(I \ J) F" by (metis interval_cbox assms uniform_limit_integral_cbox) subsection \Integration by parts\ lemma integration_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" proof - interpret bounded_bilinear prod by fact have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral (prod (f b) (g b) - prod (f a) (g a))) {a..b}" using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) (auto intro!: continuous_intros continuous_on has_vector_derivative) from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) qed lemma integration_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" proof - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" unfolding integrable_on_def by blast hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp from integration_by_parts_interior_strong[OF assms(1-7) this] show ?thesis unfolding integrable_on_def by blast qed lemma integrable_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) subsection \Integration by substitution\ lemma has_integral_substitution_general: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) show "(g has_vector_derivative g' x) (at x within {a..b})" using deriv has_field_derivative_iff_has_vector_derivative that by blast show "((\x. integral {c..x} f) has_vector_derivative f (g x)) (at (g x) within g ` {a..b})" using that le subset by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_Icc_at o_def) have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_strong: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" "g a \ g b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) by (cases "g a = g b") auto lemma has_integral_substitution: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes "a \ b" "g a \ g b" "g ` {a..b} \ {c..d}" and "continuous_on {c..d} f" and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" by (intro has_integral_substitution_strong[of "{}" a b g c d] assms) (auto intro: DERIV_continuous_on assms) lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto subsection \Compute a double integral using iterated integrals and switching the order of integration\ lemma continuous_on_imp_integrable_on_Pair1: fixes f :: "_ \ 'b::banach" assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" shows "(\y. f (x, y)) integrable_on (cbox c d)" proof - have "f \ (\y. (x, y)) integrable_on (cbox c d)" apply (rule integrable_continuous) apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]]) using x apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con) done then show ?thesis by (simp add: o_def) qed lemma integral_integrable_2dim: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) f" shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" proof (cases "content(cbox c d) = 0") case True then show ?thesis by (simp add: True integrable_const) next case False have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" by (simp add: assms compact_cbox compact_uniformly_continuous) { fix x::'a and e::real assume x: "x \ cbox a b" and e: "0 < e" then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" by (auto simp: False content_lt_nz e) then obtain dd where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ \ norm (f x' - f x) \ e/(2 * content (cbox c d))" "dd>0" using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] by (auto simp: dist_norm intro: less_imp_le) have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" apply (rule_tac x=dd in exI) using dd e2_gt assms x apply clarify apply (rule le_less_trans [OF _ e2_less]) apply (rule integrable_bound) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis apply (rule integrable_continuous) apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) done qed lemma integral_split: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on (cbox a b)" and k: "k \ Basis" shows "integral (cbox a b) f = integral (cbox a b \ {x. x\k \ c}) f + integral (cbox a b \ {x. x\k \ c}) f" apply (rule integral_unique [OF has_integral_split [where c=c]]) using k f apply (auto simp: has_integral_integral [symmetric]) done lemma integral_swap_operativeI: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes f: "continuous_on s f" and e: "0 < e" shows "operative conj True (\k. \a b c d. cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s \ norm(integral (cbox (a,c) (b,d)) f - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) \ e * content (cbox (a,c) (b,d)))" proof (standard, auto) fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b assume *: "box (a, c) (b, d) = {}" and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" and cb2: "cbox (u, w) (v, z) \ s" then have c0: "content (cbox (a, c) (b, d)) = 0" using * unfolding content_eq_0_interior by simp have c0': "content (cbox (u, w) (v, z)) = 0" by (fact content_0_subset [OF c0 cb1]) show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u,w) (v,z))" using content_cbox_pair_eq0_D [OF c0'] by (force simp add: c0') next fix a::'a and c::'b and b::'a and d::'b and M::real and i::'a and j::'b and u::'a and v::'a and w::'b and z::'b assume ij: "(i,j) \ Basis" and n1: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and n2: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" have fcont: "continuous_on (cbox (u, w) (v, z)) f" using assms(1) continuous_on_subset subs(2) by blast then have fint: "f integrable_on cbox (u, w) (v, z)" by (metis integrable_continuous) consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij by (auto simp: Euclidean_Space.Basis_prod_def) then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z))" (is ?normle) proof cases case 1 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" using 1 f subs integral_integrable_2dim continuous_on_subset by (blast intro: integral_split) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 1 subs apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) apply (simp_all add: interval_split ij) apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) done next case 2 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" using 2 subs apply (simp_all add: interval_split) apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]]) apply (auto simp: cbox_Pair_eq interval_split [symmetric]) done with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" by (simp add: integral_add [symmetric] integral_split [symmetric] continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 2 subs apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) apply (simp_all add: interval_split ij) apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) done qed qed lemma integral_Pair_const: "integral (cbox (a,c) (b,d)) (\x. k) = integral (cbox a b) (\x. integral (cbox c d) (\y. k))" by (simp add: content_Pair) lemma integral_prod_continuous: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" proof (cases "content ?CBOX = 0") case True then show ?thesis by (auto simp: content_Pair) next case False then have cbp: "content ?CBOX > 0" using content_lt_nz by blast have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" proof (rule dense_eq0_I, simp) fix e :: real assume "0 < e" with \content ?CBOX > 0\ have "0 < e/content ?CBOX" by simp have f_int_acbd: "f integrable_on ?CBOX" by (rule integrable_continuous [OF assms]) { fix p assume p: "p division_of ?CBOX" then have "finite p" by blast define e' where "e' = e/content ?CBOX" with \0 < e\ \0 < e/content ?CBOX\ have "0 < e'" by simp interpret operative conj True "\k. \a' b' c' d'. cbox (a', c') (b', d') \ k \ cbox (a', c') (b', d') \ ?CBOX \ norm (integral (cbox (a', c') (b', d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f ((x, y))))) \ e' * content (cbox (a', c') (b', d'))" using assms \0 < e'\ by (rule integral_swap_operativeI) have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e' * content ?CBOX" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e' * content (cbox (u, w) (v, z))" using that division [of p "(a, c)" "(b, d)"] p \finite p\ by (auto simp add: comm_monoid_set_F_and) with False have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this { fix k::real and \ and u::'a and v w and z::'b and t1 t2 l assume k: "0 < k" and nf: "\x y u v. \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ \ norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" and p_acbd: "\ tagged_division_of cbox (a,c) (b,d)" and fine: "(\x. ball x k) fine \" "((t1,t2), l) \ \" and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" have t: "t1 \ cbox a b" "t2 \ cbox c d" by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ have ls: "l \ ball (t1,t2) k" using fine by (simp add: fine_def Ball_def) { fix x1 x2 assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" then have x: "x1 \ cbox a b" "x2 \ cbox c d" using uwvz_sub by auto have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" by (simp add: norm_Pair norm_minus_commute) also have "norm (t1 - x1, t2 - x2) < k" using xuvwz ls uwvz_sub unfolding ball_def by (force simp add: cbox_Pair_eq dist_norm ) finally have "norm (f (x1,x2) - f (t1,t2)) \ e/content ?CBOX/2" using nf [OF t x] by simp } note nf' = this have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" using f_int_acbd uwvz_sub integrable_on_subcbox by blast have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" using assms continuous_on_subset uwvz_sub by (blast intro: continuous_on_imp_integrable_on_Pair1) have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) using cbp \0 < e/content ?CBOX\ nf' apply (auto simp: integrable_diff f_int_uwvz integrable_const) done have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast have normint_wz: "\x. x \ cbox u v \ norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) \ e * content (cbox w z) / content (cbox (a, c) (b, d))/2" apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) using cbp \0 < e/content ?CBOX\ nf' apply (auto simp: integrable_diff f_int_uv integrable_const) done have "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" apply (rule integrable_bound [OF _ _ normint_wz]) using cbp \0 < e/content ?CBOX\ apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const) done also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp add: content_Pair field_split_simps) finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp only: integral_diff [symmetric] int_integrable integrable_const) have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves by (simp add: norm_minus_commute) have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k proof - obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" and fine: "(\x. ball x k) fine p" using fine_division_exists \0 < k\ by blast show ?thesis apply (rule op_acbd [OF division_of_tagged_division [OF ptag]]) using that fine ptag \0 < k\ by (auto simp: *) qed then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp qed lemma integral_swap_2dim: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" proof - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) apply (force simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)+ done then show ?thesis by force qed theorem integral_swap_continuous: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" proof - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" using integral_prod_continuous [OF assms] by auto also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" by (rule integral_swap_2dim [OF assms]) also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" using integral_prod_continuous [OF swap_continuous] assms by auto finally show ?thesis . qed subsection \Definite integrals for exponential and power function\ lemma has_integral_exp_minus_to_infinity: assumes a: "a > 0" shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" { fix k :: nat assume k: "of_nat k \ c" from k a have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (insert a, simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" { fix k :: nat have "(f k has_integral F k) {0..c}" proof (cases "inverse (of_nat (Suc k)) \ c") case True { fix x assume x: "x \ inverse (1 + real k)" have "0 < inverse (1 + real k)" by simp also note x finally have "x > 0" . } note x = this hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const simp: has_field_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed } note has_integral_f = this have integral_f: "integral {0..c} (f k) = F k" for k using has_integral_f[of k] by (rule integral_unique) have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) also note x finally have "inverse (real (Suc (Suc k))) \ x" . } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" by (rule has_integral_on_superset) (insert False, simp_all add: f_def) with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (insert assms, simp add: integral_f F_def field_split_simps) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) from conjunct2[OF *] and this have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) with conjunct1[OF *] show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Lebesgue_Measure.thy b/src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy +++ b/src/HOL/Analysis/Lebesgue_Measure.thy @@ -1,1813 +1,1841 @@ (* Title: HOL/Analysis/Lebesgue_Measure.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Author: Jeremy Avigad Author: Luke Serafin *) section \Lebesgue Measure\ theory Lebesgue_Measure imports Finite_Product_Measure Caratheodory Complete_Measure Summation_Tests Regularity begin lemma measure_eqI_lessThan: fixes M N :: "real measure" assumes sets: "sets M = sets borel" "sets N = sets borel" assumes fin: "\x. emeasure M {x <..} < \" assumes "\x. emeasure M {x <..} = emeasure N {x <..}" shows "M = N" proof (rule measure_eqI_generator_eq_countable) let ?LT = "\a::real. {a <..}" let ?E = "range ?LT" show "Int_stable ?E" by (auto simp: Int_stable_def lessThan_Int_lessThan) show "?E \ Pow UNIV" "sets M = sigma_sets UNIV ?E" "sets N = sigma_sets UNIV ?E" unfolding sets borel_Ioi by auto show "?LT`Rats \ ?E" "(\i\Rats. ?LT i) = UNIV" "\a. a \ ?LT`Rats \ emeasure M a \ \" using fin by (auto intro: Rats_no_bot_less simp: less_top) qed (auto intro: assms countable_rat) subsection \Measures defined by monotonous functions\ text \ Every right-continuous and nondecreasing function gives rise to a measure on the reals: \ definition\<^marker>\tag important\ interval_measure :: "(real \ real) \ real measure" where "interval_measure F = extend_measure UNIV {(a, b). a \ b} (\(a, b). {a<..b}) (\(a, b). ennreal (F b - F a))" lemma\<^marker>\tag important\ emeasure_interval_measure_Ioc: assumes "a \ b" assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "emeasure (interval_measure F) {a<..b} = F b - F a" proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \a \ b\]) show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \ b}" proof (unfold_locales, safe) fix a b c d :: real assume *: "a \ b" "c \ d" then show "\C\{{a<..b} |a b. a \ b}. finite C \ disjoint C \ {a<..b} - {c<..d} = \C" proof cases let ?C = "{{a<..b}}" assume "b < c \ d \ a \ d \ c" with * have "?C \ {{a<..b} |a b. a \ b} \ finite ?C \ disjoint ?C \ {a<..b} - {c<..d} = \?C" by (auto simp add: disjoint_def) thus ?thesis .. next let ?C = "{{a<..c}, {d<..b}}" assume "\ (b < c \ d \ a \ d \ c)" with * have "?C \ {{a<..b} |a b. a \ b} \ finite ?C \ disjoint ?C \ {a<..b} - {c<..d} = \?C" by (auto simp add: disjoint_def Ioc_inj) (metis linear)+ thus ?thesis .. qed qed (auto simp: Ioc_inj, metis linear) next fix l r :: "nat \ real" and a b :: real assume l_r[simp]: "\n. l n \ r n" and "a \ b" and disj: "disjoint_family (\n. {l n<..r n})" assume lr_eq_ab: "(\i. {l i<..r i}) = {a<..b}" have [intro, simp]: "\a b. a \ b \ F a \ F b" by (auto intro!: l_r mono_F) { fix S :: "nat set" assume "finite S" moreover note \a \ b\ moreover have "\i. i \ S \ {l i <.. r i} \ {a <.. b}" unfolding lr_eq_ab[symmetric] by auto ultimately have "(\i\S. F (r i) - F (l i)) \ F b - F a" proof (induction S arbitrary: a rule: finite_psubset_induct) case (psubset S) show ?case proof cases assume "\i\S. l i < r i" with \finite S\ have "Min (l ` {i\S. l i < r i}) \ l ` {i\S. l i < r i}" by (intro Min_in) auto then obtain m where m: "m \ S" "l m < r m" "l m = Min (l ` {i\S. l i < r i})" by fastforce have "(\i\S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\i\S - {m}. F (r i) - F (l i))" using m psubset by (intro sum.remove) auto also have "(\i\S - {m}. F (r i) - F (l i)) \ F b - F (r m)" proof (intro psubset.IH) show "S - {m} \ S" using \m\S\ by auto show "r m \ b" using psubset.prems(2)[OF \m\S\] \l m < r m\ by auto next fix i assume "i \ S - {m}" then have i: "i \ S" "i \ m" by auto { assume i': "l i < r i" "l i < r m" with \finite S\ i m have "l m \ l i" by auto with i' have "{l i <.. r i} \ {l m <.. r m} \ {}" by auto then have False using disjoint_family_onD[OF disj, of i m] i by auto } then have "l i \ r i \ r m \ l i" unfolding not_less[symmetric] using l_r[of i] by auto then show "{l i <.. r i} \ {r m <.. b}" using psubset.prems(2)[OF \i\S\] by auto qed also have "F (r m) - F (l m) \ F (r m) - F a" using psubset.prems(2)[OF \m \ S\] \l m < r m\ by (auto simp add: Ioc_subset_iff intro!: mono_F) finally show ?case by (auto intro: add_mono) qed (auto simp add: \a \ b\ less_le) qed } note claim1 = this (* second key induction: a lower bound on the measures of any finite collection of Ai's that cover an interval {u..v} *) { fix S u v and l r :: "nat \ real" assume "finite S" "\i. i\S \ l i < r i" "{u..v} \ (\i\S. {l i<..< r i})" then have "F v - F u \ (\i\S. F (r i) - F (l i))" proof (induction arbitrary: v u rule: finite_psubset_induct) case (psubset S) show ?case proof cases assume "S = {}" then show ?case using psubset by (simp add: mono_F) next assume "S \ {}" then obtain j where "j \ S" by auto let ?R = "r j < u \ l j > v \ (\i\S-{j}. l i \ l j \ r j \ r i)" show ?case proof cases assume "?R" with \j \ S\ psubset.prems have "{u..v} \ (\i\S-{j}. {l i<..< r i})" apply (auto simp: subset_eq Ball_def) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis order_trans less_le_not_le linear) done with \j \ S\ have "F v - F u \ (\i\S - {j}. F (r i) - F (l i))" by (intro psubset) auto also have "\ \ (\i\S. F (r i) - F (l i))" using psubset.prems by (intro sum_mono2 psubset) (auto intro: less_imp_le) finally show ?thesis . next assume "\ ?R" then have j: "u \ r j" "l j \ v" "\i. i \ S - {j} \ r i < r j \ l i > l j" by (auto simp: not_less) let ?S1 = "{i \ S. l i < l j}" let ?S2 = "{i \ S. r i > r j}" have "(\i\S. F (r i) - F (l i)) \ (\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i))" using \j \ S\ \finite S\ psubset.prems j by (intro sum_mono2) (auto intro: less_imp_le) also have "(\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i)) = (\i\?S1. F (r i) - F (l i)) + (\i\?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))" using psubset(1) psubset.prems(1) j apply (subst sum.union_disjoint) apply simp_all apply (subst sum.union_disjoint) apply auto apply (metis less_le_not_le) done also (xtrans) have "(\i\?S1. F (r i) - F (l i)) \ F (l j) - F u" using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis less_le_trans not_le) done also (xtrans) have "(\i\?S2. F (r i) - F (l i)) \ F v - F (r j)" using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis le_less_trans not_le) done finally (xtrans) show ?case by (auto simp: add_mono) qed qed qed } note claim2 = this (* now prove the inequality going the other way *) have "ennreal (F b - F a) \ (\i. ennreal (F (r i) - F (l i)))" proof (rule ennreal_le_epsilon) fix epsilon :: real assume egt0: "epsilon > 0" have "\i. \d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)" proof fix i note right_cont_F [of "r i"] thus "\d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)" apply - apply (subst (asm) continuous_at_right_real_increasing) apply (rule mono_F, assumption) apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec) apply (erule impE) using egt0 by (auto simp add: field_simps) qed then obtain delta where deltai_gt0: "\i. delta i > 0" and deltai_prop: "\i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)" by metis have "\a' > a. F a' - F a < epsilon / 2" apply (insert right_cont_F [of a]) apply (subst (asm) continuous_at_right_real_increasing) using mono_F apply force apply (drule_tac x = "epsilon / 2" in spec) using egt0 unfolding mult.commute [of 2] by force then obtain a' where a'lea [arith]: "a' > a" and a_prop: "F a' - F a < epsilon / 2" by auto define S' where "S' = {i. l i < r i}" obtain S :: "nat set" where "S \ S'" and finS: "finite S" and Sprop: "{a'..b} \ (\i \ S. {l i<..i. i \ S' \ open ({l i<.. {a <.. b}" by auto also have "{a <.. b} = (\i\S'. {l i<..r i})" unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans) also have "\ \ (\i \ S'. {l i<.. (\i \ S'. {l i<..i. i \ S \ l i < r i" by auto from finS have "\n. \i \ S. i \ n" by (subst finite_nat_set_iff_bounded_le [symmetric]) then obtain n where Sbound [rule_format]: "\i \ S. i \ n" .. have "F b - F a' \ (\i\S. F (r i + delta i) - F (l i))" apply (rule claim2 [rule_format]) using finS Sprop apply auto apply (frule Sprop2) apply (subgoal_tac "delta i > 0") apply arith by (rule deltai_gt0) also have "... \ (\i \ S. F(r i) - F(l i) + epsilon / 2^(i+2))" apply (rule sum_mono) apply simp apply (rule order_trans) apply (rule less_imp_le) apply (rule deltai_prop) by auto also have "... = (\i \ S. F(r i) - F(l i)) + (epsilon / 4) * (\i \ S. (1 / 2)^i)" (is "_ = ?t + _") by (subst sum.distrib) (simp add: field_simps sum_distrib_left) also have "... \ ?t + (epsilon / 4) * (\ i < Suc n. (1 / 2)^i)" apply (rule add_left_mono) apply (rule mult_left_mono) apply (rule sum_mono2) using egt0 apply auto by (frule Sbound, auto) also have "... \ ?t + (epsilon / 2)" apply (rule add_left_mono) apply (subst geometric_sum) apply auto apply (rule mult_left_mono) using egt0 apply auto done finally have aux2: "F b - F a' \ (\i\S. F (r i) - F (l i)) + epsilon / 2" by simp have "F b - F a = (F b - F a') + (F a' - F a)" by auto also have "... \ (F b - F a') + epsilon / 2" using a_prop by (intro add_left_mono) simp also have "... \ (\i\S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2" apply (intro add_right_mono) apply (rule aux2) done also have "... = (\i\S. F (r i) - F (l i)) + epsilon" by auto also have "... \ (\i\n. F (r i) - F (l i)) + epsilon" using finS Sbound Sprop by (auto intro!: add_right_mono sum_mono2) finally have "ennreal (F b - F a) \ (\i\n. ennreal (F (r i) - F (l i))) + epsilon" using egt0 by (simp add: sum_nonneg flip: ennreal_plus) then show "ennreal (F b - F a) \ (\i. ennreal (F (r i) - F (l i))) + (epsilon :: real)" by (rule order_trans) (auto intro!: add_mono sum_le_suminf simp del: sum_ennreal) qed moreover have "(\i. ennreal (F (r i) - F (l i))) \ ennreal (F b - F a)" using \a \ b\ by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1) ultimately show "(\n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)" by (rule antisym[rotated]) qed (auto simp: Ioc_inj mono_F) lemma measure_interval_measure_Ioc: assumes "a \ b" and "\x y. x \ y \ F x \ F y" and "\a. continuous (at_right a) F" shows "measure (interval_measure F) {a <.. b} = F b - F a" unfolding measure_def by (simp add: assms emeasure_interval_measure_Ioc) lemma emeasure_interval_measure_Ioc_eq: "(\x y. x \ y \ F x \ F y) \ (\a. continuous (at_right a) F) \ emeasure (interval_measure F) {a <.. b} = (if a \ b then F b - F a else 0)" using emeasure_interval_measure_Ioc[of a b F] by auto lemma\<^marker>\tag important\ sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel" apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc) apply (rule sigma_sets_eqI) apply auto apply (case_tac "a \ ba") apply (auto intro: sigma_sets.Empty) done lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV" by (simp add: interval_measure_def space_extend_measure) lemma emeasure_interval_measure_Icc: assumes "a \ b" assumes mono_F: "\x y. x \ y \ F x \ F y" assumes cont_F : "continuous_on UNIV F" shows "emeasure (interval_measure F) {a .. b} = F b - F a" proof (rule tendsto_unique) { fix a b :: real assume "a \ b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a" using cont_F by (subst emeasure_interval_measure_Ioc) (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) } note * = this let ?F = "interval_measure F" show "((\a. F b - F a) \ emeasure ?F {a..b}) (at_left a)" proof (rule tendsto_at_left_sequentially) show "a - 1 < a" by simp fix X assume "\n. X n < a" "incseq X" "X \ a" with \a \ b\ have "(\n. emeasure ?F {X n<..b}) \ emeasure ?F (\n. {X n <..b})" apply (intro Lim_emeasure_decseq) apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *) apply force apply (subst (asm ) *) apply (auto intro: less_le_trans less_imp_le) done also have "(\n. {X n <..b}) = {a..b}" using \\n. X n < a\ apply auto apply (rule LIMSEQ_le_const2[OF \X \ a\]) apply (auto intro: less_imp_le) apply (auto intro: less_le_trans) done also have "(\n. emeasure ?F {X n<..b}) = (\n. F b - F (X n))" using \\n. X n < a\ \a \ b\ by (subst *) (auto intro: less_imp_le less_le_trans) finally show "(\n. F b - F (X n)) \ emeasure ?F {a..b}" . qed show "((\a. ennreal (F b - F a)) \ F b - F a) (at_left a)" by (rule continuous_on_tendsto_compose[where g="\x. x" and s=UNIV]) (auto simp: continuous_on_ennreal continuous_on_diff cont_F) qed (rule trivial_limit_at_left_real) lemma\<^marker>\tag important\ sigma_finite_interval_measure: assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "sigma_finite_measure (interval_measure F)" apply unfold_locales apply (intro exI[of _ "(\(a, b). {a <.. b}) ` (\ \ \)"]) apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms]) done subsection \Lebesgue-Borel measure\ definition\<^marker>\tag important\ lborel :: "('a :: euclidean_space) measure" where "lborel = distr (\\<^sub>M b\Basis. interval_measure (\x. x)) borel (\f. \b\Basis. f b *\<^sub>R b)" abbreviation\<^marker>\tag important\ lebesgue :: "'a::euclidean_space measure" where "lebesgue \ completion lborel" abbreviation\<^marker>\tag important\ lebesgue_on :: "'a set \ 'a::euclidean_space measure" where "lebesgue_on \ \ restrict_space (completion lborel) \" lemma lebesgue_on_mono: assumes major: "AE x in lebesgue_on S. P x" and minor: "\x.\P x; x \ S\ \ Q x" shows "AE x in lebesgue_on S. Q x" proof - have "AE a in lebesgue_on S. P a \ Q a" using minor space_restrict_space by fastforce then show ?thesis using major by auto qed lemma integral_eq_zero_null_sets: assumes "S \ null_sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = 0" proof (rule integral_eq_zero_AE) show "AE x in lebesgue_on S. f x = 0" by (metis (no_types, lifting) assms AE_not_in lebesgue_on_mono null_setsD2 null_sets_restrict_space order_refl) qed lemma shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel" and space_lborel[simp]: "space lborel = space borel" and measurable_lborel1[simp]: "measurable M lborel = measurable M borel" and measurable_lborel2[simp]: "measurable lborel M = measurable borel M" by (simp_all add: lborel_def) lemma space_lebesgue_on [simp]: "space (lebesgue_on S) = S" by (simp add: space_restrict_space) lemma sets_lebesgue_on_refl [iff]: "S \ sets (lebesgue_on S)" by (metis inf_top.right_neutral sets.top space_borel space_completion space_lborel space_restrict_space) lemma Compl_in_sets_lebesgue: "-A \ sets lebesgue \ A \ sets lebesgue" by (metis Compl_eq_Diff_UNIV double_compl space_borel space_completion space_lborel Sigma_Algebra.sets.compl_sets) lemma measurable_lebesgue_cong: assumes "\x. x \ S \ f x = g x" shows "f \ measurable (lebesgue_on S) M \ g \ measurable (lebesgue_on S) M" by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space) lemma lebesgue_on_UNIV_eq: "lebesgue_on UNIV = lebesgue" proof - have "measure_of UNIV (sets lebesgue) (emeasure lebesgue) = lebesgue" by (metis measure_of_of_measure space_borel space_completion space_lborel) then show ?thesis by (auto simp: restrict_space_def) qed lemma integral_restrict_Int: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "T \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on (S \ T)) f" proof - have "(\x. indicat_real T x *\<^sub>R (if x \ S then f x else 0)) = (\x. indicat_real (S \ T) x *\<^sub>R f x)" by (force simp: indicator_def) then show ?thesis by (simp add: assms sets.Int Bochner_Integration.integral_restrict_space) qed lemma integral_restrict: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ T" "S \ sets lebesgue" "T \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" using integral_restrict_Int [of S T f] assms by (simp add: Int_absorb2) lemma integral_restrict_UNIV: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "integral\<^sup>L lebesgue (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" using integral_restrict_Int [of S UNIV f] assms by (simp add: lebesgue_on_UNIV_eq) lemma integrable_lebesgue_on_empty [iff]: fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" shows "integrable (lebesgue_on {}) f" by (simp add: integrable_restrict_space) lemma integral_lebesgue_on_empty [simp]: fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" shows "integral\<^sup>L (lebesgue_on {}) f = 0" by (simp add: Bochner_Integration.integral_empty) lemma has_bochner_integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \: "\ \ space M \ sets M" shows "has_bochner_integral (restrict_space M \) f i \ has_bochner_integral M (\x. indicator \ x *\<^sub>R f x) i" by (simp add: integrable_restrict_space [OF assms] integral_restrict_space [OF assms] has_bochner_integral_iff) lemma integrable_restrict_UNIV: fixes f :: "'a::euclidean_space \ 'b::{banach, second_countable_topology}" assumes S: "S \ sets lebesgue" shows "integrable lebesgue (\x. if x \ S then f x else 0) \ integrable (lebesgue_on S) f" using has_bochner_integral_restrict_space [of S lebesgue f] assms by (simp add: integrable.simps indicator_scaleR_eq_if) lemma integral_mono_lebesgue_on_AE: fixes f::"_ \ real" assumes f: "integrable (lebesgue_on T) f" and gf: "AE x in (lebesgue_on S). g x \ f x" and f0: "AE x in (lebesgue_on T). 0 \ f x" and "S \ T" and S: "S \ sets lebesgue" and T: "T \ sets lebesgue" shows "(\x. g x \(lebesgue_on S)) \ (\x. f x \(lebesgue_on T))" proof - have "(\x. g x \(lebesgue_on S)) = (\x. (if x \ S then g x else 0) \lebesgue)" by (simp add: Lebesgue_Measure.integral_restrict_UNIV S) also have "\ \ (\x. (if x \ T then f x else 0) \lebesgue)" proof (rule Bochner_Integration.integral_mono_AE') show "integrable lebesgue (\x. if x \ T then f x else 0)" by (simp add: integrable_restrict_UNIV T f) show "AE x in lebesgue. (if x \ S then g x else 0) \ (if x \ T then f x else 0)" using assms by (auto simp: AE_restrict_space_iff) show "AE x in lebesgue. 0 \ (if x \ T then f x else 0)" using f0 by (simp add: AE_restrict_space_iff T) qed also have "\ = (\x. f x \(lebesgue_on T))" using Lebesgue_Measure.integral_restrict_UNIV T by blast finally show ?thesis . qed subsection \Borel measurability\ lemma borel_measurable_if_I: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" proof - have eq: "{x. x \ S} \ {x. f x \ Y} = {x. x \ S} \ {x. f x \ Y} \ S" for Y by blast show ?thesis using f S apply (simp add: vimage_def in_borel_measurable_borel Ball_def) apply (elim all_forward imp_forward asm_rl) apply (simp only: Collect_conj_eq Collect_disj_eq imp_conv_disj eq) apply (auto simp: Compl_eq [symmetric] Compl_in_sets_lebesgue sets_restrict_space_iff) done qed lemma borel_measurable_if_D: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using assms apply (simp add: in_borel_measurable_borel Ball_def) apply (elim all_forward imp_forward asm_rl) apply (force simp: space_restrict_space sets_restrict_space image_iff intro: rev_bexI) done lemma borel_measurable_if: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue \ f \ borel_measurable (lebesgue_on S)" using assms borel_measurable_if_D borel_measurable_if_I by blast lemma borel_measurable_if_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "T \ sets lebesgue" "S \ T" shows "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on T) \ f \ borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using measurable_restrict_mono [OF _ \S \ T\] by (subst measurable_lebesgue_cong [where g = "(\x. if x \ S then f x else 0)"]) auto next assume ?rhs then show ?lhs by (simp add: \S \ sets lebesgue\ borel_measurable_if_I measurable_restrict_space1) qed lemma borel_measurable_vimage_halfspace_component_lt: "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_less]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_ge: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i \ a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_ge]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_gt: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i > a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_greater]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i \ a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_le]) apply (fastforce simp add: space_restrict_space) done lemma fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows borel_measurable_vimage_open_interval: "f \ borel_measurable (lebesgue_on S) \ (\a b. {x \ S. f x \ box a b} \ sets (lebesgue_on S))" (is ?thesis1) and borel_measurable_vimage_open: "f \ borel_measurable (lebesgue_on S) \ (\T. open T \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is ?thesis2) proof - have "{x \ S. f x \ box a b} \ sets (lebesgue_on S)" if "f \ borel_measurable (lebesgue_on S)" for a b proof - have "S = S \ space lebesgue" by simp then have "S \ (f -` box a b) \ sets (lebesgue_on S)" by (metis (no_types) box_borel in_borel_measurable_borel inf_sup_aci(1) space_restrict_space that) then show ?thesis by (simp add: Collect_conj_eq vimage_def) qed moreover have "{x \ S. f x \ T} \ sets (lebesgue_on S)" if T: "\a b. {x \ S. f x \ box a b} \ sets (lebesgue_on S)" "open T" for T proof - obtain \ where "countable \" and \: "\X. X \ \ \ \a b. X = box a b" "\\ = T" using open_countable_Union_open_box that \open T\ by metis then have eq: "{x \ S. f x \ T} = (\U \ \. {x \ S. f x \ U})" by blast have "{x \ S. f x \ U} \ sets (lebesgue_on S)" if "U \ \" for U using that T \ by blast then show ?thesis by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \countable \\]) qed moreover have eq: "{x \ S. f x \ i < a} = {x \ S. f x \ {y. y \ i < a}}" for i a by auto have "f \ borel_measurable (lebesgue_on S)" if "\T. open T \ {x \ S. f x \ T} \ sets (lebesgue_on S)" by (metis (no_types) eq borel_measurable_vimage_halfspace_component_lt open_halfspace_component_lt that) ultimately show "?thesis1" "?thesis2" by blast+ qed lemma borel_measurable_vimage_closed: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\T. closed T \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof - have eq: "{x \ S. f x \ T} = S - {x \ S. f x \ (- T)}" for T by auto show ?thesis apply (simp add: borel_measurable_vimage_open, safe) apply (simp_all (no_asm) add: eq) apply (intro sets.Diff sets_lebesgue_on_refl, force simp: closed_open) apply (intro sets.Diff sets_lebesgue_on_refl, force simp: open_closed) done qed lemma borel_measurable_vimage_closed_interval: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a b. {x \ S. f x \ cbox a b} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using borel_measurable_vimage_closed by blast next assume RHS: ?rhs have "{x \ S. f x \ T} \ sets (lebesgue_on S)" if "open T" for T proof - obtain \ where "countable \" and \: "\ \ Pow T" "\X. X \ \ \ \a b. X = cbox a b" "\\ = T" using open_countable_Union_open_cbox that \open T\ by metis then have eq: "{x \ S. f x \ T} = (\U \ \. {x \ S. f x \ U})" by blast have "{x \ S. f x \ U} \ sets (lebesgue_on S)" if "U \ \" for U using that \ by (metis RHS) then show ?thesis by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \countable \\]) qed then show ?lhs by (simp add: borel_measurable_vimage_open) qed lemma borel_measurable_vimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof assume f: ?lhs then show ?rhs using measurable_sets [OF f] by (simp add: Collect_conj_eq inf_sup_aci(1) space_restrict_space vimage_def) qed (simp add: borel_measurable_vimage_open_interval) lemma lebesgue_measurable_vimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable lebesgue" "T \ sets borel" shows "{x. f x \ T} \ sets lebesgue" using assms borel_measurable_vimage_borel [of f UNIV] by auto lemma borel_measurable_lebesgue_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable lebesgue \ (\T. T \ sets borel \ {x. f x \ T} \ sets lebesgue)" apply (intro iffI allI impI lebesgue_measurable_vimage_borel) apply (auto simp: in_borel_measurable_borel vimage_def) done subsection \<^marker>\tag unimportant\ \Measurability of continuous functions\ lemma continuous_imp_measurable_on_sets_lebesgue: assumes f: "continuous_on S f" and S: "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" proof - have "sets (restrict_space borel S) \ sets (lebesgue_on S)" by (simp add: mono_restrict_space subsetI) then show ?thesis by (simp add: borel_measurable_continuous_on_restrict [OF f] borel_measurable_subalgebra space_restrict_space) qed lemma id_borel_measurable_lebesgue [iff]: "id \ borel_measurable lebesgue" by (simp add: measurable_completion) lemma id_borel_measurable_lebesgue_on [iff]: "id \ borel_measurable (lebesgue_on S)" by (simp add: measurable_completion measurable_restrict_space1) context begin interpretation sigma_finite_measure "interval_measure (\x. x)" by (rule sigma_finite_interval_measure) auto interpretation finite_product_sigma_finite "\_. interval_measure (\x. x)" Basis proof qed simp lemma lborel_eq_real: "lborel = interval_measure (\x. x)" unfolding lborel_def Basis_real_def using distr_id[of "interval_measure (\x. x)"] by (subst distr_component[symmetric]) (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong) lemma lborel_eq: "lborel = distr (\\<^sub>M b\Basis. lborel) borel (\f. \b\Basis. f b *\<^sub>R b)" by (subst lborel_def) (simp add: lborel_eq_real) lemma nn_integral_lborel_prod: assumes [measurable]: "\b. b \ Basis \ f b \ borel_measurable borel" assumes nn[simp]: "\b x. b \ Basis \ 0 \ f b x" shows "(\\<^sup>+x. (\b\Basis. f b (x \ b)) \lborel) = (\b\Basis. (\\<^sup>+x. f b x \lborel))" by (simp add: lborel_def nn_integral_distr product_nn_integral_prod product_nn_integral_singleton) lemma emeasure_lborel_Icc[simp]: fixes l u :: real assumes [simp]: "l \ u" shows "emeasure lborel {l .. u} = u - l" proof - have "((\f. f 1) -` {l..u} \ space (Pi\<^sub>M {1} (\b. interval_measure (\x. x)))) = {1::real} \\<^sub>E {l..u}" by (auto simp: space_PiM) then show ?thesis by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc) qed lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \ u then u - l else 0)" by simp lemma\<^marker>\tag important\ emeasure_lborel_cbox[simp]: assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows "emeasure lborel (cbox l u) = (\b\Basis. (u - l) \ b)" proof - have "(\x. \b\Basis. indicator {l\b .. u\b} (x \ b) :: ennreal) = indicator (cbox l u)" by (auto simp: fun_eq_iff cbox_def split: split_indicator) then have "emeasure lborel (cbox l u) = (\\<^sup>+x. (\b\Basis. indicator {l\b .. u\b} (x \ b)) \lborel)" by simp also have "\ = (\b\Basis. (u - l) \ b)" by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left) finally show ?thesis . qed lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \ c" using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c] by (auto simp add: power_0_left) lemma emeasure_lborel_Ioo[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l <..< u} = ennreal (u - l)" proof - have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_Ioc[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l <.. u} = ennreal (u - l)" proof - have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_Ico[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l ..< u} = ennreal (u - l)" proof - have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_box[simp]: assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows "emeasure lborel (box l u) = (\b\Basis. (u - l) \ b)" proof - have "(\x. \b\Basis. indicator {l\b <..< u\b} (x \ b) :: ennreal) = indicator (box l u)" by (auto simp: fun_eq_iff box_def split: split_indicator) then have "emeasure lborel (box l u) = (\\<^sup>+x. (\b\Basis. indicator {l\b <..< u\b} (x \ b)) \lborel)" by simp also have "\ = (\b\Basis. (u - l) \ b)" by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left) finally show ?thesis . qed lemma emeasure_lborel_cbox_eq: "emeasure lborel (cbox l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le) lemma emeasure_lborel_box_eq: "emeasure lborel (box l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0" using emeasure_lborel_cbox[of x x] nonempty_Basis by (auto simp del: emeasure_lborel_cbox nonempty_Basis) +lemma emeasure_lborel_cbox_finite: "emeasure lborel (cbox a b) < \" + by (auto simp: emeasure_lborel_cbox_eq) + +lemma emeasure_lborel_box_finite: "emeasure lborel (box a b) < \" + by (auto simp: emeasure_lborel_box_eq) + +lemma emeasure_lborel_ball_finite: "emeasure lborel (ball c r) < \" +proof - + have "bounded (ball c r)" by simp + from bounded_subset_cbox_symmetric[OF this] obtain a where a: "ball c r \ cbox (-a) a" + by auto + hence "emeasure lborel (ball c r) \ emeasure lborel (cbox (-a) a)" + by (intro emeasure_mono) auto + also have "\ < \" by (simp add: emeasure_lborel_cbox_eq) + finally show ?thesis . +qed + +lemma emeasure_lborel_cball_finite: "emeasure lborel (cball c r) < \" +proof - + have "bounded (cball c r)" by simp + from bounded_subset_cbox_symmetric[OF this] obtain a where a: "cball c r \ cbox (-a) a" + by auto + hence "emeasure lborel (cball c r) \ emeasure lborel (cbox (-a) a)" + by (intro emeasure_mono) auto + also have "\ < \" by (simp add: emeasure_lborel_cbox_eq) + finally show ?thesis . +qed + lemma fmeasurable_cbox [iff]: "cbox a b \ fmeasurable lborel" and fmeasurable_box [iff]: "box a b \ fmeasurable lborel" by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq) lemma fixes l u :: real assumes [simp]: "l \ u" shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l" and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l" and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l" and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l" by (simp_all add: measure_def) lemma assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows measure_lborel_box[simp]: "measure lborel (box l u) = (\b\Basis. (u - l) \ b)" and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\b\Basis. (u - l) \ b)" by (simp_all add: measure_def inner_diff_left prod_nonneg) lemma measure_lborel_cbox_eq: "measure lborel (cbox l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le) lemma measure_lborel_box_eq: "measure lborel (box l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force lemma measure_lborel_singleton[simp]: "measure lborel {x} = 0" by (simp add: measure_def) lemma sigma_finite_lborel: "sigma_finite_measure lborel" proof show "\A::'a set set. countable A \ A \ sets lborel \ \A = space lborel \ (\a\A. emeasure lborel a \ \)" by (intro exI[of _ "range (\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One))"]) (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV) qed end lemma emeasure_lborel_UNIV [simp]: "emeasure lborel (UNIV::'a::euclidean_space set) = \" proof - { fix n::nat let ?Ba = "Basis :: 'a set" have "real n \ (2::real) ^ card ?Ba * real n" by (simp add: mult_le_cancel_right1) also have "... \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" apply (rule mult_left_mono) apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc) apply (simp) done finally have "real n \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" . } note [intro!] = this show ?thesis unfolding UN_box_eq_UNIV[symmetric] apply (subst SUP_emeasure_incseq[symmetric]) apply (auto simp: incseq_def subset_box inner_add_left simp del: Sup_eq_top_iff SUP_eq_top_iff intro!: ennreal_SUP_eq_top) done qed lemma emeasure_lborel_countable: fixes A :: "'a::euclidean_space set" assumes "countable A" shows "emeasure lborel A = 0" proof - have "A \ (\i. {from_nat_into A i})" using from_nat_into_surj assms by force then have "emeasure lborel A \ emeasure lborel (\i. {from_nat_into A i})" by (intro emeasure_mono) auto also have "emeasure lborel (\i. {from_nat_into A i}) = 0" by (rule emeasure_UN_eq_0) auto finally show ?thesis by (auto simp add: ) qed lemma countable_imp_null_set_lborel: "countable A \ A \ null_sets lborel" by (simp add: null_sets_def emeasure_lborel_countable sets.countable) lemma finite_imp_null_set_lborel: "finite A \ A \ null_sets lborel" by (intro countable_imp_null_set_lborel countable_finite) lemma insert_null_sets_iff [simp]: "insert a N \ null_sets lebesgue \ N \ null_sets lebesgue" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson completion.complete2 subset_insertI) next assume ?rhs then show ?lhs by (simp add: null_sets.insert_in_sets null_setsI) qed lemma insert_null_sets_lebesgue_on_iff [simp]: assumes "a \ S" "S \ sets lebesgue" shows "insert a N \ null_sets (lebesgue_on S) \ N \ null_sets (lebesgue_on S)" by (simp add: assms null_sets_restrict_space) lemma lborel_neq_count_space[simp]: "lborel \ count_space (A::('a::ordered_euclidean_space) set)" proof assume asm: "lborel = count_space A" have "space lborel = UNIV" by simp hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space) have "emeasure lborel {undefined::'a} = 1" by (subst asm, subst emeasure_count_space_finite) auto moreover have "emeasure lborel {undefined} \ 1" by simp ultimately show False by contradiction qed lemma mem_closed_if_AE_lebesgue_open: assumes "open S" "closed C" assumes "AE x \ S in lebesgue. x \ C" assumes "x \ S" shows "x \ C" proof (rule ccontr) assume xC: "x \ C" with openE[of "S - C"] assms obtain e where e: "0 < e" "ball x e \ S - C" by blast then obtain a b where box: "x \ box a b" "box a b \ S - C" by (metis rational_boxes order_trans) then have "0 < emeasure lebesgue (box a b)" by (auto simp: emeasure_lborel_box_eq mem_box algebra_simps intro!: prod_pos) also have "\ \ emeasure lebesgue (S - C)" using assms box by (auto intro!: emeasure_mono) also have "\ = 0" using assms by (auto simp: eventually_ae_filter completion.complete2 set_diff_eq null_setsD1) finally show False by simp qed lemma mem_closed_if_AE_lebesgue: "closed C \ (AE x in lebesgue. x \ C) \ x \ C" using mem_closed_if_AE_lebesgue_open[OF open_UNIV] by simp subsection \Affine transformation on the Lebesgue-Borel\ lemma\<^marker>\tag important\ lborel_eqI: fixes M :: "'a::euclidean_space measure" assumes emeasure_eq: "\l u. (\b. b \ Basis \ l \ b \ u \ b) \ emeasure M (box l u) = (\b\Basis. (u - l) \ b)" assumes sets_eq: "sets M = sets borel" shows "lborel = M" proof (rule measure_eqI_generator_eq) let ?E = "range (\(a, b). box a b::'a set)" show "Int_stable ?E" by (auto simp: Int_stable_def box_Int_box) show "?E \ Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E" by (simp_all add: borel_eq_box sets_eq) let ?A = "\n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set" show "range ?A \ ?E" "(\i. ?A i) = UNIV" unfolding UN_box_eq_UNIV by auto { fix i show "emeasure lborel (?A i) \ \" by auto } { fix X assume "X \ ?E" then show "emeasure lborel X = emeasure M X" apply (auto simp: emeasure_eq emeasure_lborel_box_eq) apply (subst box_eq_empty(1)[THEN iffD2]) apply (auto intro: less_imp_le simp: not_le) done } qed lemma\<^marker>\tag important\ lborel_affine_euclidean: fixes c :: "'a::euclidean_space \ real" and t defines "T x \ t + (\j\Basis. (c j * (x \ j)) *\<^sub>R j)" assumes c: "\j. j \ Basis \ c j \ 0" shows "lborel = density (distr lborel borel T) (\_. (\j\Basis. \c j\))" (is "_ = ?D") proof (rule lborel_eqI) let ?B = "Basis :: 'a set" fix l u assume le: "\b. b \ ?B \ l \ b \ u \ b" have [measurable]: "T \ borel \\<^sub>M borel" by (simp add: T_def[abs_def]) have eq: "T -` box l u = box (\j\Basis. (((if 0 < c j then l - t else u - t) \ j) / c j) *\<^sub>R j) (\j\Basis. (((if 0 < c j then u - t else l - t) \ j) / c j) *\<^sub>R j)" using c by (auto simp: box_def T_def field_simps inner_simps divide_less_eq) with le c show "emeasure ?D (box l u) = (\b\?B. (u - l) \ b)" by (auto simp: emeasure_density emeasure_distr nn_integral_multc emeasure_lborel_box_eq inner_simps field_split_simps ennreal_mult'[symmetric] prod_nonneg prod.distrib[symmetric] intro!: prod.cong) qed simp lemma lborel_affine: fixes t :: "'a::euclidean_space" shows "c \ 0 \ lborel = density (distr lborel borel (\x. t + c *\<^sub>R x)) (\_. \c\^DIM('a))" using lborel_affine_euclidean[where c="\_::'a. c" and t=t] unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation prod_constant by simp lemma lborel_real_affine: "c \ 0 \ lborel = density (distr lborel borel (\x. t + c * x)) (\_. ennreal (abs c))" using lborel_affine[of c t] by simp lemma AE_borel_affine: fixes P :: "real \ bool" shows "c \ 0 \ Measurable.pred borel P \ AE x in lborel. P x \ AE x in lborel. P (t + c * x)" by (subst lborel_real_affine[where t="- t / c" and c="1 / c"]) (simp_all add: AE_density AE_distr_iff field_simps) lemma nn_integral_real_affine: fixes c :: real assumes [measurable]: "f \ borel_measurable borel" and c: "c \ 0" shows "(\\<^sup>+x. f x \lborel) = \c\ * (\\<^sup>+x. f (t + c * x) \lborel)" by (subst lborel_real_affine[OF c, of t]) (simp add: nn_integral_density nn_integral_distr nn_integral_cmult) lemma lborel_integrable_real_affine: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "integrable lborel f" shows "c \ 0 \ integrable lborel (\x. f (t + c * x))" using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top) lemma lborel_integrable_real_affine_iff: fixes f :: "real \ 'a :: {banach, second_countable_topology}" shows "c \ 0 \ integrable lborel (\x. f (t + c * x)) \ integrable lborel f" using lborel_integrable_real_affine[of f c t] lborel_integrable_real_affine[of "\x. f (t + c * x)" "1/c" "-t/c"] by (auto simp add: field_simps) lemma\<^marker>\tag important\ lborel_integral_real_affine: fixes f :: "real \ 'a :: {banach, second_countable_topology}" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lborel) = \c\ *\<^sub>R (\x. f (t + c * x) \lborel)" proof cases assume f[measurable]: "integrable lborel f" then show ?thesis using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t] by (subst lborel_real_affine[OF c, of t]) (simp add: integral_density integral_distr) next assume "\ integrable lborel f" with c show ?thesis by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq) qed lemma fixes c :: "'a::euclidean_space \ real" and t assumes c: "\j. j \ Basis \ c j \ 0" defines "T == (\x. t + (\j\Basis. (c j * (x \ j)) *\<^sub>R j))" shows lebesgue_affine_euclidean: "lebesgue = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" (is "_ = ?D") and lebesgue_affine_measurable: "T \ lebesgue \\<^sub>M lebesgue" proof - have T_borel[measurable]: "T \ borel \\<^sub>M borel" by (auto simp: T_def[abs_def]) { fix A :: "'a set" assume A: "A \ sets borel" then have "emeasure lborel A = 0 \ emeasure (density (distr lborel borel T) (\_. (\j\Basis. \c j\))) A = 0" unfolding T_def using c by (subst lborel_affine_euclidean[symmetric]) auto also have "\ \ emeasure (distr lebesgue lborel T) A = 0" using A c by (simp add: distr_completion emeasure_density nn_integral_cmult prod_nonneg cong: distr_cong) finally have "emeasure lborel A = 0 \ emeasure (distr lebesgue lborel T) A = 0" . } then have eq: "null_sets lborel = null_sets (distr lebesgue lborel T)" by (auto simp: null_sets_def) show "T \ lebesgue \\<^sub>M lebesgue" by (rule completion.measurable_completion2) (auto simp: eq measurable_completion) have "lebesgue = completion (density (distr lborel borel T) (\_. (\j\Basis. \c j\)))" using c by (subst lborel_affine_euclidean[of c t]) (simp_all add: T_def[abs_def]) also have "\ = density (completion (distr lebesgue lborel T)) (\_. (\j\Basis. \c j\))" using c by (auto intro!: always_eventually prod_pos completion_density_eq simp: distr_completion cong: distr_cong) also have "\ = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" by (subst completion.completion_distr_eq) (auto simp: eq measurable_completion) finally show "lebesgue = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" . qed corollary lebesgue_real_affine: "c \ 0 \ lebesgue = density (distr lebesgue lebesgue (\x. t + c * x)) (\_. ennreal (abs c))" using lebesgue_affine_euclidean [where c= "\x::real. c"] by simp lemma nn_integral_real_affine_lebesgue: fixes c :: real assumes f[measurable]: "f \ borel_measurable lebesgue" and c: "c \ 0" shows "(\\<^sup>+x. f x \lebesgue) = ennreal\c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" proof - have "(\\<^sup>+x. f x \lebesgue) = (\\<^sup>+x. f x \density (distr lebesgue lebesgue (\x. t + c * x)) (\x. ennreal \c\))" using lebesgue_real_affine c by auto also have "\ = \\<^sup>+ x. ennreal \c\ * f x \distr lebesgue lebesgue (\x. t + c * x)" by (subst nn_integral_density) auto also have "\ = ennreal \c\ * integral\<^sup>N (distr lebesgue lebesgue (\x. t + c * x)) f" using f measurable_distr_eq1 nn_integral_cmult by blast also have "\ = \c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" using lebesgue_affine_measurable[where c= "\x::real. c"] by (subst nn_integral_distr) (force+) finally show ?thesis . qed lemma lebesgue_measurable_scaling[measurable]: "(*\<^sub>R) x \ lebesgue \\<^sub>M lebesgue" proof cases assume "x = 0" then have "(*\<^sub>R) x = (\x. 0::'a)" by (auto simp: fun_eq_iff) then show ?thesis by auto next assume "x \ 0" then show ?thesis using lebesgue_affine_measurable[of "\_. x" 0] unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation by (auto simp add: ac_simps) qed lemma fixes m :: real and \ :: "'a::euclidean_space" defines "T r d x \ r *\<^sub>R x + d" shows emeasure_lebesgue_affine: "emeasure lebesgue (T m \ ` S) = \m\ ^ DIM('a) * emeasure lebesgue S" (is ?e) and measure_lebesgue_affine: "measure lebesgue (T m \ ` S) = \m\ ^ DIM('a) * measure lebesgue S" (is ?m) proof - show ?e proof cases assume "m = 0" then show ?thesis by (simp add: image_constant_conv T_def[abs_def]) next let ?T = "T m \" and ?T' = "T (1 / m) (- ((1/m) *\<^sub>R \))" assume "m \ 0" then have s_comp_s: "?T' \ ?T = id" "?T \ ?T' = id" by (auto simp: T_def[abs_def] fun_eq_iff scaleR_add_right scaleR_diff_right) then have "inv ?T' = ?T" "bij ?T'" by (auto intro: inv_unique_comp o_bij) then have eq: "T m \ ` S = T (1 / m) ((-1/m) *\<^sub>R \) -` S \ space lebesgue" using bij_vimage_eq_inv_image[OF \bij ?T'\, of S] by auto have trans_eq_T: "(\x. \ + (\j\Basis. (m * (x \ j)) *\<^sub>R j)) = T m \" for m \ unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] by (auto simp add: euclidean_representation ac_simps) have T[measurable]: "T r d \ lebesgue \\<^sub>M lebesgue" for r d using lebesgue_affine_measurable[of "\_. r" d] by (cases "r = 0") (auto simp: trans_eq_T T_def[abs_def]) show ?thesis proof cases assume "S \ sets lebesgue" with \m \ 0\ show ?thesis unfolding eq apply (subst lebesgue_affine_euclidean[of "\_. m" \]) apply (simp_all add: emeasure_density trans_eq_T nn_integral_cmult emeasure_distr del: space_completion emeasure_completion) apply (simp add: vimage_comp s_comp_s) done next assume "S \ sets lebesgue" moreover have "?T ` S \ sets lebesgue" proof assume "?T ` S \ sets lebesgue" then have "?T -` (?T ` S) \ space lebesgue \ sets lebesgue" by (rule measurable_sets[OF T]) also have "?T -` (?T ` S) \ space lebesgue = S" by (simp add: vimage_comp s_comp_s eq) finally show False using \S \ sets lebesgue\ by auto qed ultimately show ?thesis by (simp add: emeasure_notin_sets) qed qed show ?m unfolding measure_def \?e\ by (simp add: enn2real_mult prod_nonneg) qed lemma lebesgue_real_scale: assumes "c \ 0" shows "lebesgue = density (distr lebesgue lebesgue (\x. c * x)) (\x. ennreal \c\)" using assms by (subst lebesgue_affine_euclidean[of "\_. c" 0]) simp_all lemma divideR_right: fixes x y :: "'a::real_normed_vector" shows "r \ 0 \ y = x /\<^sub>R r \ r *\<^sub>R y = x" using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp lemma lborel_has_bochner_integral_real_affine_iff: fixes x :: "'a :: {banach, second_countable_topology}" shows "c \ 0 \ has_bochner_integral lborel f x \ has_bochner_integral lborel (\x. f (t + c * x)) (x /\<^sub>R \c\)" unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)" by (subst lborel_real_affine[of "-1" 0]) (auto simp: density_1 one_ennreal_def[symmetric]) lemma lborel_distr_mult: assumes "(c::real) \ 0" shows "distr lborel borel ((*) c) = density lborel (\_. inverse \c\)" proof- have "distr lborel borel ((*) c) = distr lborel lborel ((*) c)" by (simp cong: distr_cong) also from assms have "... = density lborel (\_. inverse \c\)" by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr) finally show ?thesis . qed lemma lborel_distr_mult': assumes "(c::real) \ 0" shows "lborel = density (distr lborel borel ((*) c)) (\_. \c\)" proof- have "lborel = density lborel (\_. 1)" by (rule density_1[symmetric]) also from assms have "(\_. 1 :: ennreal) = (\_. inverse \c\ * \c\)" by (intro ext) simp also have "density lborel ... = density (density lborel (\_. inverse \c\)) (\_. \c\)" by (subst density_density_eq) (auto simp: ennreal_mult) also from assms have "density lborel (\_. inverse \c\) = distr lborel borel ((*) c)" by (rule lborel_distr_mult[symmetric]) finally show ?thesis . qed lemma lborel_distr_plus: fixes c :: "'a::euclidean_space" shows "distr lborel borel ((+) c) = lborel" by (subst lborel_affine[of 1 c], auto simp: density_1) interpretation lborel: sigma_finite_measure lborel by (rule sigma_finite_lborel) interpretation lborel_pair: pair_sigma_finite lborel lborel .. lemma\<^marker>\tag important\ lborel_prod: "lborel \\<^sub>M lborel = (lborel :: ('a::euclidean_space \ 'b::euclidean_space) measure)" proof (rule lborel_eqI[symmetric], clarify) fix la ua :: 'a and lb ub :: 'b assume lu: "\a b. (a, b) \ Basis \ (la, lb) \ (a, b) \ (ua, ub) \ (a, b)" have [simp]: "\b. b \ Basis \ la \ b \ ua \ b" "\b. b \ Basis \ lb \ b \ ub \ b" "inj_on (\u. (u, 0)) Basis" "inj_on (\u. (0, u)) Basis" "(\u. (u, 0)) ` Basis \ (\u. (0, u)) ` Basis = {}" "box (la, lb) (ua, ub) = box la ua \ box lb ub" using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def) show "emeasure (lborel \\<^sub>M lborel) (box (la, lb) (ua, ub)) = ennreal (prod ((\) ((ua, ub) - (la, lb))) Basis)" by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def prod.union_disjoint prod.reindex ennreal_mult inner_diff_left prod_nonneg) qed (simp add: borel_prod[symmetric]) (* FIXME: conversion in measurable prover *) lemma lborelD_Collect[measurable (raw)]: "{x\space borel. P x} \ sets borel \ {x\space lborel. P x} \ sets lborel" by simp lemma lborelD[measurable (raw)]: "A \ sets borel \ A \ sets lborel" by simp lemma emeasure_bounded_finite: assumes "bounded A" shows "emeasure lborel A < \" proof - obtain a b where "A \ cbox a b" by (meson bounded_subset_cbox_symmetric \bounded A\) then have "emeasure lborel A \ emeasure lborel (cbox a b)" by (intro emeasure_mono) auto then show ?thesis by (auto simp: emeasure_lborel_cbox_eq prod_nonneg less_top[symmetric] top_unique split: if_split_asm) qed lemma emeasure_compact_finite: "compact A \ emeasure lborel A < \" using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded) lemma borel_integrable_compact: fixes f :: "'a::euclidean_space \ 'b::{banach, second_countable_topology}" assumes "compact S" "continuous_on S f" shows "integrable lborel (\x. indicator S x *\<^sub>R f x)" proof cases assume "S \ {}" have "continuous_on S (\x. norm (f x))" using assms by (intro continuous_intros) from continuous_attains_sup[OF \compact S\ \S \ {}\ this] obtain M where M: "\x. x \ S \ norm (f x) \ M" by auto show ?thesis proof (rule integrable_bound) show "integrable lborel (\x. indicator S x * M)" using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lborel" using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact) show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \ norm (indicator S x * M)" by (auto split: split_indicator simp: abs_real_def dest!: M) qed qed simp lemma borel_integrable_atLeastAtMost: fixes f :: "real \ real" assumes f: "\x. a \ x \ x \ b \ isCont f x" shows "integrable lborel (\x. f x * indicator {a .. b} x)" (is "integrable _ ?f") proof - have "integrable lborel (\x. indicator {a .. b} x *\<^sub>R f x)" proof (rule borel_integrable_compact) from f show "continuous_on {a..b} f" by (auto intro: continuous_at_imp_continuous_on) qed simp then show ?thesis by (auto simp: mult.commute) qed subsection \Lebesgue measurable sets\ abbreviation\<^marker>\tag important\ lmeasurable :: "'a::euclidean_space set set" where "lmeasurable \ fmeasurable lebesgue" lemma not_measurable_UNIV [simp]: "UNIV \ lmeasurable" by (simp add: fmeasurable_def) lemma\<^marker>\tag important\ lmeasurable_iff_integrable: "S \ lmeasurable \ integrable lebesgue (indicator S :: 'a::euclidean_space \ real)" by (auto simp: fmeasurable_def integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator) lemma lmeasurable_cbox [iff]: "cbox a b \ lmeasurable" and lmeasurable_box [iff]: "box a b \ lmeasurable" by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq) lemma fixes a::real shows lmeasurable_interval [iff]: "{a..b} \ lmeasurable" "{a<.. lmeasurable" apply (metis box_real(2) lmeasurable_cbox) by (metis box_real(1) lmeasurable_box) lemma fmeasurable_compact: "compact S \ S \ fmeasurable lborel" using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact) lemma lmeasurable_compact: "compact S \ S \ lmeasurable" using fmeasurable_compact by (force simp: fmeasurable_def) lemma measure_frontier: "bounded S \ measure lebesgue (frontier S) = measure lebesgue (closure S) - measure lebesgue (interior S)" using closure_subset interior_subset by (auto simp: frontier_def fmeasurable_compact intro!: measurable_measure_Diff) lemma lmeasurable_closure: "bounded S \ closure S \ lmeasurable" by (simp add: lmeasurable_compact) lemma lmeasurable_frontier: "bounded S \ frontier S \ lmeasurable" by (simp add: compact_frontier_bounded lmeasurable_compact) lemma lmeasurable_open: "bounded S \ open S \ S \ lmeasurable" using emeasure_bounded_finite[of S] by (intro fmeasurableI) (auto simp: borel_open) lemma lmeasurable_ball [simp]: "ball a r \ lmeasurable" by (simp add: lmeasurable_open) lemma lmeasurable_cball [simp]: "cball a r \ lmeasurable" by (simp add: lmeasurable_compact) lemma lmeasurable_interior: "bounded S \ interior S \ lmeasurable" by (simp add: bounded_interior lmeasurable_open) lemma null_sets_cbox_Diff_box: "cbox a b - box a b \ null_sets lborel" proof - have "emeasure lborel (cbox a b - box a b) = 0" by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox) then have "cbox a b - box a b \ null_sets lborel" by (auto simp: null_sets_def) then show ?thesis by (auto dest!: AE_not_in) qed lemma bounded_set_imp_lmeasurable: assumes "bounded S" "S \ sets lebesgue" shows "S \ lmeasurable" by (metis assms bounded_Un emeasure_bounded_finite emeasure_completion fmeasurableI main_part_null_part_Un) lemma finite_measure_lebesgue_on: "S \ lmeasurable \ finite_measure (lebesgue_on S)" by (auto simp: finite_measureI fmeasurable_def emeasure_restrict_space) lemma integrable_const_ivl [iff]: fixes a::"'a::ordered_euclidean_space" shows "integrable (lebesgue_on {a..b}) (\x. c)" by (metis cbox_interval finite_measure.integrable_const finite_measure_lebesgue_on lmeasurable_cbox) subsection\<^marker>\tag unimportant\\Translation preserves Lebesgue measure\ lemma sigma_sets_image: assumes S: "S \ sigma_sets \ M" and "M \ Pow \" "f ` \ = \" "inj_on f \" and M: "\y. y \ M \ f ` y \ M" shows "(f ` S) \ sigma_sets \ M" using S proof (induct S rule: sigma_sets.induct) case (Basic a) then show ?case by (simp add: M) next case Empty then show ?case by (simp add: sigma_sets.Empty) next case (Compl a) then have "\ - a \ \" "a \ \" by (auto simp: sigma_sets_into_sp [OF \M \ Pow \\]) then show ?case by (auto simp: inj_on_image_set_diff [OF \inj_on f \\] assms intro: Compl sigma_sets.Compl) next case (Union a) then show ?case by (metis image_UN sigma_sets.simps) qed lemma null_sets_translation: assumes "N \ null_sets lborel" shows "{x. x - a \ N} \ null_sets lborel" proof - have [simp]: "(\x. x + a) ` N = {x. x - a \ N}" by force show ?thesis using assms emeasure_lebesgue_affine [of 1 a N] by (auto simp: null_sets_def) qed lemma lebesgue_sets_translation: fixes f :: "'a \ 'a::euclidean_space" assumes S: "S \ sets lebesgue" shows "((\x. a + x) ` S) \ sets lebesgue" proof - have im_eq: "(+) a ` A = {x. x - a \ A}" for A by force have "((\x. a + x) ` S) = ((\x. -a + x) -` S) \ (space lebesgue)" using image_iff by fastforce also have "\ \ sets lebesgue" proof (rule measurable_sets [OF measurableI assms]) fix A :: "'b set" assume A: "A \ sets lebesgue" have vim_eq: "(\x. x - a) -` A = (+) a ` A" for A by force have "\s n N'. (+) a ` (S \ N) = s \ n \ s \ sets borel \ N' \ null_sets lborel \ n \ N'" if "S \ sets borel" and "N' \ null_sets lborel" and "N \ N'" for S N N' proof (intro exI conjI) show "(+) a ` (S \ N) = (\x. a + x) ` S \ (\x. a + x) ` N" by auto show "(\x. a + x) ` N' \ null_sets lborel" using that by (auto simp: null_sets_translation im_eq) qed (use that im_eq in auto) with A have "(\x. x - a) -` A \ sets lebesgue" by (force simp: vim_eq completion_def intro!: sigma_sets_image) then show "(+) (- a) -` A \ space lebesgue \ sets lebesgue" by (auto simp: vimage_def im_eq) qed auto finally show ?thesis . qed lemma measurable_translation: "S \ lmeasurable \ ((+) a ` S) \ lmeasurable" using emeasure_lebesgue_affine [of 1 a S] apply (auto intro: lebesgue_sets_translation simp add: fmeasurable_def cong: image_cong_simp) apply (simp add: ac_simps) done lemma measurable_translation_subtract: "S \ lmeasurable \ ((\x. x - a) ` S) \ lmeasurable" using measurable_translation [of S "- a"] by (simp cong: image_cong_simp) lemma measure_translation: "measure lebesgue ((+) a ` S) = measure lebesgue S" using measure_lebesgue_affine [of 1 a S] by (simp add: ac_simps cong: image_cong_simp) lemma measure_translation_subtract: "measure lebesgue ((\x. x - a) ` S) = measure lebesgue S" using measure_translation [of "- a"] by (simp cong: image_cong_simp) subsection \A nice lemma for negligibility proofs\ lemma summable_iff_suminf_neq_top: "(\n. f n \ 0) \ \ summable f \ (\i. ennreal (f i)) = top" by (metis summable_suminf_not_top) proposition\<^marker>\tag important\ starlike_negligible_bounded_gmeasurable: fixes S :: "'a :: euclidean_space set" assumes S: "S \ sets lebesgue" and "bounded S" and eq1: "\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1" shows "S \ null_sets lebesgue" proof - obtain M where "0 < M" "S \ ball 0 M" using \bounded S\ by (auto dest: bounded_subset_ballD) let ?f = "\n. root DIM('a) (Suc n)" have vimage_eq_image: "(*\<^sub>R) (?f n) -` S = (*\<^sub>R) (1 / ?f n) ` S" for n apply safe subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto subgoal by auto done have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n by (simp add: field_simps) { fix n x assume x: "root DIM('a) (1 + real n) *\<^sub>R x \ S" have "1 * norm x \ root DIM('a) (1 + real n) * norm x" by (rule mult_mono) auto also have "\ < M" using x \S \ ball 0 M\ by auto finally have "norm x < M" by simp } note less_M = this have "(\n. ennreal (1 / Suc n)) = top" using not_summable_harmonic[where 'a=real] summable_Suc_iff[where f="\n. 1 / (real n)"] by (intro summable_iff_suminf_neq_top) (auto simp add: inverse_eq_divide) then have "top * emeasure lebesgue S = (\n. (1 / ?f n)^DIM('a) * emeasure lebesgue S)" unfolding ennreal_suminf_multc eq by simp also have "\ = (\n. emeasure lebesgue ((*\<^sub>R) (?f n) -` S))" unfolding vimage_eq_image using emeasure_lebesgue_affine[of "1 / ?f n" 0 S for n] by simp also have "\ = emeasure lebesgue (\n. (*\<^sub>R) (?f n) -` S)" proof (intro suminf_emeasure) show "disjoint_family (\n. (*\<^sub>R) (?f n) -` S)" unfolding disjoint_family_on_def proof safe fix m n :: nat and x assume "m \ n" "?f m *\<^sub>R x \ S" "?f n *\<^sub>R x \ S" with eq1[of "?f m / ?f n" "?f n *\<^sub>R x"] show "x \ {}" by auto qed have "(*\<^sub>R) (?f i) -` S \ sets lebesgue" for i using measurable_sets[OF lebesgue_measurable_scaling[of "?f i"] S] by auto then show "range (\i. (*\<^sub>R) (?f i) -` S) \ sets lebesgue" by auto qed also have "\ \ emeasure lebesgue (ball 0 M :: 'a set)" using less_M by (intro emeasure_mono) auto also have "\ < top" using lmeasurable_ball by (auto simp: fmeasurable_def) finally have "emeasure lebesgue S = 0" by (simp add: ennreal_top_mult split: if_split_asm) then show "S \ null_sets lebesgue" unfolding null_sets_def using \S \ sets lebesgue\ by auto qed corollary starlike_negligible_compact: "compact S \ (\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1) \ S \ null_sets lebesgue" using starlike_negligible_bounded_gmeasurable[of S] by (auto simp: compact_eq_bounded_closed) proposition outer_regular_lborel_le: assumes B[measurable]: "B \ sets borel" and "0 < (e::real)" obtains U where "open U" "B \ U" and "emeasure lborel (U - B) \ e" proof - let ?\ = "emeasure lborel" let ?B = "\n::nat. ball 0 n :: 'a set" let ?e = "\n. e*((1/2)^Suc n)" have "\n. \U. open U \ ?B n \ B \ U \ ?\ (U - B) < ?e n" proof fix n :: nat let ?A = "density lborel (indicator (?B n))" have emeasure_A: "X \ sets borel \ emeasure ?A X = ?\ (?B n \ X)" for X by (auto simp: emeasure_density borel_measurable_indicator indicator_inter_arith[symmetric]) have finite_A: "emeasure ?A (space ?A) \ \" using emeasure_bounded_finite[of "?B n"] by (auto simp: emeasure_A) interpret A: finite_measure ?A by rule fact have "emeasure ?A B + ?e n > (INF U\{U. B \ U \ open U}. emeasure ?A U)" using \0 by (auto simp: outer_regular[OF _ finite_A B, symmetric]) then obtain U where U: "B \ U" "open U" and muU: "?\ (?B n \ B) + ?e n > ?\ (?B n \ U)" unfolding INF_less_iff by (auto simp: emeasure_A) moreover { have "?\ ((?B n \ U) - B) = ?\ ((?B n \ U) - (?B n \ B))" using U by (intro arg_cong[where f="?\"]) auto also have "\ = ?\ (?B n \ U) - ?\ (?B n \ B)" using U A.emeasure_finite[of B] by (intro emeasure_Diff) (auto simp del: A.emeasure_finite simp: emeasure_A) also have "\ < ?e n" using U muU A.emeasure_finite[of B] by (subst minus_less_iff_ennreal) (auto simp del: A.emeasure_finite simp: emeasure_A less_top ac_simps intro!: emeasure_mono) finally have "?\ ((?B n \ U) - B) < ?e n" . } ultimately show "\U. open U \ ?B n \ B \ U \ ?\ (U - B) < ?e n" by (intro exI[of _ "?B n \ U"]) auto qed then obtain U where U: "\n. open (U n)" "\n. ?B n \ B \ U n" "\n. ?\ (U n - B) < ?e n" by metis show ?thesis proof { fix x assume "x \ B" moreover obtain n where "norm x < real n" using reals_Archimedean2 by blast ultimately have "x \ (\n. U n)" using U(2)[of n] by auto } note * = this then show "open (\n. U n)" "B \ (\n. U n)" using U by auto have "?\ (\n. U n - B) \ (\n. ?\ (U n - B))" using U(1) by (intro emeasure_subadditive_countably) auto also have "\ \ (\n. ennreal (?e n))" using U(3) by (intro suminf_le) (auto intro: less_imp_le) also have "\ = ennreal (e * 1)" using \0 by (intro suminf_ennreal_eq sums_mult power_half_series) auto finally show "emeasure lborel ((\n. U n) - B) \ ennreal e" by simp qed qed lemma\<^marker>\tag important\ outer_regular_lborel: assumes B: "B \ sets borel" and "0 < (e::real)" obtains U where "open U" "B \ U" "emeasure lborel (U - B) < e" proof - obtain U where U: "open U" "B \ U" and "emeasure lborel (U-B) \ e/2" using outer_regular_lborel_le [OF B, of "e/2"] \e > 0\ by force moreover have "ennreal (e/2) < ennreal e" using \e > 0\ by (simp add: ennreal_lessI) ultimately have "emeasure lborel (U-B) < e" by auto with U show ?thesis using that by auto qed lemma completion_upper: assumes A: "A \ sets (completion M)" obtains A' where "A \ A'" "A' \ sets M" "A' - A \ null_sets (completion M)" "emeasure (completion M) A = emeasure M A'" proof - from AE_notin_null_part[OF A] obtain N where N: "N \ null_sets M" "null_part M A \ N" unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto let ?A' = "main_part M A \ N" show ?thesis proof show "A \ ?A'" using \null_part M A \ N\ by (subst main_part_null_part_Un[symmetric, OF A]) auto have "main_part M A \ A" using assms main_part_null_part_Un by auto then have "?A' - A \ N" by blast with N show "?A' - A \ null_sets (completion M)" by (blast intro: null_sets_completionI completion.complete_measure_axioms complete_measure.complete2) show "emeasure (completion M) A = emeasure M (main_part M A \ N)" using A \N \ null_sets M\ by (simp add: emeasure_Un_null_set) qed (use A N in auto) qed lemma sets_lebesgue_outer_open: fixes e::real assumes S: "S \ sets lebesgue" and "e > 0" obtains T where "open T" "S \ T" "(T - S) \ lmeasurable" "emeasure lebesgue (T - S) < ennreal e" proof - obtain S' where S': "S \ S'" "S' \ sets borel" and null: "S' - S \ null_sets lebesgue" and em: "emeasure lebesgue S = emeasure lborel S'" using completion_upper[of S lborel] S by auto then have f_S': "S' \ sets borel" using S by (auto simp: fmeasurable_def) with outer_regular_lborel[OF _ \0] obtain U where U: "open U" "S' \ U" "emeasure lborel (U - S') < e" by blast show thesis proof show "open U" "S \ U" using f_S' U S' by auto have "(U - S) = (U - S') \ (S' - S)" using S' U by auto then have eq: "emeasure lebesgue (U - S) = emeasure lborel (U - S')" using null by (simp add: U(1) emeasure_Un_null_set f_S' sets.Diff) have "(U - S) \ sets lebesgue" by (simp add: S U(1) sets.Diff) then show "(U - S) \ lmeasurable" unfolding fmeasurable_def using U(3) eq less_le_trans by fastforce with eq U show "emeasure lebesgue (U - S) < e" by (simp add: eq) qed qed lemma sets_lebesgue_inner_closed: fixes e::real assumes "S \ sets lebesgue" "e > 0" obtains T where "closed T" "T \ S" "S-T \ lmeasurable" "emeasure lebesgue (S - T) < ennreal e" proof - have "-S \ sets lebesgue" using assms by (simp add: Compl_in_sets_lebesgue) then obtain T where "open T" "-S \ T" and T: "(T - -S) \ lmeasurable" "emeasure lebesgue (T - -S) < e" using sets_lebesgue_outer_open assms by blast show thesis proof show "closed (-T)" using \open T\ by blast show "-T \ S" using \- S \ T\ by auto show "S - ( -T) \ lmeasurable" "emeasure lebesgue (S - (- T)) < e" using T by (auto simp: Int_commute) qed qed lemma lebesgue_openin: "\openin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_open openin_open sets.Int sets_completionI_sets sets_lborel) lemma lebesgue_closedin: "\closedin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_closed closedin_closed sets.Int sets_completionI_sets sets_lborel) subsection\\F_sigma\ and \G_delta\ sets.\ \ \\<^url>\https://en.wikipedia.org/wiki/F-sigma_set\\ inductive\<^marker>\tag important\ fsigma :: "'a::topological_space set \ bool" where "(\n::nat. closed (F n)) \ fsigma (\(F ` UNIV))" inductive\<^marker>\tag important\ gdelta :: "'a::topological_space set \ bool" where "(\n::nat. open (F n)) \ gdelta (\(F ` UNIV))" lemma fsigma_Union_compact: fixes S :: "'a::{real_normed_vector,heine_borel} set" shows "fsigma S \ (\F::nat \ 'a set. range F \ Collect compact \ S = \(F ` UNIV))" proof safe assume "fsigma S" then obtain F :: "nat \ 'a set" where F: "range F \ Collect closed" "S = \(F ` UNIV)" by (meson fsigma.cases image_subsetI mem_Collect_eq) then have "\D::nat \ 'a set. range D \ Collect compact \ \(D ` UNIV) = F i" for i using closed_Union_compact_subsets [of "F i"] by (metis image_subsetI mem_Collect_eq range_subsetD) then obtain D :: "nat \ nat \ 'a set" where D: "\i. range (D i) \ Collect compact \ \((D i) ` UNIV) = F i" by metis let ?DD = "\n. (\(i,j). D i j) (prod_decode n)" show "\F::nat \ 'a set. range F \ Collect compact \ S = \(F ` UNIV)" proof (intro exI conjI) show "range ?DD \ Collect compact" using D by clarsimp (metis mem_Collect_eq rangeI split_conv subsetCE surj_pair) show "S = \ (range ?DD)" proof show "S \ \ (range ?DD)" using D F by clarsimp (metis UN_iff old.prod.case prod_decode_inverse prod_encode_eq) show "\ (range ?DD) \ S" using D F by fastforce qed qed next fix F :: "nat \ 'a set" assume "range F \ Collect compact" and "S = \(F ` UNIV)" then show "fsigma (\(F ` UNIV))" by (simp add: compact_imp_closed fsigma.intros image_subset_iff) qed lemma gdelta_imp_fsigma: "gdelta S \ fsigma (- S)" proof (induction rule: gdelta.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" by auto then show ?case by (simp add: fsigma.intros closed_Compl 1) qed lemma fsigma_imp_gdelta: "fsigma S \ gdelta (- S)" proof (induction rule: fsigma.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" by auto then show ?case by (simp add: 1 gdelta.intros open_closed) qed lemma gdelta_complement: "gdelta(- S) \ fsigma S" using fsigma_imp_gdelta gdelta_imp_fsigma by force lemma lebesgue_set_almost_fsigma: assumes "S \ sets lebesgue" obtains C T where "fsigma C" "T \ null_sets lebesgue" "C \ T = S" "disjnt C T" proof - { fix n::nat obtain T where "closed T" "T \ S" "S-T \ lmeasurable" "emeasure lebesgue (S - T) < ennreal (1 / Suc n)" using sets_lebesgue_inner_closed [OF assms] by (metis of_nat_0_less_iff zero_less_Suc zero_less_divide_1_iff) then have "\T. closed T \ T \ S \ S - T \ lmeasurable \ measure lebesgue (S-T) < 1 / Suc n" by (metis emeasure_eq_measure2 ennreal_leI not_le) } then obtain F where F: "\n::nat. closed (F n) \ F n \ S \ S - F n \ lmeasurable \ measure lebesgue (S - F n) < 1 / Suc n" by metis let ?C = "\(F ` UNIV)" show thesis proof show "fsigma ?C" using F by (simp add: fsigma.intros) show "(S - ?C) \ null_sets lebesgue" proof (clarsimp simp add: completion.null_sets_outer_le) fix e :: "real" assume "0 < e" then obtain n where n: "1 / Suc n < e" using nat_approx_posE by metis show "\T \ lmeasurable. S - (\x. F x) \ T \ measure lebesgue T \ e" proof (intro bexI conjI) show "measure lebesgue (S - F n) \ e" by (meson F n less_trans not_le order.asym) qed (use F in auto) qed show "?C \ (S - ?C) = S" using F by blast show "disjnt ?C (S - ?C)" by (auto simp: disjnt_def) qed qed lemma lebesgue_set_almost_gdelta: assumes "S \ sets lebesgue" obtains C T where "gdelta C" "T \ null_sets lebesgue" "S \ T = C" "disjnt S T" proof - have "-S \ sets lebesgue" using assms Compl_in_sets_lebesgue by blast then obtain C T where C: "fsigma C" "T \ null_sets lebesgue" "C \ T = -S" "disjnt C T" using lebesgue_set_almost_fsigma by metis show thesis proof show "gdelta (-C)" by (simp add: \fsigma C\ fsigma_imp_gdelta) show "S \ T = -C" "disjnt S T" using C by (auto simp: disjnt_def) qed (use C in auto) qed end diff --git a/src/HOL/Analysis/Vitali_Covering_Theorem.thy b/src/HOL/Analysis/Vitali_Covering_Theorem.thy --- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy +++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy @@ -1,657 +1,658 @@ (* Title: HOL/Analysis/Vitali_Covering_Theorem.thy Authors: LC Paulson, based on material from HOL Light *) section \Vitali Covering Theorem and an Application to Negligibility\ theory Vitali_Covering_Theorem - imports Ball_Volume "HOL-Library.Permutations" + imports Equivalence_Lebesgue_Henstock_Integration "HOL-Library.Permutations" begin lemma stretch_Galois: fixes x :: "real^'n" shows "(\k. m k \ 0) \ ((y = (\ k. m k * x$k)) \ (\ k. y$k / m k) = x)" by auto lemma lambda_swap_Galois: "(x = (\ i. y $ Fun.swap m n id i) \ (\ i. x $ Fun.swap m n id i) = y)" by (auto; simp add: pointfree_idE vec_eq_iff) lemma lambda_add_Galois: fixes x :: "real^'n" shows "m \ n \ (x = (\ i. if i = m then y$m + y$n else y$i) \ (\ i. if i = m then x$m - x$n else x$i) = y)" by (safe; simp add: vec_eq_iff) lemma Vitali_covering_lemma_cballs_balls: fixes a :: "'a \ 'b::euclidean_space" assumes "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" proof (cases "K = {}") case True with that show ?thesis by auto next case False then have "B > 0" using assms less_le_trans by auto have rgt0[simp]: "\i. i \ K \ 0 < r i" using assms by auto let ?djnt = "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j)))" have "\C. \n. (C n \ K \ (\i \ C n. B/2 ^ n \ r i) \ ?djnt (C n) \ (\i \ K. B/2 ^ n < r i \ (\j. j \ C n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)))) \ (C n \ C(Suc n))" proof (rule dependent_nat_choice, safe) fix C n define D where "D \ {i \ K. B/2 ^ Suc n < r i \ (\j\C. disjnt (cball(a i)(r i)) (cball (a j) (r j)))}" let ?cover_ar = "\i j. \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" assume "C \ K" and Ble: "\i\C. B/2 ^ n \ r i" and djntC: "?djnt C" and cov_n: "\i\K. B/2 ^ n < r i \ (\j. j \ C \ ?cover_ar i j)" have *: "\C\chains {C. C \ D \ ?djnt C}. \C \ {C. C \ D \ ?djnt C}" proof (clarsimp simp: chains_def) fix C assume C: "C \ {C. C \ D \ ?djnt C}" and "chain\<^sub>\ C" show "\C \ D \ ?djnt (\C)" unfolding pairwise_def proof (intro ballI conjI impI) show "\C \ D" using C by blast next fix x y assume "x \ \C" and "y \ \C" and "x \ y" then obtain X Y where XY: "x \ X" "X \ C" "y \ Y" "Y \ C" by blast then consider "X \ Y" | "Y \ X" by (meson \chain\<^sub>\ C\ chain_subset_def) then show "disjnt (cball (a x) (r x)) (cball (a y) (r y))" proof cases case 1 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast next case 2 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast qed qed qed obtain E where "E \ D" and djntE: "?djnt E" and maximalE: "\X. \X \ D; ?djnt X; E \ X\ \ X = E" using Zorn_Lemma [OF *] by safe blast show "\L. (L \ K \ (\i\L. B/2 ^ Suc n \ r i) \ ?djnt L \ (\i\K. B/2 ^ Suc n < r i \ (\j. j \ L \ ?cover_ar i j))) \ C \ L" proof (intro exI conjI ballI) show "C \ E \ K" using D_def \C \ K\ \E \ D\ by blast show "B/2 ^ Suc n \ r i" if i: "i \ C \ E" for i using i proof assume "i \ C" have "B/2 ^ Suc n \ B/2 ^ n" using \B > 0\ by (simp add: field_split_simps) also have "\ \ r i" using Ble \i \ C\ by blast finally show ?thesis . qed (use D_def \E \ D\ in auto) show "?djnt (C \ E)" using D_def \C \ K\ \E \ D\ djntC djntE unfolding pairwise_def disjnt_def by blast next fix i assume "i \ K" show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ ?cover_ar i j)" proof (cases "r i \ B/2^n") case False then show ?thesis using cov_n \i \ K\ by auto next case True have "cball (a i) (r i) \ ball (a j) (5 * r j)" if less: "B/2 ^ Suc n < r i" and j: "j \ C \ E" and nondis: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" for j proof - obtain x where x: "dist (a i) x \ r i" "dist (a j) x \ r j" using nondis by (force simp: disjnt_def) have "dist (a i) (a j) \ dist (a i) x + dist x (a j)" by (simp add: dist_triangle) also have "\ \ r i + r j" by (metis add_mono_thms_linordered_semiring(1) dist_commute x) finally have aij: "dist (a i) (a j) + r i < 5 * r j" if "r i < 2 * r j" using that by auto show ?thesis using j proof assume "j \ C" have "B/2^n < 2 * r j" using Ble True \j \ C\ less by auto with aij True show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) next assume "j \ E" then have "B/2 ^ n < 2 * r j" using D_def \E \ D\ by auto with True have "r i < 2 * r j" by auto with aij show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) qed qed moreover have "\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j))" if "B/2 ^ Suc n < r i" proof (rule classical) assume NON: "\ ?thesis" show ?thesis proof (cases "i \ D") case True have "insert i E = E" proof (rule maximalE) show "insert i E \ D" by (simp add: True \E \ D\) show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (insert i E)" using False NON by (auto simp: pairwise_insert djntE disjnt_sym) qed auto then show ?thesis using \i \ K\ assms by fastforce next case False with that show ?thesis by (auto simp: D_def disjnt_def \i \ K\) qed qed ultimately show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j))" by blast qed qed auto qed (use assms in force) then obtain F where FK: "\n. F n \ K" and Fle: "\n i. i \ F n \ B/2 ^ n \ r i" and Fdjnt: "\n. ?djnt (F n)" and FF: "\n i. \i \ K; B/2 ^ n < r i\ \ \j. j \ F n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" and inc: "\n. F n \ F(Suc n)" by (force simp: all_conj_distrib) show thesis proof have *: "countable I" if "I \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) I" for I proof - show ?thesis proof (rule countable_image_inj_on [of "\i. cball(a i)(r i)"]) show "countable ((\i. cball (a i) (r i)) ` I)" proof (rule countable_disjoint_nonempty_interior_subsets) show "disjoint ((\i. cball (a i) (r i)) ` I)" by (auto simp: dest: pairwiseD [OF pw] intro: pairwise_imageI) show "\S. \S \ (\i. cball (a i) (r i)) ` I; interior S = {}\ \ S = {}" using \I \ K\ by (auto simp: not_less [symmetric]) qed next have "\x y. \x \ I; y \ I; a x = a y; r x = r y\ \ x = y" using pw \I \ K\ assms apply (clarsimp simp: pairwise_def disjnt_def) by (metis assms centre_in_cball subsetD empty_iff inf.idem less_eq_real_def) then show "inj_on (\i. cball (a i) (r i)) I" using \I \ K\ by (fastforce simp: inj_on_def cball_eq_cball_iff dest: assms) qed qed show "(Union(range F)) \ K" using FK by blast moreover show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (Union(range F))" proof (rule pairwise_chain_Union) show "chain\<^sub>\ (range F)" unfolding chain_subset_def by clarify (meson inc lift_Suc_mono_le linear subsetCE) qed (use Fdjnt in blast) ultimately show "countable (Union(range F))" by (blast intro: *) next fix i assume "i \ K" then obtain n where "(1/2) ^ n < r i / B" using \B > 0\ assms real_arch_pow_inv by fastforce then have B2: "B/2 ^ n < r i" using \B > 0\ by (simp add: field_split_simps) have "0 < r i" "r i \ B" by (auto simp: \i \ K\ assms) show "\j. j \ (Union(range F)) \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" using FF [OF \i \ K\ B2] by auto qed qed subsection\Vitali covering theorem\ lemma Vitali_covering_lemma_cballs: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. cball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "S \ (\i\C. cball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. cball (a i) (r i)) \ (\i\C. cball (a i) (5 * r i))" using cov subset_iff by fastforce with S show "S \ (\i\C. cball (a i) (5 * r i))" by blast qed (use C in auto) qed lemma Vitali_covering_lemma_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. ball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "S \ (\i\C. ball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. ball (a i) (r i)) \ (\i\C. ball (a i) (5 * r i))" using cov subset_iff by clarsimp (meson less_imp_le mem_ball mem_cball subset_eq) with S show "S \ (\i\C. ball (a i) (5 * r i))" by blast show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" using pw by (clarsimp simp: pairwise_def) (meson ball_subset_cball disjnt_subset1 disjnt_subset2) qed (use C in auto) qed theorem Vitali_covering_theorem_cballs: fixes a :: "'a \ 'n::euclidean_space" assumes r: "\i. i \ K \ 0 < r i" and S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "negligible(S - (\i \ C. cball (a i) (r i)))" proof - let ?\ = "measure lebesgue" have *: "\C. countable C \ C \ K \ pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C \ negligible(S - (\i \ C. cball (a i) (r i)))" if r01: "\i. i \ K \ 0 < r i \ r i \ 1" and Sd: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" for K r and a :: "'a \ 'n" proof - obtain C where C: "countable C" "C \ K" and pwC: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [of K r 1 a]) (auto simp: r01) have ar_injective: "\x y. \x \ C; y \ C; a x = a y; r x = r y\ \ x = y" using \C \ K\ pwC cov by (force simp: pairwise_def disjnt_def) show ?thesis proof (intro exI conjI) show "negligible (S - (\i\C. cball (a i) (r i)))" proof (clarsimp simp: negligible_on_intervals [of "S-T" for T]) fix l u show "negligible ((S - (\i\C. cball (a i) (r i))) \ cbox l u)" unfolding negligible_outer_le proof (intro allI impI) fix e::real assume "e > 0" define D where "D \ {i \ C. \ disjnt (ball(a i) (5 * r i)) (cbox l u)}" then have "D \ C" by auto have "countable D" unfolding D_def using \countable C\ by simp have UD: "(\i\D. cball (a i) (r i)) \ lmeasurable" proof (rule fmeasurableI2) show "cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One) \ lmeasurable" by blast have "y \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" if "i \ C" and x: "x \ cbox l u" and ai: "dist (a i) y \ r i" "dist (a i) x < 5 * r i" for i x y proof - have d6: "dist y x < 6 * r i" using dist_triangle3 [of y x "a i"] that by linarith show ?thesis proof (clarsimp simp: mem_box algebra_simps) fix j::'n assume j: "j \ Basis" then have xyj: "\x \ j - y \ j\ \ dist y x" by (metis Basis_le_norm dist_commute dist_norm inner_diff_left) have "l \ j \ x \ j" using \j \ Basis\ mem_box \x \ cbox l u\ by blast also have "\ \ y \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ y \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have l: "l \ j \ y \ j + 6" . have "y \ j \ x \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ u \ j + 6 * r i" using j x by (auto simp: mem_box) also have "\ \ u \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have u: "y \ j \ u \ j + 6" . show "l \ j \ y \ j + 6 \ y \ j \ u \ j + 6" using l u by blast qed qed then show "(\i\D. cball (a i) (r i)) \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" by (force simp: D_def disjnt_def) show "(\i\D. cball (a i) (r i)) \ sets lebesgue" using \countable D\ by auto qed obtain D1 where "D1 \ D" "finite D1" and measD1: "?\ (\i\D. cball (a i) (r i)) - e / 5 ^ DIM('n) < ?\ (\i\D1. cball (a i) (r i))" proof (rule measure_countable_Union_approachable [where e = "e / 5 ^ (DIM('n))"]) show "countable ((\i. cball (a i) (r i)) ` D)" using \countable D\ by auto show "\d. d \ (\i. cball (a i) (r i)) ` D \ d \ lmeasurable" by auto show "\D'. \D' \ (\i. cball (a i) (r i)) ` D; finite D'\ \ ?\ (\D') \ ?\ (\i\D. cball (a i) (r i))" by (fastforce simp add: intro!: measure_mono_fmeasurable UD) qed (use \e > 0\ in \auto dest: finite_subset_image\) show "\T. (S - (\i\C. cball (a i) (r i))) \ cbox l u \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "(S - (\i\C. cball (a i) (r i))) \ cbox l u \ (\i\D - D1. ball (a i) (5 * r i))" proof clarify fix x assume x: "x \ cbox l u" "x \ S" "x \ (\i\C. cball (a i) (r i))" have "closed (\i\D1. cball (a i) (r i))" using \finite D1\ by blast moreover have "x \ (\j\D1. cball (a j) (r j))" using x \D1 \ D\ unfolding D_def by blast ultimately obtain q where "q > 0" and q: "ball x q \ - (\i\D1. cball (a i) (r i))" by (metis (no_types, lifting) ComplI open_contains_ball closed_def) obtain i where "i \ K" and xi: "x \ cball (a i) (r i)" and ri: "r i < q/2" using Sd [OF \x \ S\] \q > 0\ half_gt_zero by blast then obtain j where "j \ C" and nondisj: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" and sub5j: "cball (a i) (r i) \ ball (a j) (5 * r j)" using cov [OF \i \ K\] by metis show "x \ (\i\D - D1. ball (a i) (5 * r i))" proof show "j \ D - D1" proof show "j \ D" using \j \ C\ sub5j \x \ cbox l u\ xi by (auto simp: D_def disjnt_def) obtain y where yi: "dist (a i) y \ r i" and yj: "dist (a j) y \ r j" using disjnt_def nondisj by fastforce have "dist x y \ r i + r i" by (metis add_mono dist_commute dist_triangle_le mem_cball xi yi) also have "\ < q" using ri by linarith finally have "y \ ball x q" by simp with yj q show "j \ D1" by (auto simp: disjoint_UN_iff) qed show "x \ ball (a j) (5 * r j)" using xi sub5j by blast qed qed have 3: "?\ (\i\D2. ball (a i) (5 * r i)) \ e" if D2: "D2 \ D - D1" and "finite D2" for D2 proof - have rgt0: "0 < r i" if "i \ D2" for i using \C \ K\ D_def \i \ D2\ D2 r01 by (simp add: subset_iff) then have inj: "inj_on (\i. ball (a i) (5 * r i)) D2" using \C \ K\ D2 by (fastforce simp: inj_on_def D_def ball_eq_ball_iff intro: ar_injective) have "?\ (\i\D2. ball (a i) (5 * r i)) \ sum (?\) ((\i. ball (a i) (5 * r i)) ` D2)" using that by (force intro: measure_Union_le) also have "\ = (\i\D2. ?\ (ball (a i) (5 * r i)))" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = (\i\D2. 5 ^ DIM('n) * ?\ (ball (a i) (r i)))" proof (rule sum.cong [OF refl]) - fix i - assume "i \ D2" - show "?\ (ball (a i) (5 * r i)) = 5 ^ DIM('n) * ?\ (ball (a i) (r i))" - using rgt0 [OF \i \ D2\] by (simp add: content_ball) + fix i assume "i \ D2" + thus "?\ (ball (a i) (5 * r i)) = 5 ^ DIM('n) * ?\ (ball (a i) (r i))" + using content_ball_conv_unit_ball[of "5 * r i" "a i"] + content_ball_conv_unit_ball[of "r i" "a i"] rgt0[of i] by auto qed also have "\ = (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" by (simp add: sum_distrib_left mult.commute) finally have "?\ (\i\D2. ball (a i) (5 * r i)) \ (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" . moreover have "(\i\D2. ?\ (ball (a i) (r i))) \ e / 5 ^ DIM('n)" proof - have D12_dis: "((\x\D1. cball (a x) (r x)) \ (\x\D2. cball (a x) (r x))) \ {}" proof clarify fix w d1 d2 assume "d1 \ D1" "w d1 d2 \ cball (a d1) (r d1)" "d2 \ D2" "w d1 d2 \ cball (a d2) (r d2)" then show "w d1 d2 \ {}" by (metis DiffE disjnt_iff subsetCE D2 \D1 \ D\ \D \ C\ pairwiseD [OF pwC, of d1 d2]) qed have inj: "inj_on (\i. cball (a i) (r i)) D2" using rgt0 D2 \D \ C\ by (force simp: inj_on_def cball_eq_cball_iff intro!: ar_injective) have ds: "disjoint ((\i. cball (a i) (r i)) ` D2)" using D2 \D \ C\ by (auto intro: pairwiseI pairwiseD [OF pwC]) have "(\i\D2. ?\ (ball (a i) (r i))) = (\i\D2. ?\ (cball (a i) (r i)))" - using rgt0 by (simp add: content_ball content_cball less_eq_real_def) + by (simp add: content_cball_conv_ball) also have "\ = sum ?\ ((\i. cball (a i) (r i)) ` D2)" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = ?\ (\i\D2. cball (a i) (r i))" by (auto intro: measure_Union' [symmetric] ds simp add: \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) = ?\ (\i\D1. cball (a i) (r i)) + ?\ (\i\D2. cball (a i) (r i))" by simp also have "\ = ?\ (\i \ D1 \ D2. cball (a i) (r i))" using D12_dis by (simp add: measure_Un3 \finite D1\ \finite D2\ fmeasurable.finite_UN) also have "\ \ ?\ (\i\D. cball (a i) (r i))" using D2 \D1 \ D\ by (fastforce intro!: measure_mono_fmeasurable [OF _ _ UD] \finite D1\ \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) \ ?\ (\i\D. cball (a i) (r i))" . with measD1 show ?thesis by simp qed ultimately show ?thesis by (simp add: field_split_simps) qed have co: "countable (D - D1)" by (simp add: \countable D\) show "(\i\D - D1. ball (a i) (5 * r i)) \ lmeasurable" using \e > 0\ by (auto simp: fmeasurable_UN_bound [OF co _ 3]) show "?\ (\i\D - D1. ball (a i) (5 * r i)) \ e" using \e > 0\ by (auto simp: measure_UN_bound [OF co _ 3]) qed qed qed qed (use C pwC in auto) qed define K' where "K' \ {i \ K. r i \ 1}" have 1: "\i. i \ K' \ 0 < r i \ r i \ 1" using K'_def r by auto have 2: "\i. i \ K' \ x \ cball (a i) (r i) \ r i < d" if "x \ S \ 0 < d" for x d using that by (auto simp: K'_def dest!: S [where d = "min d 1"]) have "K' \ K" using K'_def by auto then show thesis using * [OF 1 2] that by fastforce qed theorem Vitali_covering_theorem_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ ball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "negligible(S - (\i \ C. ball (a i) (r i)))" proof - have 1: "\i. i \ {i \ K. 0 < r i} \ x \ cball (a i) (r i) \ r i < d" if xd: "x \ S" "d > 0" for x d by (metis (mono_tags, lifting) assms ball_eq_empty less_eq_real_def mem_Collect_eq mem_ball mem_cball not_le xd(1) xd(2)) obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and neg: "negligible(S - (\i \ C. cball (a i) (r i)))" by (rule Vitali_covering_theorem_cballs [of "{i \ K. 0 < r i}" r S a, OF _ 1]) auto show thesis proof show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" apply (rule pairwise_mono [OF pw]) apply (auto simp: disjnt_def) by (meson disjoint_iff_not_equal less_imp_le mem_cball) have "negligible (\i\C. sphere (a i) (r i))" by (auto intro: negligible_sphere \countable C\) then have "negligible (S - (\i \ C. cball(a i)(r i)) \ (\i \ C. sphere (a i) (r i)))" by (rule negligible_Un [OF neg]) then show "negligible (S - (\i\C. ball (a i) (r i)))" by (rule negligible_subset) force qed (use C in auto) qed lemma negligible_eq_zero_density_alt: "negligible S \ (\x \ S. \e > 0. \d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d))" (is "_ = (\x \ S. \e > 0. ?Q x e)") proof (intro iffI ballI allI impI) fix x and e :: real assume "negligible S" and "x \ S" and "e > 0" then show "\d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)" apply (rule_tac x=e in exI) apply (rule_tac x="S \ ball x e" in exI) - apply (auto simp: negligible_imp_measurable negligible_Int negligible_imp_measure0 zero_less_measure_iff) + apply (auto simp: negligible_imp_measurable negligible_Int negligible_imp_measure0 zero_less_measure_iff + intro: mult_pos_pos content_ball_pos) done next assume R [rule_format]: "\x \ S. \e > 0. ?Q x e" let ?\ = "measure lebesgue" have "\U. openin (top_of_set S) U \ z \ U \ negligible U" if "z \ S" for z proof (intro exI conjI) show "openin (top_of_set S) (S \ ball z 1)" by (simp add: openin_open_Int) show "z \ S \ ball z 1" using \z \ S\ by auto show "negligible (S \ ball z 1)" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" let ?K = "{(x,d). x \ S \ 0 < d \ ball x d \ ball z 1 \ (\U. S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d))}" obtain C where "countable C" and Csub: "C \ ?K" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible((S \ ball z 1) - (\i \ C. ball (fst i) (snd i)))" proof (rule Vitali_covering_theorem_balls [of "S \ ball z 1" ?K fst snd]) fix x and d :: "real" assume x: "x \ S \ ball z 1" and "d > 0" obtain k where "k > 0" and k: "ball x k \ ball z 1" by (meson Int_iff open_ball openE x) let ?\ = "min (e / ?\ (ball z 1) / 2) (min (d / 2) k)" obtain r U where r: "r > 0" "r \ ?\" and U: "S \ ball x r \ U" "U \ lmeasurable" and mU: "?\ U < ?\ * ?\ (ball x r)" - using R [of x ?\] \d > 0\ \e > 0\ \k > 0\ x by auto + using R [of x ?\] \d > 0\ \e > 0\ \k > 0\ x by (auto simp: content_ball_pos) show "\i. i \ ?K \ x \ ball (fst i) (snd i) \ snd i < d" proof (rule exI [of _ "(x,r)"], simp, intro conjI exI) have "ball x r \ ball x k" using r by (simp add: ball_subset_ball_iff) also have "\ \ ball z 1" using ball_subset_ball_iff k by auto finally show "ball x r \ ball z 1" . have "?\ * ?\ (ball x r) \ e * content (ball x r) / content (ball z 1)" - using r \e > 0\ by (simp add: ord_class.min_def field_split_simps) + using r \e > 0\ by (simp add: ord_class.min_def field_split_simps content_ball_pos) with mU show "?\ U < e * content (ball x r) / content (ball z 1)" by auto qed (use r U x in auto) qed have "\U. case p of (x,d) \ S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d)" if "p \ C" for p using that Csub unfolding case_prod_unfold by blast then obtain U where U: "\p. p \ C \ case p of (x,d) \ S \ ball x d \ U p \ U p \ lmeasurable \ ?\ (U p) < e / ?\ (ball z 1) * ?\ (ball x d)" by (rule that [OF someI_ex]) let ?T = "((S \ ball z 1) - (\(x,d)\C. ball x d)) \ \(U ` C)" show "\T. S \ ball z 1 \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "S \ ball z 1 \ ?T" using U by fastforce { have Um: "U i \ lmeasurable" if "i \ C" for i using that U by blast have lee: "?\ (\i\I. U i) \ e" if "I \ C" "finite I" for I proof - have "?\ (\(x,d)\I. ball x d) \ ?\ (ball z 1)" apply (rule measure_mono_fmeasurable) using \I \ C\ \finite I\ Csub by (force simp: prod.case_eq_if sets.finite_UN)+ then have le1: "(?\ (\(x,d)\I. ball x d) / ?\ (ball z 1)) \ 1" - by simp + by (simp add: content_ball_pos) have "?\ (\i\I. U i) \ (\i\I. ?\ (U i))" using that U by (blast intro: measure_UNION_le) also have "\ \ (\(x,r)\I. e / ?\ (ball z 1) * ?\ (ball x r))" by (rule sum_mono) (use \I \ C\ U in force) also have "\ = (e / ?\ (ball z 1)) * (\(x,r)\I. ?\ (ball x r))" by (simp add: case_prod_app prod.case_distrib sum_distrib_left) also have "\ = e * (?\ (\(x,r)\I. ball x r) / ?\ (ball z 1))" apply (subst measure_UNION') using that pwC by (auto simp: case_prod_unfold elim: pairwise_mono) also have "\ \ e" by (metis mult.commute mult.left_neutral real_mult_le_cancel_iff1 \e > 0\ le1) finally show ?thesis . qed have "\(U ` C) \ lmeasurable" "?\ (\(U ` C)) \ e" using \e > 0\ Um lee by(auto intro!: fmeasurable_UN_bound [OF \countable C\] measure_UN_bound [OF \countable C\]) } moreover have "?\ ?T = ?\ (\(U ` C))" proof (rule measure_negligible_symdiff [OF \\(U ` C) \ lmeasurable\]) show "negligible((\(U ` C) - ?T) \ (?T - \(U ` C)))" by (force intro!: negligible_subset [OF negC]) qed ultimately show "?T \ lmeasurable" "?\ ?T \ e" by (simp_all add: fmeasurable.Un negC negligible_imp_measurable split_def) qed qed qed with locally_negligible_alt show "negligible S" by metis qed proposition negligible_eq_zero_density: "negligible S \ (\x\S. \r>0. \e>0. \d. 0 < d \ d \ r \ (\U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)))" proof - let ?Q = "\x d e. \U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d)" have "(\e>0. \d>0. d \ e \ ?Q x d e) = (\r>0. \e>0. \d>0. d \ r \ ?Q x d e)" if "x \ S" for x proof (intro iffI allI impI) fix r :: "real" and e :: "real" assume L [rule_format]: "\e>0. \d>0. d \ e \ ?Q x d e" and "r > 0" "e > 0" show "\d>0. d \ r \ ?Q x d e" using L [of "min r e"] apply (rule ex_forward) - using \r > 0\ \e > 0\ by (auto intro: less_le_trans elim!: ex_forward) + using \r > 0\ \e > 0\ by (auto intro: less_le_trans elim!: ex_forward simp: content_ball_pos) qed auto then show ?thesis by (force simp: negligible_eq_zero_density_alt) qed end