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,3498 +1,3498 @@ (* 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 \0 < (MAX k. \m k\) * B\ \0 < B\ zero_less_mult_pos2 by fastforce 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, opaque_lifting) 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 $ Transposition.transpose m n 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 $ Transposition.transpose m n 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 transpose_involutory)+ 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 Transposition.transpose m n i = j then 1 else 0) = (\ i j. if j = Transposition.transpose m n i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Transposition.transpose m n 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]) 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_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: 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, opaque_lifting) 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_vol + e * ?unit_vol) * r ^ CARD('n)" proof - 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_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: 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_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, opaque_lifting) less_le_trans mult.commute not_le 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 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_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_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, opaque_lifting) 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 intro: content_cball_pos) also have "\ = e * content (ball x (min d r))" 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, opaque_lifting) 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: 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 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, opaque_lifting) 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 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) 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, opaque_lifting) 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) 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_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_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_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 integral_indicator) apply (simp add: indicator_def of_bool_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 \x \ S\ h_le_f indicator_simps(1) indicator_simps(2) mem_Collect_eq mult.commute mult.left_neutral mult_zero_left 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_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_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_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_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_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_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_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_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_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_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 of_bool_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 of_bool_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_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_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_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_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_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 text \when @{term "n=1"}\ lemma has_absolute_integral_change_of_variables_1': fixes f :: "real \ real" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_field_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 - have "(\x. \g' x\ *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) \ (\x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \ integral (g ` S) (\x. vec (f x)) = (vec b :: real ^ 1)" - using assms unfolding has_field_derivative_iff_has_vector_derivative + using assms unfolding has_real_derivative_iff_has_vector_derivative by (intro has_absolute_integral_change_of_variables_1 assms) auto thus ?thesis by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) qed 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/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,5084 +1,5084 @@ (* 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 LIMSEQ_if_less: "(\k. if i < k then a else b) \ a" by (rule_tac k="Suc i" in LIMSEQ_offset) auto text \Note that the rhs is an implication. This lemma plays a specific role in one proof.\ 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 by (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") (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 using has_integral_const[of "1::real" l u] by (simp flip: has_integral_restrict[OF box_subset_cbox] add: has_integral_spike_interior) 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 "(\k. ?f k x) \ (if x \ \(F ` UNIV) \ box a b then 1 else 0)" for x by (auto simp: * sum.If_cases Iio_Int_singleton if_distrib LIMSEQ_if_less cong: if_cong) 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 of_bool_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 of_bool_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] have U_le_f: "U i x \ f x" for i x by (metis (no_types) LIMSEQ_le_const UNIV_I incseq_def le_fun_def seq.hyps(4) seq.hyps(5) space_borel) { 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 of_bool_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 of_bool_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" by (simp add: indicator_scaleR_eq_if) 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 of_bool_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" unfolding absolutely_integrable_on_def by (metis (mono_tags, lifting) integrable_eq integrable_reflect) 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] of_bool_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 (simp add: indicator_def [abs_def] of_bool_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) moreover have "integral T (\a. if a \ S then 1::real else 0) = integral (S \ T \ UNIV) (\a. 1)" by (simp add: Henstock_Kurzweil_Integration.integral_restrict_Int) moreover have "integral T (indicat_real S) = integral T (\a. if a \ S then 1 else 0)" by (simp add: indicator_def [abs_def] of_bool_def) ultimately show ?thesis by (simp add: assms lmeasure_integral) 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 then show ?lhs apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) 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 by (metis euclidean_all_zero_iff inner_zero_right) moreover have "c = 1" if "a \ (x + c *\<^sub>R w) = b" "a \ (x + w) = b" for c w using that by (metis (no_types, opaque_lifting) add_diff_eq diff_0 diff_minus_eq_add inner_diff_right inner_scaleR_right mult_cancel_right2 right_minus_eq x) ultimately show ?thesis using starlike_negligible [OF closed_hyperplane, of x] by simp 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 \ interior (convex hull insert 0 B)" proof (rule interior_simplex_nonempty [OF indB]) show "finite B" by (simp add: indB independent_finite) show "card B = DIM('N)" by (simp add: cardB 2) qed then have a: "a \ interior S" by (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) 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" using eq mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"] unfolding frontier_def by (metis Diff_iff add_diff_cancel_left' add_diff_eq diff_gt_0_iff_gt group_cancel.rule0 not_le) 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 = {})" by (metis Diff_empty closure_subset frontier_def interior_subset negligible_convex_frontier negligible_subset open_interior open_not_negligible) 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)" by (metis (full_types) completion.null_sets_outer negligible_iff_null_sets negligible_imp_measure0 order_refl) 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") by (metis lmeasurable_iff_indicator_has_integral negligible_iff_measure) 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 (auto 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 by (rule_tac x="min (1/2) \" in exI) (simp add: U dist_norm norm_minus_commute subset_iff) 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 proof (rule_tac c = "abs B + 1" in that) show "S \ cbox (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One)" using norm_bound_Basis_le Basis_le_norm by (fastforce simp: mem_box dest!: B intro: order_trans) show "box (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One) \ {}" by (simp add: box_eq_empty(1)) qed 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 by (metis Int_iff interior_cbox cbox Ksub) 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)" by (simp add: \X \ \\ inner_diff_left prj1_idem cong: prod.cong) 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 \ \'\ \0 < B\ flip: prj1_eq) using MleN 0 1 uvz \X \ \\ by (fastforce simp add: box_ne_empty power_decreasing) 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 "\K. K \ \' \ \a b. K = cbox a b" using cbox that by blast then have div: "\' division_of \\'" using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def by (force simp: \finite \'\ \{} \ \'\ division_of_def) 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" using \0 < B\ by (intro mult_left_mono [OF le_meaT]) (force simp add: algebra_simps) 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 "norm x \ real (nat \max B (norm x)\)" by linarith then have "x \ ?S (nat (ceiling (max B (norm x))))" using T no by (force simp: \x \ S\ algebra_simps) 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:) show ?thesis proof (intro exI conjI ballI) show "norm (f y - f x) \ (B + 1) * norm (y - x)" if "y \ S \ ball 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 (metis IntE d dist_norm mem_ball norm_minus_commute that) finally show ?thesis by (simp add: algebra_simps) qed qed (use \d>0\ in auto) 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. lift (x, 0)) ` S)" proof - have "negligible {x. x\j = 0}" by (metis \j \ Basis\ negligible_standard_hyperplane) moreover have "(\m. lift (m, 0)) ` S \ {n. n \ j = 0}" using j by force ultimately show ?thesis using negligible_subset by auto qed moreover have "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 linear_fst]) apply (force simp: image_comp o_def) done moreover have "f = f \ fst \ drop \ (\x. lift (x, 0))" by (simp add: o_def) ultimately show ?thesis by (metis (no_types) image_comp negligible_differentiable_image_negligible order_refl) 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 "indicat_real (\ (range S)) integrable_on UNIV \ (\n. integral UNIV (indicat_real (S n))) \ integral UNIV (indicat_real (\ (range S)))" proof (rule monotone_convergence_increasing) show "\n. (indicat_real (S n)) integrable_on UNIV" using S measurable_integrable by blast show "\n x::'a. indicat_real (S n) x \ (indicat_real (S (Suc n)) x)" by (simp add: indicator_leI nest rev_subsetD) have "\x. (\n. x \ S n) \ (\\<^sub>F n in sequentially. x \ S n)" by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) then show "\x. (\n. indicat_real (S n) x) \ (indicat_real (\(S ` UNIV)) x)" by (simp add: indicator_def tendsto_eventually) show "bounded (range (\n. integral UNIV (indicat_real (S n))))" using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) qed then have "(\n. S n) \ lmeasurable \ ?Lim" by (simp add: lmeasure_integral_UNIV S cong: conj_cong) (simp add: measurable_integrable) 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)" proof (rule measure_mono_fmeasurable) show "\ (S ` {..n}) \ sets lebesgue" using S by blast qed (use ab in auto) 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 by (metis locally_negligible_alt negligible_subset openin_imp_subset openin_subtopology_self) 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" shows "\sum (\x. norm(f x)) S - sum (\x. norm(g x)) S\ < e" (is "?lhs < e") proof - have "?lhs \ (\i\S. \norm (f i) - norm (g i)\)" by (metis (no_types) sum_abs sum_subtractf) also have "... \ (\x\S. norm (f x - g x))" by (simp add: norm_triangle_ineq3 sum_mono) also have "... < e" using assms(1) by blast finally show ?thesis . qed 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 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'}" 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 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 in_p': "(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 Ex_p_p': "\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))" proof (rule sum.over_tagged_division_lemma[OF p'']) show "\u v. box u v = {} \ norm (integral (cbox u v) f) = 0" by (auto intro: integral_null simp: content_eq_0_interior) qed 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'_div: "d' division_of cbox u v" unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) moreover have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" by (simp add: sum_norm_le) moreover have "f integrable_on K" using f integrable_on_subcbox uv uvab by blast moreover have "d' division_of K" using d'_div uv by blast ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" by (simp add: integral_combine_division_topdown) also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof (rule sum.mono_neutral_left) show "finite {K \ L |L. L \ snd ` p}" by (simp add: snd_p(1)) show "\i\{K \ L |L. L \ snd ` p} - d'. norm (integral i f) = 0" "d' \ {K \ L |L. L \ snd ` p}" using d'_def image_eqI uv by auto qed also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" unfolding Setcompr_eq_image proof (rule sum.reindex_nontrivial [unfolded o_def]) show "finite (snd ` p)" using snd_p(1) by blast show "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 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))" unfolding sum_p' proof (rule sum.mono_neutral_right) show "finite {i \ l |i l. i \ d \ l \ snd ` p}" by (metis * finite_imageI[OF fin_d_sndp]) show "snd ` p' \ {i \ l |i l. i \ d \ l \ snd ` p}" by (clarsimp simp: p'_def) (metis image_eqI snd_conv) show "\i\{i \ l |i l. i \ d \ l \ snd ` p} - snd ` p'. norm (integral i f) = 0" by clarsimp (metis Henstock_Kurzweil_Integration.integral_empty disjoint_iff image_eqI in_p' snd_conv) 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 proof (rule sum.mono_neutral_left) show "finite {(x, i \ l) |x i l. (x, l) \ p \ i \ d}" by (simp add: "*" fin_pd) qed (use p'alt in \force+\) 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 = {}" using that p'(5) d'(5) by (metis snd_conv) 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 by force+ 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 by (fastforce intro: sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" proof (rule sum.mono_neutral_right) show "finite {k \ cbox u v |k. k \ d}" by (simp add: d'(1)) qed (fastforce simp: inf.commute)+ finally have "(\i\d. \content (l \ i)\) = content (cbox u v)" using additive_content_division[OF division_inter_1[OF d(1)]] uv xl(2) by auto then show "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" unfolding sum_distrib_right[symmetric] using uv by auto 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)}" define SDF where "SDF \ SUP d\?D. ?f d" have D_1: "?D \ {}" by (rule elementary_interval) auto have D_2: "bdd_above (?f`?D)" using assms(2) by auto have f_int: "\a b. f absolutely_integrable_on cbox a b" using assms integrable_on_subcbox by (blast intro!: bounded_variation_absolutely_integrable_interval) have "\B>0. \a b. ball 0 B \ cbox a b \ \integral (cbox a b) (\x. norm (f x)) - SDF\ < e" if "0 < e" for e proof - have "\y \ ?f ` ?D. \ y \ SDF - e" proof (rule ccontr) assume "\ ?thesis" then have "SDF \ SDF - e" unfolding SDF_def by (metis (mono_tags) D_1 cSUP_least image_eqI) then show False using that by auto qed then obtain d K where ddiv: "d division_of \d" and "K = ?f d" "SDF - e < K" by (auto simp add: image_iff not_le) then have d: "SDF - e < ?f d" by auto note d'=division_ofD[OF ddiv] have "bounded (\d)" using ddiv by blast then obtain K where K: "0 < K" "\x\\d. norm x \ K" using bounded_pos by blast show ?thesis proof (intro conjI impI allI exI) fix a b :: 'n assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. \SDF - e < s1; s1 \ s; s < SDF + e\ \ \s - SDF\ < e" by arith show "\integral (cbox a b) (\x. norm (f x)) - SDF\ < e" unfolding real_norm_def proof (rule * [OF d]) have "?f d \ 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))" by (metis (full_types) absolutely_integrable_on_def d'(4) ddiv f_int integral_combine_division_bottomup) also have "\ \ integral (cbox a b) (\x. norm (f x))" 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 "?f d \ integral (cbox a b) (\x. norm (f x))" . 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 \ SDF; \sf - si\ < e/2; \sf' - di\ < e/2\ \ di < SDF + e" by arith have "integral (cbox a b) (\x. norm (f x)) < SDF + 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 proof (rule absdiff_norm_less) show "(\p\p. norm (content (snd p) *\<^sub>R f (fst p) - integral (snd p) f)) < e/2" using d2[of p] p(1,3) by (auto simp: tagged_division_of_def split_def) qed 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]) have "(\(x,k) \ p. norm (integral k f)) = (\k\snd ` p. norm (integral k f))" apply (rule sum.over_tagged_division_lemma[OF p(1)]) by (metis Henstock_Kurzweil_Integration.integral_empty integral_open_interval norm_zero) also have "... \ SDF" using partial_division_of_tagged_division[of p "cbox a b"] p(1) by (auto simp: SDF_def tagged_partial_division_of_def intro!: cSUP_upper2 D_1 D_2) finally show "(\(x,k) \ p. norm (integral k f)) \ SDF" . qed then show "integral (cbox a b) (\x. norm (f x)) < SDF + e" by simp qed qed (use K in auto) qed moreover have "\a b. (\x. norm (f x)) integrable_on cbox a b" using absolutely_integrable_on_def f_int by auto ultimately have "((\x. norm (f x)) has_integral SDF) UNIV" by (auto simp: has_integral_alt') 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)" proof (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) show "\K. K \ \ \ K \ S \ lmeasurable" using \div \S \ lmeasurable\ by blast show "pairwise (\K y. negligible (K \ S \ (y \ S))) \" unfolding pairwise_def by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) qed 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)" by (metis Diff_partition Int_commute calculation closure_Un_frontier frontier_def inf.absorb_iff2 inf_bot_right inf_sup_absorb negligible_Un_eq open_interior open_not_negligible sup_commute) 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) 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 ` (\\))" proof - have "?\ (cbox a b - S) = ?\ (cbox a b) - ?\ S" by (simp add: measurable_measure_Diff \S \ cbox a b\ fmeasurableD that(2)) then have "(?\ S - e / (1 + m)) \ (content (cbox a b) - ?\ (\ \))" using \m \ 0\ le by auto then show ?thesis using \m \ 0\ \e > 0\ by (simp add: mult_left_mono im fUD [OF \countable \\ cbox intdisj] flip: right_diff_distrib) qed also have "\ = ?\ (f ` cbox a b - f ` \\)" proof (rule measurable_measure_Diff [symmetric]) show "f ` cbox a b \ lmeasurable" by (simp add: assms(1) measurable_linear_image_interval) show "f ` \ \ \ sets lebesgue" by (simp add: \countable \\ cbox fUD fmeasurableD intdisj) show "f ` \ \ \ f ` cbox a b" by (simp add: Sup_le_iff cbox image_mono) qed 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: "?\ ((f ` S) \ cbox a b) = m * ?\ (S \ h ` 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\ \ \?\ S - ?\ (S \ h ` cbox a b)\ * m" by (metis "2" \m \ 0\ abs_minus_commute abs_mult_pos mult.commute order_refl right_diff_distrib') also have "\ \ e / (1 + m) * m" by (metis \m \ 0\ abs_minus_commute abs_of_nonneg meas_adiff mult.commute mult_left_mono) also have "\ < e" using \e > 0\ \m \ 0\ by (simp 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 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)" (is "?lhs = ?rhs") if "f integrable_on A" proof assume ?lhs then show ?rhs by (metis absolutely_integrable_on_def Topology_Euclidean_Space.norm_nth_le absolutely_integrable_integrable_bound integrable_component that) next assume R: ?rhs have "f absolutely_integrable_on A" proof (rule absolutely_integrable_integrable_bound) show "(\x. \i\Basis. norm (f x \ i)) integrable_on A" using R by (force intro: integrable_sum) qed (use that norm_le_l1 in auto) then show ?lhs using absolutely_integrable_on_def by auto 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" by (simp add: absolutely_integrable_linear assms 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 *: "(\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" using assms \i \ Basis\ unfolding o_def by (intro absolutely_integrable_norm [unfolded o_def]) (auto simp: algebra_simps dest: absolutely_integrable_component) ultimately show ?thesis by (subst comp_assoc) (blast intro: absolutely_integrable_linear) qed 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) show ?thesis unfolding eq by (rule absolutely_integrable_sum) (force simp: intro!: *)+ 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" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) 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" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) 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" proof (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (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" proof (rule set_integral_add [OF f nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (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) subgoal by (metis vector_one_nth) subgoal apply (erule all_forward imp_forward ex_forward asm_rl)+ by (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}" proof (clarsimp) fix t assume "\t\ < B" then show "a $ 1 \ t \ t \ b $ 1" using subsetD [OF B] by (metis (mono_tags, opaque_lifting) mem_ball_0 mem_box_cart(2) norm_real vec_component) qed 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) subgoal by (metis vector_one_nth) apply (erule thin_rl) apply (erule all_forward ex_forward conj_forward)+ by (blast intro!: *)+ 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] + (auto simp: has_real_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 of_bool_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 of_bool_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] + unfolding has_real_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" by (auto intro: filterlim_compose[OF lim filterlim_tendsto_add_at_top] filterlim_real_sequentially) 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 "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = LBINT x. F x * g x * indicat_real {a..b} x + f x * G x * indicat_real {a..b} x" by (meson vector_space_over_itself.scale_left_distrib) also have "... = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" proof (intro Bochner_Integration.integral_add borel_integrable_atLeastAtMost cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed finally 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" . moreover have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" proof (intro integral_FTC_Icc_real derivative_eq_intros cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed auto ultimately show ?thesis 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 subsection \A non-negative continuous function whose integral is zero must be zero\ 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 "continuous_on (cbox a b) f" and "\x. x \ box a b \ 0 \ f x" assumes "(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 corollary integral_cbox_eq_0_iff: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (cbox a b) f" and "box a b \ {}" and "\x. x \ cbox a b \ f x \ 0" shows "integral (cbox a b) f = 0 \ (\x \ cbox a b. f x = 0)" (is "?lhs = ?rhs") proof assume int0: ?lhs show ?rhs using has_integral_0_cbox_imp_0 [of a b f] assms by (metis box_subset_cbox eq_integralD int0 integrable_continuous subsetD) next assume ?rhs then show ?lhs by (meson has_integral_is_0_cbox integral_unique) qed lemma integral_eq_0_iff: fixes f :: "real \ real" assumes "continuous_on {a..b} f" and "a < b" and "\x. x \ {a..b} \ f x \ 0" shows "integral {a..b} f = 0 \ (\x \ {a..b}. f x = 0)" using integral_cbox_eq_0_iff [of a b f] assms by simp lemma integralL_eq_0_iff: fixes f :: "real \ real" assumes contf: "continuous_on {a..b} f" and "a < b" and "\x. x \ {a..b} \ f x \ 0" shows "integral\<^sup>L (lebesgue_on {a..b}) f = 0 \ (\x \ {a..b}. f x = 0)" using integral_eq_0_iff [OF assms] by (simp add: contf continuous_imp_integrable_real lebesgue_integral_eq_integral) text \In fact, strict inequality is required only at a single point within the box.\ lemma integral_less: fixes f :: "'n::euclidean_space \ real" assumes cont: "continuous_on (cbox a b) f" "continuous_on (cbox a b) g" and "box a b \ {}" and fg: "\x. x \ box a b \ f x < g x" shows "integral (cbox a b) f < integral (cbox a b) g" proof - obtain int: "f integrable_on (cbox a b)" "g integrable_on (cbox a b)" using cont integrable_continuous by blast then have "integral (cbox a b) f \ integral (cbox a b) g" by (metis fg integrable_on_open_interval integral_le integral_open_interval less_eq_real_def) moreover have "integral (cbox a b) f \ integral (cbox a b) g" proof (rule ccontr) assume "\ integral (cbox a b) f \ integral (cbox a b) g" then have 0: "((\x. g x - f x) has_integral 0) (cbox a b)" by (metis (full_types) cancel_comm_monoid_add_class.diff_cancel has_integral_diff int integrable_integral) have cgf: "continuous_on (cbox a b) (\x. g x - f x)" using cont continuous_on_diff by blast show False using has_integral_0_cbox_imp_0 [OF cgf _ 0] assms(3) box_subset_cbox fg less_eq_real_def by fastforce qed ultimately show ?thesis by linarith qed lemma integral_less_real: fixes f :: "real \ real" assumes "continuous_on {a..b} f" "continuous_on {a..b} g" and "{a<.. {}" and "\x. x \ {a<.. f x < g x" shows "integral {a..b} f < integral {a..b} g" by (metis assms box_real integral_less) 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, opaque_lifting) 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 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" by (metis abs_absolutely_integrableI_1 absolutely_integrable_measurable_real assms measurable_bounded_by_integrable_imp_integrable order_refl real_norm_def set_integrable_abs set_lebesgue_integral_eq_integral(1)) 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 using that noxy by (rule_tac x = "ball (f x) 1" in exI) (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 (y - x) / real (max i j) \ norm (y - x) / real i" using \i \ 0\ \j \ 0\ by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) also have "... \ norm (f y - f x)" using B [of "y-x"] le norm_triangle_ineq3 [of "f y - f x" "f' x (y - x)"] by linarith finally show ?thesis . 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 with \k > 0\ \B > 0\ \d > 0\ \a \ S\ 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)" by (rule_tac x="min k (d / B)" in exI) (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,7606 +1,7645 @@ (* 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" proof (intro exI conjI) show "(?S(h \ f) has_integral h z) (cbox a b)" by (simp add: * has_integral_linear_cbox[OF z(1) h]) show "norm (h z - h y) < e" by (metis B diff le_less_trans pos_less_divide_eq z(2)) qed 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" proof (intro exI conjI) show "(?S(\x. f x + g x) has_integral (w + z)) (cbox a b)" by (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) show "norm (w + z - (k + l)) < e" by (metis dist_norm dist_triangle_add_half w(2) z(2)) qed 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)" by (meson has_integral_iff has_integral_linear) 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 integrable_cong: + assumes "\x. x \ A \ f x = g x" + shows "f integrable_on A \ g integrable_on A" + using has_integral_cong [OF assms] by fast + 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, opaque_lifting) 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 "gauge (\x. \{\ i x |i. i \ {0..n}})" for n using \ by (intro gauge_Inter) 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 by (auto simp: integrable_split intro: integral_unique [OF has_integral_split[OF _ _ k]]) 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 (auto intro: has_integral_split[OF _ _ k]) 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: "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" by simp 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)" by (metis additive_content_division p eq_iff sumeq) 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 (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)" proof - have "\xk. xk \ p \ norm (f (fst xk)) \ e" using assms(2) p tag_in_interval by force moreover have "(\i\p. \content (snd i)\ * e) \ e * content (cbox a b)" unfolding sum_distrib_right[symmetric] using con sum_le by (auto simp: mult.commute intro: mult_left_mono [OF _ e]) ultimately show ?thesis unfolding split_def norm_scaleR by (metis (no_types, lifting) mult_left_mono[OF _ abs_ge_zero] order_trans[OF sum_mono]) qed 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)" using order_trans[OF _ rsum_bound[OF assms]] by (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) 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" using has_integral_component_le assms by blast 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 using True assms has_integral_component_nonneg by blast 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" using assms has_integral_component_lbound by blast 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)" using assms has_integral_component_ubound by blast 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)" using \M \ 0\ m n by (intro add_mono; force 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 * by (simp add: sum_nonneg split: prod.split) 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))" using q(1) by (intro sum_Sigma_product [symmetric]) auto 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 have \
: "\a b z. \\x. x \ T \ x \ S \ g x = f x; ((\x. if x \ T then f x else 0) has_integral z) (cbox a b)\ \ ((\x. if x \ T then g x else 0) has_integral z) (cbox a b)" by (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (auto split: if_split_asm) apply (blast dest: *) using \
by meson 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 using that by (simp add: indicator_def) 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" by (metis insert_is_Un negligible_Un_eq negligible_sing) 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)" by (meson Diff_iff gf has_integral_spike[OF negligible_frontier_interval _ f]) 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 using assms that by (metis content_eq_0_interior integrable_on_null interior_cbox norm_zero right_minus_eq) { 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}" using fg 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 .. 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" proof (intro exI conjI strip) 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 proof (subst has_integral, safe) fix e :: real assume e: "e > 0" show "?P (e * content (cbox a b)) (\)" if \
[rule_format]: "\\>0. ?P \ (<)" using \
[of "e * content (cbox a b)"] by (meson F e less_imp_le mult_pos_pos) show "?P e (<)" if \
[rule_format]: "\\>0. ?P (\ * content (cbox a b)) (\)" using \
[of "e/2 / content (cbox a b)"] using F e by (force simp add: algebra_simps) 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)" proof (rule d) show "norm (u - x) < d x" using \u \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \u \ K\ kab in auto) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" proof (rule d) show "norm (v - x) < d x" using \v \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \v \ K\ kab in auto) 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 has_complex_derivative_imp_has_vector_derivative: + fixes f :: "complex \ complex" + assumes "(f has_field_derivative f') (at (of_real a) within (cbox (of_real x) (of_real y)))" + shows "((f o of_real) has_vector_derivative f') (at a within {x..y})" + using has_derivative_in_compose[of of_real of_real a "{x..y}" f "(*) f'"] assms + by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def o_def cbox_complex_of_real) + 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}" unfolding power2_eq_square by (rule fundamental_theorem_of_calculus [OF assms] 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] + unfolding has_real_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] + unfolding has_real_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_exp [simp]: + fixes a::real + assumes "a \ b" shows "integral {a..b} exp = exp b - exp a" + by (meson DERIV_exp assms fundamental_theorem_of_calculus has_real_derivative_iff_has_vector_derivative + has_vector_derivative_at_within integral_unique) + + 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 unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ 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 unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ 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 proof qed show ?thesis proof show "\a b. box a b = {} \ (f integrable_on cbox a b) = True" by (simp add: content_eq_0_interior integrable_on_null) show "\a b c k. k \ Basis \ (f integrable_on cbox a b) \ (f integrable_on cbox a b \ {x. x \ k \ c} \ f integrable_on cbox a b \ {x. c \ x \ k})" unfolding integrable_on_def by (auto intro!: has_integral_split) qed 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\" (is "?lhs \ ?rhs") 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}" proof (rule has_integral_diff) show "(f has_integral integral {x..y} f) {x..y}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (y - x) *\<^sub>R f x) {x..y}" using has_integral_const_real [of "f x" x y] False by simp qed have "?lhs \ e * content {x..y}" using yx False d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ ?rhs" using False by auto finally show ?thesis . 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}" proof (rule has_integral_diff) show "(f has_integral integral {y..x} f) {y..x}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" using has_integral_const_real [of "f x" y x] True by simp qed have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * content {y..x}" using yx True d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ e * \y - x\" using True by auto finally have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" . 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) + by (auto simp: has_real_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}" by clarsimp (metis (no_types, lifting) gh 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)\(\(y,L). (g y, g ` L)) ` p. content K *\<^sub>R f x) = (\u\p. case case u of (x,K) \ (g x, g ` K) of (y,L) \ content L *\<^sub>R f y)" by (metis (mono_tags, lifting) "*" sum.reindex_cong) also have "... = (\(x,K)\p. r *\<^sub>R content K *\<^sub>R f (g x))" using r by (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) finally 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" by (simp add: scaleR_right.sum split_def) also have "\ = r *\<^sub>R ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" using \0 < r\ by (auto simp: scaleR_diff_right) finally 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 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) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) 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) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) 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)" proof (rule has_integral_twiddle) show "\w z. (\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox u v = cbox w z" "\w z. (\x. m *\<^sub>R x + c) ` cbox u v = cbox w z" for u v using interval_image_affinity_interval by blast+ show "content ((\x. m *\<^sub>R x + c) ` cbox u v) = \m\ ^ DIM('a) * content (cbox u v)" for u v using content_image_affinity_cbox by blast qed (use assms zero_less_power in \auto simp: field_simps\) 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 has_integral_affinity assms unfolding integrable_on_def by blast 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_cmul_iff': + assumes "c \ 0" + shows "((\x. c *\<^sub>R f x) has_integral I) A \ (f has_integral I /\<^sub>R c) A" + using assms by (metis divideR_right has_integral_cmul_iff) + 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: "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" (is "?lhs = ?rhs") proof - have "?lhs = (\ k. f k (x \ axis k 1))" by (simp add: cart_eq_inner_axis) also have "... = (\u\UNIV. f u (x \ axis u 1) *\<^sub>R axis u 1)" by (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) also have "... = ?rhs" by (simp add: Basis_vec_def UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) finally show ?thesis . qed 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)" proof - have "x \ uminus ` cbox a b" if "x \ cbox (- b) (- a)" for x proof - have "-x \ cbox a b" using that by (auto simp: mem_box) then show ?thesis by force qed then show ?thesis by (auto simp: mem_box) qed 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, THEN conjunct2, rule_format] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" by (simp add: bounded_linear_scaleR_left) 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") \\The explicit quantifier is required by the following step\ proof fix x assume "x \ box a b" with e show "?Q x" using derf_exp [of x "e/2"] by auto qed then 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)" unfolding bgauge_existence_lemma 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 show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' a)) *\<^sub>R f' a) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' a)" using False ab e by (auto simp add: field_simps) qed 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 show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' b)) *\<^sub>R f' b) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' b)" using False ab e by (auto simp add: field_simps) qed 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))))" using p(1) ab e by (subst sum.union_disjoint) 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] by (subst additive_tagged_division_1[OF \a \ b\ ptag]) 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 by (metis (no_types, lifting) Diff_Diff_Int Diff_subset norm_le p(1) sum.subset_diff) 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}" using \a < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) 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}" using \t < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) 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}" proof (intro tagged_division_Un_interval_real) show "{(c, {t..c})} tagged_division_of {a..c} \ {x. t \ x \ 1}" using \t \ c\ by (auto simp: eqs tagged_division_of_self_real) qed (auto simp: eqs ptag) 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" using assms t by (auto simp: algebra_simps 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 then show ?thesis by (force simp: dist_norm algebra_simps intro: indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\]) next case 2 then show ?thesis by (force simp: dist_norm norm_minus_commute algebra_simps intro: indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\]) 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) + by (auto simp: has_real_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})" using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps intro: has_derivative_subset [OF derf]) 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_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" by (meson \0 < B\ \0 < e\ add_pos_pos le_less_trans zero_less_one norm_pths(2)) 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) show ?thesis using as by (force intro!: has_integral_component_le[OF k \
]) 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: ?r show ?l unfolding negligible_def proof safe fix a b have "negligible (s \ cbox a b)" by (simp add: R) then show "(indicator s has_integral 0) (cbox a b)" by (meson Diff_iff Int_iff has_integral_negligible indicator_simps(2)) 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 negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" using negligible_insert by fastforce 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) +lemma has_integral_Icc_iff_Ioo: + fixes f :: "real \ 'a :: banach" + shows "(f has_integral I) {a..b} \ (f has_integral I) {a<.. {a..b} - {a<.. 0}" + by (rule negligible_subset [of "{a,b}"]) auto + show "negligible {x \ {a<.. 0}" + by (rule negligible_subset [of "{}"]) auto +qed + +lemma integrable_on_Icc_iff_Ioo: + fixes f :: "real \ 'a :: banach" + shows "f integrable_on {a..b} \ f integrable_on {a<..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" by (meson assms has_integral_subset_component_le integrable_integral) 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" using assms has_integral_subset_le by blast 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)" by (simp add: has_integral_iff rhs) 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" using fgh by (force intro: has_integral_le) 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" using ** Bg ballBg Bh ballBh by blast 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_UN: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite I" and int: "\i. i \ I \ (f has_integral (g i)) (\ i)" and neg: "pairwise (\i i'. negligible (\ i \ \ i')) I" shows "(f has_integral (sum g I)) (\i\I. \ i)" proof - let ?\ = "((\(a,b). \ a \ \ b) ` {(a,b). a \ I \ b \ I-{a}})" have "((\x. if x \ (\i\I. \ i) then f x else 0) has_integral sum g I) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (I \ I)" by (simp add: \finite I\) moreover have "{(a,b). a \ I \ b \ I-{a}} \ I \ I" by auto ultimately show "finite ?\" by (simp add: finite_subset) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ (\i\I. \ i) then f x else 0) = (\i\I. if x \ \ i then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix i assume i: "i \ I" "x \ \ i" then have "\j\I. x \ \ j \ j = i" using that by blast with i show "f x = (\i\I. if x \ \ i then f x else 0)" by (simp add: sum.delta[OF \finite I\]) qed next show "((\x. (\i\I. if x \ \ i then f x else 0)) has_integral sum g I) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum [OF \finite I\]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite \" and "\S. S \ \ \ (f has_integral (i S)) S" and "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - have "(f has_integral (sum i \)) (\S\\. S)" by (intro has_integral_UN assms) then show ?thesis by force 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) \" by (meson assms integral_unique 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) \" using assms has_integral_combine_division_topdown by blast 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. (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" proof - have "snd ` p division_of S" by (simp add: assms(1) division_of_tagged_division) with assms show ?thesis by (metis (mono_tags, lifting) has_integral_combine_division has_integral_integrable_integral imageE prod.collapse) qed 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: "p tagged_division_of (cbox a b)" and f: "\x k. (x,k)\p \ f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" by (simp add: has_integral_combine_tagged_division[OF p] integral_unique f integrable_integral) 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 (simp add: has_integral_combine_tagged_division p) 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" using assms by (auto intro: integral_unique [OF has_integral_combine_tagged_division_topdown]) 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, opaque_lifting) 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 "interior T \ interior (\(snd ` p)) = {}" if "T \ r" for T proof (rule Int_interior_Union_intervals) show "\U. U \ snd ` p \ \a b. U = cbox a b" using q q'(4) by blast show "\U. U \ snd ` p \ interior T \ interior U = {}" by (metis DiffE q q'(5) r_def subsetD that) qed (use p' in auto) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" by (metis Int_commute) 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" proof (subst (asm) sum.reindex_nontrivial [OF \finite r\]) qed (auto simp: split_paired_all sum.neutral) 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 dest: tagged_partial_division_ofD) 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 tagged_partial_division_ofD 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 by (force simp add: fine)+ 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" proof (rule eventually_sequentiallyI [of k]) show "\j. k \ j \ f k x \ f j x" using le x by (force intro: transitive_stepwise_le) qed 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] by (auto simp add: field_simps) 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" proof (rule tagged_partial_division_subset[of \]) show "\ tagged_partial_division_of cbox a b" using ptag tagged_division_of_def by blast qed auto 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 proof - have "\xa. k \ xa \ f k x \ f xa x" using le by (force intro: transitive_stepwise_le that) then show ?thesis using tendsto_lowerbound [OF lim [OF that]] eventually_sequentiallyI by force qed 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 proof (rule transitive_stepwise_le [OF \n \ m\ order_refl]) show "\u y z. \f u x \ f y x; f y x \ f z x\ \ f u x \ f z x" using dual_order.trans by blast qed (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 (meson K(1) K(2) content_pos_le mult_left_mono nle subsetD) 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)" using integral_norm_bound_integral[OF f integrable_linear[OF g]] by (simp add: bounded_linear_inner_left fg) 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] + (auto simp: has_real_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_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)" proof (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) show "content (cbox a b) * e < e'" using \e' > 0\ by (simp add: divide_simps e_def not_less) qed 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] show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_scaleR_left (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))) = blinfun_apply (blinfun_scaleR_left (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_vector_derivative_def\) 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 proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_mult_right (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))) = blinfun_apply (blinfun_mult_right (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_field_derivative_def\) 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 + using deriv has_real_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)" proof (intro integrable_continuous continuous_on_compose [OF _ continuous_on_subset [OF con]]) show "continuous_on (cbox c d) (Pair x)" by (simp add: continuous_on_Pair) show "Pair x ` cbox c d \ cbox (a,c) (b,d)" using x by blast qed 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" using dd e2_gt assms x apply (rule_tac x=dd in exI) apply clarify apply (rule le_less_trans [OF integrable_bound e2_less]) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis proof (rule integrable_continuous) show "continuous_on (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" by (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) qed 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" using k f by (auto simp: has_integral_integral intro: integral_unique [OF has_integral_split]) 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 flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split 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 integral_integrable_2dim [OF continuous_on_subset [OF f]]; 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 flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split 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" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (auto simp: integrable_diff f_int_uwvz integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) 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" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (auto simp: integrable_diff f_int_uv integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) 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)" using cbp \0 < e/content ?CBOX\ apply (intro integrable_bound [OF _ _ normint_wz]) 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 using that fine ptag \0 < k\ by (auto simp: * intro: op_acbd [OF division_of_tagged_division [OF ptag]]) 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)))" proof (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]) show "\u v. content (prod.swap ` cbox u v) = content (cbox u v)" by (metis content_Pair mult.commute old.prod.exhaust swap_cbox_Pair) show "((\(x, y). f y x) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (cbox (c, a) (d, b))" by (simp add: assms integrable_continuous integrable_integral swap_continuous) qed (use isCont_swap in \fastforce+\) 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]) + simp: has_real_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]) + simp: has_real_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) + simp: has_real_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/Infinite_Sum.thy b/src/HOL/Analysis/Infinite_Sum.thy --- a/src/HOL/Analysis/Infinite_Sum.thy +++ b/src/HOL/Analysis/Infinite_Sum.thy @@ -1,2576 +1,2610 @@ (* Title: HOL/Analysis/Infinite_Sum.thy Author: Dominique Unruh, University of Tartu Manuel Eberl, University of Innsbruck A theory of sums over possibly infinite sets. *) section \Infinite sums\ \<^latex>\\label{section:Infinite_Sum}\ text \In this theory, we introduce the definition of infinite sums, i.e., sums ranging over an infinite, potentially uncountable index set with no particular ordering. (This is different from series. Those are sums indexed by natural numbers, and the order of the index set matters.) Our definition is quite standard: $s:=\sum_{x\in A} f(x)$ is the limit of finite sums $s_F:=\sum_{x\in F} f(x)$ for increasing $F$. That is, $s$ is the limit of the net $s_F$ where $F$ are finite subsets of $A$ ordered by inclusion. We believe that this is the standard definition for such sums. See, e.g., Definition 4.11 in \cite{conway2013course}. This definition is quite general: it is well-defined whenever $f$ takes values in some commutative monoid endowed with a Hausdorff topology. (Examples are reals, complex numbers, normed vector spaces, and more.)\ theory Infinite_Sum imports Elementary_Topology "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Complex_Order" begin subsection \Definition and syntax\ definition has_sum :: \('a \ 'b :: {comm_monoid_add, topological_space}) \ 'a set \ 'b \ bool\ where \has_sum f A x \ (sum f \ x) (finite_subsets_at_top A)\ definition summable_on :: "('a \ 'b::{comm_monoid_add, topological_space}) \ 'a set \ bool" (infixr "summable'_on" 46) where "f summable_on A \ (\x. has_sum f A x)" definition infsum :: "('a \ 'b::{comm_monoid_add,t2_space}) \ 'a set \ 'b" where "infsum f A = (if f summable_on A then Lim (finite_subsets_at_top A) (sum f) else 0)" abbreviation abs_summable_on :: "('a \ 'b::real_normed_vector) \ 'a set \ bool" (infixr "abs'_summable'_on" 46) where "f abs_summable_on A \ (\x. norm (f x)) summable_on A" syntax (ASCII) "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10) syntax "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(2\\<^sub>\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\\<^sub>\i\A. b" \ "CONST infsum (\i. b) A" syntax (ASCII) "_univinfsum" :: "pttrn \ 'a \ 'a" ("(3INFSUM _./ _)" [0, 10] 10) syntax "_univinfsum" :: "pttrn \ 'a \ 'a" ("(2\\<^sub>\_./ _)" [0, 10] 10) translations "\\<^sub>\x. t" \ "CONST infsum (\x. t) (CONST UNIV)" syntax (ASCII) "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10) syntax "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\\<^sub>\_ | (_)./ _)" [0, 0, 10] 10) translations "\\<^sub>\x|P. t" => "CONST infsum (\x. t) {x. P}" print_translation \ let fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] = if x <> y then raise Match else let val x' = Syntax_Trans.mark_bound_body (x, Tx); val t' = subst_bound (x', t); val P' = subst_bound (x', P); in Syntax.const @{syntax_const "_qinfsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' end | sum_tr' _ = raise Match; in [(@{const_syntax infsum}, K sum_tr')] end \ subsection \General properties\ lemma infsumI: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \has_sum f A x\ shows \infsum f A = x\ by (metis assms finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim) lemma infsum_eqI: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \x = y\ assumes \has_sum f A x\ assumes \has_sum g B y\ shows \infsum f A = infsum g B\ by (metis assms(1) assms(2) assms(3) finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim) lemma infsum_eqI': fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\x. has_sum f A x \ has_sum g B x\ shows \infsum f A = infsum g B\ by (metis assms infsum_def infsum_eqI summable_on_def) lemma infsum_not_exists: fixes f :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\ f summable_on A\ shows \infsum f A = 0\ by (simp add: assms infsum_def) lemma summable_iff_has_sum_infsum: "f summable_on A \ has_sum f A (infsum f A)" using infsumI summable_on_def by blast lemma has_sum_infsum[simp]: assumes \f summable_on S\ shows \has_sum f S (infsum f S)\ using assms by (auto simp: summable_on_def infsum_def has_sum_def tendsto_Lim) lemma has_sum_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, topological_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows "has_sum f S x \ has_sum g T x" proof - have \eventually P (filtermap (sum f) (finite_subsets_at_top S)) = eventually P (filtermap (sum g) (finite_subsets_at_top T))\ for P proof assume \eventually P (filtermap (sum f) (finite_subsets_at_top S))\ then obtain F0 where \finite F0\ and \F0 \ S\ and F0_P: \\F. finite F \ F \ S \ F \ F0 \ P (sum f F)\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) define F0' where \F0' = F0 \ T\ have [simp]: \finite F0'\ \F0' \ T\ by (simp_all add: F0'_def \finite F0\) have \P (sum g F)\ if \finite F\ \F \ T\ \F \ F0'\ for F proof - have \P (sum f ((F\S) \ (F0\S)))\ apply (rule F0_P) using \F0 \ S\ \finite F0\ that by auto also have \sum f ((F\S) \ (F0\S)) = sum g F\ apply (rule sum.mono_neutral_cong) using that \finite F0\ F0'_def assms by auto finally show ?thesis . qed with \F0' \ T\ \finite F0'\ show \eventually P (filtermap (sum g) (finite_subsets_at_top T))\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) next assume \eventually P (filtermap (sum g) (finite_subsets_at_top T))\ then obtain F0 where \finite F0\ and \F0 \ T\ and F0_P: \\F. finite F \ F \ T \ F \ F0 \ P (sum g F)\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) define F0' where \F0' = F0 \ S\ have [simp]: \finite F0'\ \F0' \ S\ by (simp_all add: F0'_def \finite F0\) have \P (sum f F)\ if \finite F\ \F \ S\ \F \ F0'\ for F proof - have \P (sum g ((F\T) \ (F0\T)))\ apply (rule F0_P) using \F0 \ T\ \finite F0\ that by auto also have \sum g ((F\T) \ (F0\T)) = sum f F\ apply (rule sum.mono_neutral_cong) using that \finite F0\ F0'_def assms by auto finally show ?thesis . qed with \F0' \ S\ \finite F0'\ show \eventually P (filtermap (sum f) (finite_subsets_at_top S))\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) qed then have tendsto_x: "(sum f \ x) (finite_subsets_at_top S) \ (sum g \ x) (finite_subsets_at_top T)" for x by (simp add: le_filter_def filterlim_def) then show ?thesis by (simp add: has_sum_def) qed lemma summable_on_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, topological_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows "f summable_on S \ g summable_on T" using has_sum_cong_neutral[of T S g f, OF assms] by (simp add: summable_on_def) lemma infsum_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows \infsum f S = infsum g T\ apply (rule infsum_eqI') using assms by (rule has_sum_cong_neutral) lemma has_sum_cong: assumes "\x. x\A \ f x = g x" shows "has_sum f A x \ has_sum g A x" using assms by (intro has_sum_cong_neutral) auto lemma summable_on_cong: assumes "\x. x\A \ f x = g x" shows "f summable_on A \ g summable_on A" by (metis assms summable_on_def has_sum_cong) lemma infsum_cong: assumes "\x. x\A \ f x = g x" shows "infsum f A = infsum g A" using assms infsum_eqI' has_sum_cong by blast lemma summable_on_cofin_subset: fixes f :: "'a \ 'b::topological_ab_group_add" assumes "f summable_on A" and [simp]: "finite F" shows "f summable_on (A - F)" proof - from assms(1) obtain x where lim_f: "(sum f \ x) (finite_subsets_at_top A)" unfolding summable_on_def has_sum_def by auto define F' where "F' = F\A" with assms have "finite F'" and "A-F = A-F'" by auto have "filtermap ((\)F') (finite_subsets_at_top (A-F)) \ finite_subsets_at_top A" proof (rule filter_leI) fix P assume "eventually P (finite_subsets_at_top A)" then obtain X where [simp]: "finite X" and XA: "X \ A" and P: "\Y. finite Y \ X \ Y \ Y \ A \ P Y" unfolding eventually_finite_subsets_at_top by auto define X' where "X' = X-F" hence [simp]: "finite X'" and [simp]: "X' \ A-F" using XA by auto hence "finite Y \ X' \ Y \ Y \ A - F \ P (F' \ Y)" for Y using P XA unfolding X'_def using F'_def \finite F'\ by blast thus "eventually P (filtermap ((\) F') (finite_subsets_at_top (A - F)))" unfolding eventually_filtermap eventually_finite_subsets_at_top by (rule_tac x=X' in exI, simp) qed with lim_f have "(sum f \ x) (filtermap ((\)F') (finite_subsets_at_top (A-F)))" using tendsto_mono by blast have "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))" if "((sum f \ (\) F') \ x) (finite_subsets_at_top (A - F))" using that unfolding o_def by auto hence "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A-F))" using tendsto_compose_filtermap [symmetric] by (simp add: \(sum f \ x) (filtermap ((\) F') (finite_subsets_at_top (A - F)))\ tendsto_compose_filtermap) have "\Y. finite Y \ Y \ A - F \ sum f (F' \ Y) = sum f F' + sum f Y" by (metis Diff_disjoint Int_Diff \A - F = A - F'\ \finite F'\ inf.orderE sum.union_disjoint) hence "\\<^sub>F x in finite_subsets_at_top (A - F). sum f (F' \ x) = sum f F' + sum f x" unfolding eventually_finite_subsets_at_top using exI [where x = "{}"] by (simp add: \\P. P {} \ \x. P x\) hence "((\G. sum f F' + sum f G) \ x) (finite_subsets_at_top (A-F))" using tendsto_cong [THEN iffD1 , rotated] \((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))\ by fastforce hence "((\G. sum f F' + sum f G) \ sum f F' + (x-sum f F')) (finite_subsets_at_top (A-F))" by simp hence "(sum f \ x - sum f F') (finite_subsets_at_top (A-F))" using tendsto_add_const_iff by blast thus "f summable_on (A - F)" unfolding summable_on_def has_sum_def by auto qed lemma fixes f :: "'a \ 'b::{topological_ab_group_add}" assumes \has_sum f B b\ and \has_sum f A a\ and AB: "A \ B" shows has_sum_Diff: "has_sum f (B - A) (b - a)" proof - have finite_subsets1: "finite_subsets_at_top (B - A) \ filtermap (\F. F - A) (finite_subsets_at_top B)" proof (rule filter_leI) fix P assume "eventually P (filtermap (\F. F - A) (finite_subsets_at_top B))" then obtain X where "finite X" and "X \ B" and P: "finite Y \ X \ Y \ Y \ B \ P (Y - A)" for Y unfolding eventually_filtermap eventually_finite_subsets_at_top by auto hence "finite (X-A)" and "X-A \ B - A" by auto moreover have "finite Y \ X-A \ Y \ Y \ B - A \ P Y" for Y using P[where Y="Y\X"] \finite X\ \X \ B\ by (metis Diff_subset Int_Diff Un_Diff finite_Un inf.orderE le_sup_iff sup.orderE sup_ge2) ultimately show "eventually P (finite_subsets_at_top (B - A))" unfolding eventually_finite_subsets_at_top by meson qed have finite_subsets2: "filtermap (\F. F \ A) (finite_subsets_at_top B) \ finite_subsets_at_top A" apply (rule filter_leI) using assms unfolding eventually_filtermap eventually_finite_subsets_at_top by (metis Int_subset_iff finite_Int inf_le2 subset_trans) from assms(1) have limB: "(sum f \ b) (finite_subsets_at_top B)" using has_sum_def by auto from assms(2) have limA: "(sum f \ a) (finite_subsets_at_top A)" using has_sum_def by blast have "((\F. sum f (F\A)) \ a) (finite_subsets_at_top B)" proof (subst asm_rl [of "(\F. sum f (F\A)) = sum f o (\F. F\A)"]) show "(\F. sum f (F \ A)) = sum f \ (\F. F \ A)" unfolding o_def by auto show "((sum f \ (\F. F \ A)) \ a) (finite_subsets_at_top B)" unfolding o_def using tendsto_compose_filtermap finite_subsets2 limA tendsto_mono \(\F. sum f (F \ A)) = sum f \ (\F. F \ A)\ by fastforce qed with limB have "((\F. sum f F - sum f (F\A)) \ b - a) (finite_subsets_at_top B)" using tendsto_diff by blast have "sum f X - sum f (X \ A) = sum f (X - A)" if "finite X" and "X \ B" for X :: "'a set" using that by (metis add_diff_cancel_left' sum.Int_Diff) hence "\\<^sub>F x in finite_subsets_at_top B. sum f x - sum f (x \ A) = sum f (x - A)" by (rule eventually_finite_subsets_at_top_weakI) hence "((\F. sum f (F-A)) \ b - a) (finite_subsets_at_top B)" using tendsto_cong [THEN iffD1 , rotated] \((\F. sum f F - sum f (F \ A)) \ b - a) (finite_subsets_at_top B)\ by fastforce hence "(sum f \ b - a) (filtermap (\F. F-A) (finite_subsets_at_top B))" by (subst tendsto_compose_filtermap[symmetric], simp add: o_def) hence limBA: "(sum f \ b - a) (finite_subsets_at_top (B-A))" apply (rule tendsto_mono[rotated]) by (rule finite_subsets1) thus ?thesis by (simp add: has_sum_def) qed lemma fixes f :: "'a \ 'b::{topological_ab_group_add}" assumes "f summable_on B" and "f summable_on A" and "A \ B" shows summable_on_Diff: "f summable_on (B-A)" by (meson assms summable_on_def has_sum_Diff) lemma fixes f :: "'a \ 'b::{topological_ab_group_add,t2_space}" assumes "f summable_on B" and "f summable_on A" and AB: "A \ B" shows infsum_Diff: "infsum f (B - A) = infsum f B - infsum f A" by (metis AB assms has_sum_Diff infsumI summable_on_def) lemma has_sum_mono_neutral: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" (* Does this really require a linorder topology? (Instead of order topology.) *) assumes \has_sum f A a\ and "has_sum g B b" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows "a \ b" proof - define f' g' where \f' x = (if x \ A then f x else 0)\ and \g' x = (if x \ B then g x else 0)\ for x have [simp]: \f summable_on A\ \g summable_on B\ using assms(1,2) summable_on_def by auto have \has_sum f' (A\B) a\ apply (subst has_sum_cong_neutral[where g=f and T=A]) by (auto simp: f'_def assms(1)) then have f'_lim: \(sum f' \ a) (finite_subsets_at_top (A\B))\ by (meson has_sum_def) have \has_sum g' (A\B) b\ apply (subst has_sum_cong_neutral[where g=g and T=B]) by (auto simp: g'_def assms(2)) then have g'_lim: \(sum g' \ b) (finite_subsets_at_top (A\B))\ using has_sum_def by blast have *: \\\<^sub>F x in finite_subsets_at_top (A \ B). sum f' x \ sum g' x\ apply (rule eventually_finite_subsets_at_top_weakI) apply (rule sum_mono) using assms by (auto simp: f'_def g'_def) show ?thesis apply (rule tendsto_le) using * g'_lim f'_lim by auto qed lemma infsum_mono_neutral: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" and "g summable_on B" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows "infsum f A \ infsum g B" by (rule has_sum_mono_neutral[of f A _ g B _]) (use assms in \auto intro: has_sum_infsum\) lemma has_sum_mono: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A x" and "has_sum g A y" assumes \\x. x \ A \ f x \ g x\ shows "x \ y" apply (rule has_sum_mono_neutral) using assms by auto lemma infsum_mono: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" and "g summable_on A" assumes \\x. x \ A \ f x \ g x\ shows "infsum f A \ infsum g A" apply (rule infsum_mono_neutral) using assms by auto lemma has_sum_finite[simp]: assumes "finite F" shows "has_sum f F (sum f F)" using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_def has_sum_def principal_eq_bot_iff) lemma summable_on_finite[simp]: fixes f :: \'a \ 'b::{comm_monoid_add,topological_space}\ assumes "finite F" shows "f summable_on F" using assms summable_on_def has_sum_finite by blast lemma infsum_finite[simp]: assumes "finite F" shows "infsum f F = sum f F" using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_def principal_eq_bot_iff) lemma has_sum_finite_approximation: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" assumes "has_sum f A x" and "\ > 0" shows "\F. finite F \ F \ A \ dist (sum f F) x \ \" proof - have "(sum f \ x) (finite_subsets_at_top A)" by (meson assms(1) has_sum_def) hence *: "\\<^sub>F F in (finite_subsets_at_top A). dist (sum f F) x < \" using assms(2) by (rule tendstoD) thus ?thesis unfolding eventually_finite_subsets_at_top by fastforce qed lemma infsum_finite_approximation: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" assumes "f summable_on A" and "\ > 0" shows "\F. finite F \ F \ A \ dist (sum f F) (infsum f A) \ \" proof - from assms have "has_sum f A (infsum f A)" by (simp add: summable_iff_has_sum_infsum) from this and \\ > 0\ show ?thesis by (rule has_sum_finite_approximation) qed lemma abs_summable_summable: fixes f :: \'a \ 'b :: banach\ assumes \f abs_summable_on A\ shows \f summable_on A\ proof - from assms obtain L where lim: \(sum (\x. norm (f x)) \ L) (finite_subsets_at_top A)\ unfolding has_sum_def summable_on_def by blast then have *: \cauchy_filter (filtermap (sum (\x. norm (f x))) (finite_subsets_at_top A))\ by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) have \\P. eventually P (finite_subsets_at_top A) \ (\F F'. P F \ P F' \ dist (sum f F) (sum f F') < e)\ if \e>0\ for e proof - define d P where \d = e/4\ and \P F \ finite F \ F \ A \ dist (sum (\x. norm (f x)) F) L < d\ for F then have \d > 0\ by (simp add: d_def that) have ev_P: \eventually P (finite_subsets_at_top A)\ using lim by (auto simp add: P_def[abs_def] \0 < d\ eventually_conj_iff eventually_finite_subsets_at_top_weakI tendsto_iff) moreover have \dist (sum f F1) (sum f F2) < e\ if \P F1\ and \P F2\ for F1 F2 proof - from ev_P obtain F' where \finite F'\ and \F' \ A\ and P_sup_F': \finite F \ F \ F' \ F \ A \ P F\ for F by atomize_elim (simp add: eventually_finite_subsets_at_top) define F where \F = F' \ F1 \ F2\ have \finite F\ and \F \ A\ using F_def P_def[abs_def] that \finite F'\ \F' \ A\ by auto have dist_F: \dist (sum (\x. norm (f x)) F) L < d\ by (metis F_def \F \ A\ P_def P_sup_F' \finite F\ le_supE order_refl) have dist_F_subset: \dist (sum f F) (sum f F') < 2*d\ if F': \F' \ F\ \P F'\ for F' proof - have \dist (sum f F) (sum f F') = norm (sum f (F-F'))\ unfolding dist_norm using \finite F\ F' by (subst sum_diff) auto also have \\ \ norm (\x\F-F'. norm (f x))\ by (rule order.trans[OF sum_norm_le[OF order.refl]]) auto also have \\ = dist (\x\F. norm (f x)) (\x\F'. norm (f x))\ unfolding dist_norm using \finite F\ F' by (subst sum_diff) auto also have \\ < 2 * d\ using dist_F F' unfolding P_def dist_norm real_norm_def by linarith finally show \dist (sum f F) (sum f F') < 2*d\ . qed have \dist (sum f F1) (sum f F2) \ dist (sum f F) (sum f F1) + dist (sum f F) (sum f F2)\ by (rule dist_triangle3) also have \\ < 2 * d + 2 * d\ by (intro add_strict_mono dist_F_subset that) (auto simp: F_def) also have \\ \ e\ by (auto simp: d_def) finally show \dist (sum f F1) (sum f F2) < e\ . qed then show ?thesis using ev_P by blast qed then have \cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\ by (simp add: cauchy_filter_metric_filtermap) then obtain L' where \(sum f \ L') (finite_subsets_at_top A)\ apply atomize_elim unfolding filterlim_def apply (rule complete_uniform[where S=UNIV, simplified, THEN iffD1, rule_format]) apply (auto simp add: filtermap_bot_iff) by (meson Cauchy_convergent UNIV_I complete_def convergent_def) then show ?thesis using summable_on_def has_sum_def by blast qed text \The converse of @{thm [source] abs_summable_summable} does not hold: Consider the Hilbert space of square-summable sequences. Let $e_i$ denote the sequence with 1 in the $i$th position and 0 elsewhere. Let $f(i) := e_i/i$ for $i\geq1$. We have \<^term>\\ f abs_summable_on UNIV\ because $\lVert f(i)\rVert=1/i$ and thus the sum over $\lVert f(i)\rVert$ diverges. On the other hand, we have \<^term>\f summable_on UNIV\; the limit is the sequence with $1/i$ in the $i$th position. (We have not formalized this separating example here because to the best of our knowledge, this Hilbert space has not been formalized in Isabelle/HOL yet.)\ lemma norm_has_sum_bound: fixes f :: "'b \ 'a::real_normed_vector" and A :: "'b set" assumes "has_sum (\x. norm (f x)) A n" assumes "has_sum f A a" shows "norm a \ n" proof - have "norm a \ n + \" if "\>0" for \ proof- have "\F. norm (a - sum f F) \ \ \ finite F \ F \ A" using has_sum_finite_approximation[where A=A and f=f and \="\"] assms \0 < \\ by (metis dist_commute dist_norm) then obtain F where "norm (a - sum f F) \ \" and "finite F" and "F \ A" by (simp add: atomize_elim) hence "norm a \ norm (sum f F) + \" by (metis add.commute diff_add_cancel dual_order.refl norm_triangle_mono) also have "\ \ sum (\x. norm (f x)) F + \" using norm_sum by auto also have "\ \ n + \" apply (rule add_right_mono) apply (rule has_sum_mono_neutral[where A=F and B=A and f=\\x. norm (f x)\ and g=\\x. norm (f x)\]) using \finite F\ \F \ A\ assms by auto finally show ?thesis by assumption qed thus ?thesis using linordered_field_class.field_le_epsilon by blast qed lemma norm_infsum_bound: fixes f :: "'b \ 'a::real_normed_vector" and A :: "'b set" assumes "f abs_summable_on A" shows "norm (infsum f A) \ infsum (\x. norm (f x)) A" proof (cases "f summable_on A") case True show ?thesis apply (rule norm_has_sum_bound[where A=A and f=f and a=\infsum f A\ and n=\infsum (\x. norm (f x)) A\]) using assms True by (metis finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim)+ next case False obtain t where t_def: "(sum (\x. norm (f x)) \ t) (finite_subsets_at_top A)" using assms unfolding summable_on_def has_sum_def by blast have sumpos: "sum (\x. norm (f x)) X \ 0" for X by (simp add: sum_nonneg) have tgeq0:"t \ 0" proof(rule ccontr) define S::"real set" where "S = {s. s < 0}" assume "\ 0 \ t" hence "t < 0" by simp hence "t \ S" unfolding S_def by blast moreover have "open S" proof- have "closed {s::real. s \ 0}" using Elementary_Topology.closed_sequential_limits[where S = "{s::real. s \ 0}"] by (metis Lim_bounded2 mem_Collect_eq) moreover have "{s::real. s \ 0} = UNIV - S" unfolding S_def by auto ultimately have "closed (UNIV - S)" by simp thus ?thesis by (simp add: Compl_eq_Diff_UNIV open_closed) qed ultimately have "\\<^sub>F X in finite_subsets_at_top A. (\x\X. norm (f x)) \ S" using t_def unfolding tendsto_def by blast hence "\X. (\x\X. norm (f x)) \ S" by (metis (no_types, lifting) eventually_mono filterlim_iff finite_subsets_at_top_neq_bot tendsto_Lim) then obtain X where "(\x\X. norm (f x)) \ S" by blast hence "(\x\X. norm (f x)) < 0" unfolding S_def by auto thus False by (simp add: leD sumpos) qed have "\!h. (sum (\x. norm (f x)) \ h) (finite_subsets_at_top A)" using t_def finite_subsets_at_top_neq_bot tendsto_unique by blast hence "t = (Topological_Spaces.Lim (finite_subsets_at_top A) (sum (\x. norm (f x))))" using t_def unfolding Topological_Spaces.Lim_def by (metis the_equality) hence "Lim (finite_subsets_at_top A) (sum (\x. norm (f x))) \ 0" using tgeq0 by blast thus ?thesis unfolding infsum_def using False by auto qed lemma infsum_tendsto: assumes \f summable_on S\ shows \((\F. sum f F) \ infsum f S) (finite_subsets_at_top S)\ using assms by (simp flip: has_sum_def) lemma has_sum_0: assumes \\x. x\M \ f x = 0\ shows \has_sum f M 0\ unfolding has_sum_def apply (subst tendsto_cong[where g=\\_. 0\]) apply (rule eventually_finite_subsets_at_top_weakI) using assms by (auto simp add: subset_iff) lemma summable_on_0: assumes \\x. x\M \ f x = 0\ shows \f summable_on M\ using assms summable_on_def has_sum_0 by blast lemma infsum_0: assumes \\x. x\M \ f x = 0\ shows \infsum f M = 0\ by (metis assms finite_subsets_at_top_neq_bot infsum_def has_sum_0 has_sum_def tendsto_Lim) text \Variants of @{thm [source] infsum_0} etc. suitable as simp-rules\ lemma infsum_0_simp[simp]: \infsum (\_. 0) M = 0\ by (simp_all add: infsum_0) lemma summable_on_0_simp[simp]: \(\_. 0) summable_on M\ by (simp_all add: summable_on_0) lemma has_sum_0_simp[simp]: \has_sum (\_. 0) M 0\ by (simp_all add: has_sum_0) lemma has_sum_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add}" assumes \has_sum f A a\ assumes \has_sum g A b\ shows \has_sum (\x. f x + g x) A (a + b)\ proof - from assms have lim_f: \(sum f \ a) (finite_subsets_at_top A)\ and lim_g: \(sum g \ b) (finite_subsets_at_top A)\ by (simp_all add: has_sum_def) then have lim: \(sum (\x. f x + g x) \ a + b) (finite_subsets_at_top A)\ unfolding sum.distrib by (rule tendsto_add) then show ?thesis by (simp_all add: has_sum_def) qed lemma summable_on_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add}" assumes \f summable_on A\ assumes \g summable_on A\ shows \(\x. f x + g x) summable_on A\ by (metis (full_types) assms(1) assms(2) summable_on_def has_sum_add) lemma infsum_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes \f summable_on A\ assumes \g summable_on A\ shows \infsum (\x. f x + g x) A = infsum f A + infsum g A\ proof - have \has_sum (\x. f x + g x) A (infsum f A + infsum g A)\ by (simp add: assms(1) assms(2) has_sum_add) then show ?thesis using infsumI by blast qed lemma has_sum_Un_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes "has_sum f A a" assumes "has_sum f B b" assumes disj: "A \ B = {}" shows \has_sum f (A \ B) (a + b)\ proof - define fA fB where \fA x = (if x \ A then f x else 0)\ and \fB x = (if x \ A then f x else 0)\ for x have fA: \has_sum fA (A \ B) a\ apply (subst has_sum_cong_neutral[where T=A and g=f]) using assms by (auto simp: fA_def) have fB: \has_sum fB (A \ B) b\ apply (subst has_sum_cong_neutral[where T=B and g=f]) using assms by (auto simp: fB_def) have fAB: \f x = fA x + fB x\ for x unfolding fA_def fB_def by simp show ?thesis unfolding fAB using fA fB by (rule has_sum_add) qed lemma summable_on_Un_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes "f summable_on A" assumes "f summable_on B" assumes disj: "A \ B = {}" shows \f summable_on (A \ B)\ by (meson assms(1) assms(2) disj summable_on_def has_sum_Un_disjoint) lemma infsum_Un_disjoint: fixes f :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes "f summable_on A" assumes "f summable_on B" assumes disj: "A \ B = {}" shows \infsum f (A \ B) = infsum f A + infsum f B\ by (intro infsumI has_sum_Un_disjoint has_sum_infsum assms) -(* TODO move *) -lemma (in uniform_space) cauchy_filter_complete_converges: - assumes "cauchy_filter F" "complete A" "F \ principal A" "F \ bot" - shows "\c. F \ nhds c" - using assms unfolding complete_uniform by blast +lemma norm_summable_imp_has_sum: + fixes f :: "nat \ 'a :: banach" + assumes "summable (\n. norm (f n))" and "f sums S" + shows "has_sum f (UNIV :: nat set) S" + unfolding has_sum_def tendsto_iff eventually_finite_subsets_at_top +proof (safe, goal_cases) + case (1 \) + from assms(1) obtain S' where S': "(\n. norm (f n)) sums S'" + by (auto simp: summable_def) + with 1 obtain N where N: "\n. n \ N \ \S' - (\i < \" + by (auto simp: tendsto_iff eventually_at_top_linorder sums_def dist_norm abs_minus_commute) + + show ?case + proof (rule exI[of _ "{..n. if n \ Y then 0 else f n) sums (S - sum f Y)" + by (intro sums_If_finite_set'[OF \f sums S\]) (auto simp: sum_negf) + hence "S - sum f Y = (\n. if n \ Y then 0 else f n)" + by (simp add: sums_iff) + also have "norm \ \ (\n. norm (if n \ Y then 0 else f n))" + by (rule summable_norm[OF summable_comparison_test'[OF assms(1)]]) auto + also have "\ \ (\n. if n < N then 0 else norm (f n))" + using 2 by (intro suminf_le summable_comparison_test'[OF assms(1)]) auto + also have "(\n. if n \ {..in. if n < N then 0 else norm (f n)) = S' - (\ii \S' - (\i" by simp + also have "\ < \" by (rule N) auto + finally show ?case by (simp add: dist_norm norm_minus_commute) + qed auto +qed + +lemma norm_summable_imp_summable_on: + fixes f :: "nat \ 'a :: banach" + assumes "summable (\n. norm (f n))" + shows "f summable_on UNIV" + using norm_summable_imp_has_sum[OF assms, of "suminf f"] assms + by (auto simp: sums_iff summable_on_def dest: summable_norm_cancel) text \The following lemma indeed needs a complete space (as formalized by the premise \<^term>\complete UNIV\). The following two counterexamples show this: \begin{itemize} \item Consider the real vector space $V$ of sequences with finite support, and with the $\ell_2$-norm (sum of squares). Let $e_i$ denote the sequence with a $1$ at position $i$. Let $f : \mathbb Z \to V$ be defined as $f(n) := e_{\lvert n\rvert} / n$ (with $f(0) := 0$). We have that $\sum_{n\in\mathbb Z} f(n) = 0$ (it even converges absolutely). But $\sum_{n\in\mathbb N} f(n)$ does not exist (it would converge against a sequence with infinite support). \item Let $f$ be a positive rational valued function such that $\sum_{x\in B} f(x)$ is $\sqrt 2$ and $\sum_{x\in A} f(x)$ is 1 (over the reals, with $A\subseteq B$). Then $\sum_{x\in B} f(x)$ does not exist over the rationals. But $\sum_{x\in A} f(x)$ exists. \end{itemize} The lemma also requires uniform continuity of the addition. And example of a topological group with continuous but not uniformly continuous addition would be the positive reals with the usual multiplication as the addition. We do not know whether the lemma would also hold for such topological groups.\ lemma summable_on_subset: fixes A B and f :: \'a \ 'b::{ab_group_add, uniform_space}\ assumes \complete (UNIV :: 'b set)\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'b,y). x+y)\ assumes \f summable_on A\ assumes \B \ A\ shows \f summable_on B\ proof - let ?filter_fB = \filtermap (sum f) (finite_subsets_at_top B)\ from \f summable_on A\ obtain S where \(sum f \ S) (finite_subsets_at_top A)\ (is \(sum f \ S) ?filter_A\) using summable_on_def has_sum_def by blast then have cauchy_fA: \cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\ (is \cauchy_filter ?filter_fA\) by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) have \cauchy_filter (filtermap (sum f) (finite_subsets_at_top B))\ proof (unfold cauchy_filter_def, rule filter_leI) fix E :: \('b\'b) \ bool\ assume \eventually E uniformity\ then obtain E' where \eventually E' uniformity\ and E'E'E: \E' (x, y) \ E' (y, z) \ E (x, z)\ for x y z using uniformity_trans by blast obtain D where \eventually D uniformity\ and DE: \D (x, y) \ E' (x+c, y+c)\ for x y c using plus_cont \eventually E' uniformity\ unfolding uniformly_continuous_on_uniformity filterlim_def le_filter_def uniformity_prod_def by (auto simp: case_prod_beta eventually_filtermap eventually_prod_same uniformity_refl) have DE': "E' (x, y)" if "D (x + c, y + c)" for x y c using DE[of "x + c" "y + c" "-c"] that by simp from \eventually D uniformity\ and cauchy_fA have \eventually D (?filter_fA \\<^sub>F ?filter_fA)\ unfolding cauchy_filter_def le_filter_def by simp then obtain P1 P2 where ev_P1: \eventually (\F. P1 (sum f F)) ?filter_A\ and ev_P2: \eventually (\F. P2 (sum f F)) ?filter_A\ and P1P2E: \P1 x \ P2 y \ D (x, y)\ for x y unfolding eventually_prod_filter eventually_filtermap by auto from ev_P1 obtain F1 where F1: \finite F1\ \F1 \ A\ \\F. F\F1 \ finite F \ F\A \ P1 (sum f F)\ by (metis eventually_finite_subsets_at_top) from ev_P2 obtain F2 where F2: \finite F2\ \F2 \ A\ \\F. F\F2 \ finite F \ F\A \ P2 (sum f F)\ by (metis eventually_finite_subsets_at_top) define F0 F0A F0B where \F0 \ F1 \ F2\ and \F0A \ F0 - B\ and \F0B \ F0 \ B\ have [simp]: \finite F0\ \F0 \ A\ using \F1 \ A\ \F2 \ A\ \finite F1\ \finite F2\ unfolding F0_def by blast+ have *: "E' (sum f F1', sum f F2')" if "F1'\F0B" "F2'\F0B" "finite F1'" "finite F2'" "F1'\B" "F2'\B" for F1' F2' proof (intro DE'[where c = "sum f F0A"] P1P2E) have "P1 (sum f (F1' \ F0A))" using that assms F1(1,2) F2(1,2) by (intro F1) (auto simp: F0A_def F0B_def F0_def) thus "P1 (sum f F1' + sum f F0A)" by (subst (asm) sum.union_disjoint) (use that in \auto simp: F0A_def\) next have "P2 (sum f (F2' \ F0A))" using that assms F1(1,2) F2(1,2) by (intro F2) (auto simp: F0A_def F0B_def F0_def) thus "P2 (sum f F2' + sum f F0A)" by (subst (asm) sum.union_disjoint) (use that in \auto simp: F0A_def\) qed show \eventually E (?filter_fB \\<^sub>F ?filter_fB)\ unfolding eventually_prod_filter proof (safe intro!: exI) show "eventually (\x. E' (x, sum f F0B)) (filtermap (sum f) (finite_subsets_at_top B))" and "eventually (\x. E' (sum f F0B, x)) (filtermap (sum f) (finite_subsets_at_top B))" unfolding eventually_filtermap eventually_finite_subsets_at_top by (rule exI[of _ F0B]; use * in \force simp: F0B_def\)+ next show "E (x, y)" if "E' (x, sum f F0B)" and "E' (sum f F0B, y)" for x y using E'E'E that by blast qed qed then obtain x where \?filter_fB \ nhds x\ using cauchy_filter_complete_converges[of ?filter_fB UNIV] \complete (UNIV :: _)\ by (auto simp: filtermap_bot_iff) then have \(sum f \ x) (finite_subsets_at_top B)\ by (auto simp: filterlim_def) then show ?thesis by (auto simp: summable_on_def has_sum_def) qed text \A special case of @{thm [source] summable_on_subset} for Banach spaces with less premises.\ lemma summable_on_subset_banach: fixes A B and f :: \'a \ 'b::banach\ assumes \f summable_on A\ assumes \B \ A\ shows \f summable_on B\ by (rule summable_on_subset[OF _ _ assms]) (auto simp: complete_def convergent_def dest!: Cauchy_convergent) lemma has_sum_empty[simp]: \has_sum f {} 0\ by (meson ex_in_conv has_sum_0) lemma summable_on_empty[simp]: \f summable_on {}\ by auto lemma infsum_empty[simp]: \infsum f {} = 0\ by simp lemma sum_has_sum: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes finite: \finite A\ assumes conv: \\a. a \ A \ has_sum f (B a) (s a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \has_sum f (\a\A. B a) (sum s A)\ using assms proof (insert finite conv disj, induction) case empty then show ?case by simp next case (insert x A) have \has_sum f (B x) (s x)\ by (simp add: insert.prems) moreover have IH: \has_sum f (\a\A. B a) (sum s A)\ using insert by simp ultimately have \has_sum f (B x \ (\a\A. B a)) (s x + sum s A)\ apply (rule has_sum_Un_disjoint) using insert by auto then show ?case using insert.hyps by auto qed lemma summable_on_finite_union_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes finite: \finite A\ assumes conv: \\a. a \ A \ f summable_on (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \f summable_on (\a\A. B a)\ using finite conv disj apply induction by (auto intro!: summable_on_Un_disjoint) lemma sum_infsum: fixes f :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes finite: \finite A\ assumes conv: \\a. a \ A \ f summable_on (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \sum (\a. infsum f (B a)) A = infsum f (\a\A. B a)\ by (rule sym, rule infsumI) (use sum_has_sum[of A f B \\a. infsum f (B a)\] assms in auto) text \The lemmas \infsum_comm_additive_general\ and \infsum_comm_additive\ (and variants) below both state that the infinite sum commutes with a continuous additive function. \infsum_comm_additive_general\ is stated more for more general type classes at the expense of a somewhat less compact formulation of the premises. E.g., by avoiding the constant \<^const>\additive\ which introduces an additional sort constraint (group instead of monoid). For example, extended reals (\<^typ>\ereal\, \<^typ>\ennreal\) are not covered by \infsum_comm_additive\.\ lemma has_sum_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,topological_space} \ 'c :: {comm_monoid_add,topological_space}\ assumes f_sum: \\F. finite F \ F \ S \ sum (f o g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes cont: \f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes infsum: \has_sum g S x\ shows \has_sum (f o g) S (f x)\ proof - have \(sum g \ x) (finite_subsets_at_top S)\ using infsum has_sum_def by blast then have \((f o sum g) \ f x) (finite_subsets_at_top S)\ apply (rule tendsto_compose_at) using assms by auto then have \(sum (f o g) \ f x) (finite_subsets_at_top S)\ apply (rule tendsto_cong[THEN iffD1, rotated]) using f_sum by fastforce then show \has_sum (f o g) S (f x)\ using has_sum_def by blast qed lemma summable_on_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,topological_space} \ 'c :: {comm_monoid_add,topological_space}\ assumes \\F. finite F \ F \ S \ sum (f o g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes \\x. has_sum g S x \ f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes \g summable_on S\ shows \(f o g) summable_on S\ by (meson assms summable_on_def has_sum_comm_additive_general has_sum_def infsum_tendsto) lemma infsum_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,t2_space} \ 'c :: {comm_monoid_add,t2_space}\ assumes f_sum: \\F. finite F \ F \ S \ sum (f o g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \infsum (f o g) S = f (infsum g S)\ using assms by (intro infsumI has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def) lemma has_sum_comm_additive: fixes f :: \'b :: {ab_group_add,topological_space} \ 'c :: {ab_group_add,topological_space}\ assumes \additive f\ assumes \f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes infsum: \has_sum g S x\ shows \has_sum (f o g) S (f x)\ using assms by (intro has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def additive.sum) lemma summable_on_comm_additive: fixes f :: \'b :: {ab_group_add,t2_space} \ 'c :: {ab_group_add,topological_space}\ assumes \additive f\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \(f o g) summable_on S\ by (meson assms(1) assms(2) assms(3) summable_on_def has_sum_comm_additive has_sum_infsum isContD) lemma infsum_comm_additive: fixes f :: \'b :: {ab_group_add,t2_space} \ 'c :: {ab_group_add,t2_space}\ assumes \additive f\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \infsum (f o g) S = f (infsum g S)\ by (rule infsum_comm_additive_general; auto simp: assms additive.sum) lemma nonneg_bdd_above_has_sum: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \has_sum f A (SUP F\{F. finite F \ F\A}. sum f F)\ proof - have \(sum f \ (SUP F\{F. finite F \ F\A}. sum f F)) (finite_subsets_at_top A)\ proof (rule order_tendstoI) fix a assume \a < (SUP F\{F. finite F \ F\A}. sum f F)\ then obtain F where \a < sum f F\ and \finite F\ and \F \ A\ by (metis (mono_tags, lifting) Collect_cong Collect_empty_eq assms(2) empty_subsetI finite.emptyI less_cSUP_iff mem_Collect_eq) show \\\<^sub>F x in finite_subsets_at_top A. a < sum f x\ unfolding eventually_finite_subsets_at_top proof (rule exI[of _ F], safe) fix Y assume Y: "finite Y" "F \ Y" "Y \ A" have "a < sum f F" by fact also have "\ \ sum f Y" using assms Y by (intro sum_mono2) auto finally show "a < sum f Y" . qed (use \finite F\ \F \ A\ in auto) next fix a assume *: \(SUP F\{F. finite F \ F\A}. sum f F) < a\ have \sum f F < a\ if \F\A\ and \finite F\ for F proof - have "sum f F \ (SUP F\{F. finite F \ F\A}. sum f F)" by (rule cSUP_upper) (use that assms(2) in \auto simp: conj_commute\) also have "\ < a" by fact finally show ?thesis . qed then show \\\<^sub>F x in finite_subsets_at_top A. sum f x < a\ by (rule eventually_finite_subsets_at_top_weakI) qed then show ?thesis using has_sum_def by blast qed lemma nonneg_bdd_above_summable_on: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \f summable_on A\ using assms(1) assms(2) summable_on_def nonneg_bdd_above_has_sum by blast lemma nonneg_bdd_above_infsum: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \infsum f A = (SUP F\{F. finite F \ F\A}. sum f F)\ using assms by (auto intro!: infsumI nonneg_bdd_above_has_sum) lemma nonneg_has_sum_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \has_sum f A (SUP F\{F. finite F \ F\A}. sum f F)\ using assms nonneg_bdd_above_has_sum by blast lemma nonneg_summable_on_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \f summable_on A\ using assms nonneg_bdd_above_summable_on by blast lemma nonneg_infsum_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \infsum f A = (SUP F\{F. finite F \ F\A}. sum f F)\ using assms nonneg_bdd_above_infsum by blast lemma has_sum_nonneg: fixes f :: "'a \ 'b::{ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f M a" and "\x. x \ M \ 0 \ f x" shows "a \ 0" by (metis (no_types, lifting) DiffD1 assms(1) assms(2) empty_iff has_sum_0 has_sum_mono_neutral order_refl) lemma infsum_nonneg: fixes f :: "'a \ 'b::{ordered_comm_monoid_add,linorder_topology}" assumes "\x. x \ M \ 0 \ f x" shows "infsum f M \ 0" (is "?lhs \ _") apply (cases \f summable_on M\) apply (metis assms infsum_0_simp summable_on_0_simp infsum_mono) using assms by (auto simp add: infsum_not_exists) lemma has_sum_mono2: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A S" "has_sum f B S'" "A \ B" assumes "\x. x \ B - A \ f x \ 0" shows "S \ S'" proof - have "has_sum f (B - A) (S' - S)" by (rule has_sum_Diff) fact+ hence "S' - S \ 0" by (rule has_sum_nonneg) (use assms(4) in auto) thus ?thesis by (metis add_0 add_mono_thms_linordered_semiring(3) diff_add_cancel) qed lemma infsum_mono2: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" "f summable_on B" "A \ B" assumes "\x. x \ B - A \ f x \ 0" shows "infsum f A \ infsum f B" by (rule has_sum_mono2[OF has_sum_infsum has_sum_infsum]) (use assms in auto) lemma finite_sum_le_has_sum: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A S" "finite B" "B \ A" assumes "\x. x \ A - B \ f x \ 0" shows "sum f B \ S" proof (rule has_sum_mono2) show "has_sum f A S" by fact show "has_sum f B (sum f B)" by (rule has_sum_finite) fact+ qed (use assms in auto) lemma finite_sum_le_infsum: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" "finite B" "B \ A" assumes "\x. x \ A - B \ f x \ 0" shows "sum f B \ infsum f A" by (rule finite_sum_le_has_sum[OF has_sum_infsum]) (use assms in auto) lemma has_sum_reindex: assumes \inj_on h A\ shows \has_sum g (h ` A) x \ has_sum (g \ h) A x\ proof - have \has_sum g (h ` A) x \ (sum g \ x) (finite_subsets_at_top (h ` A))\ by (simp add: has_sum_def) also have \\ \ ((\F. sum g (h ` F)) \ x) (finite_subsets_at_top A)\ apply (subst filtermap_image_finite_subsets_at_top[symmetric]) using assms by (auto simp: filterlim_def filtermap_filtermap) also have \\ \ (sum (g \ h) \ x) (finite_subsets_at_top A)\ apply (rule tendsto_cong) apply (rule eventually_finite_subsets_at_top_weakI) apply (rule sum.reindex) using assms subset_inj_on by blast also have \\ \ has_sum (g \ h) A x\ by (simp add: has_sum_def) finally show ?thesis . qed lemma summable_on_reindex: assumes \inj_on h A\ shows \g summable_on (h ` A) \ (g \ h) summable_on A\ by (simp add: assms summable_on_def has_sum_reindex) lemma infsum_reindex: assumes \inj_on h A\ shows \infsum g (h ` A) = infsum (g \ h) A\ by (metis (no_types, opaque_lifting) assms finite_subsets_at_top_neq_bot infsum_def summable_on_reindex has_sum_def has_sum_infsum has_sum_reindex tendsto_Lim) lemma summable_on_reindex_bij_betw: assumes "bij_betw g A B" shows "(\x. f (g x)) summable_on A \ f summable_on B" proof - thm summable_on_reindex have \(\x. f (g x)) summable_on A \ f summable_on g ` A\ apply (rule summable_on_reindex[symmetric, unfolded o_def]) using assms bij_betw_imp_inj_on by blast also have \\ \ f summable_on B\ using assms bij_betw_imp_surj_on by blast finally show ?thesis . qed lemma infsum_reindex_bij_betw: assumes "bij_betw g A B" shows "infsum (\x. f (g x)) A = infsum f B" proof - have \infsum (\x. f (g x)) A = infsum f (g ` A)\ by (metis (mono_tags, lifting) assms bij_betw_imp_inj_on infsum_cong infsum_reindex o_def) also have \\ = infsum f B\ using assms bij_betw_imp_surj_on by blast finally show ?thesis . qed lemma sum_uniformity: assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'b::{uniform_space,comm_monoid_add},y). x+y)\ assumes \eventually E uniformity\ obtains D where \eventually D uniformity\ and \\M::'a set. \f f' :: 'a \ 'b. card M \ n \ (\m\M. D (f m, f' m)) \ E (sum f M, sum f' M)\ proof (atomize_elim, insert \eventually E uniformity\, induction n arbitrary: E rule:nat_induct) case 0 then show ?case by (metis card_eq_0_iff equals0D le_zero_eq sum.infinite sum.not_neutral_contains_not_neutral uniformity_refl) next case (Suc n) from plus_cont[unfolded uniformly_continuous_on_uniformity filterlim_def le_filter_def, rule_format, OF Suc.prems] obtain D1 D2 where \eventually D1 uniformity\ and \eventually D2 uniformity\ and D1D2E: \D1 (x, y) \ D2 (x', y') \ E (x + x', y + y')\ for x y x' y' apply atomize_elim by (auto simp: eventually_prod_filter case_prod_beta uniformity_prod_def eventually_filtermap) from Suc.IH[OF \eventually D2 uniformity\] obtain D3 where \eventually D3 uniformity\ and D3: \card M \ n \ (\m\M. D3 (f m, f' m)) \ D2 (sum f M, sum f' M)\ for M :: \'a set\ and f f' by metis define D where \D x \ D1 x \ D3 x\ for x have \eventually D uniformity\ using D_def \eventually D1 uniformity\ \eventually D3 uniformity\ eventually_elim2 by blast have \E (sum f M, sum f' M)\ if \card M \ Suc n\ and DM: \\m\M. D (f m, f' m)\ for M :: \'a set\ and f f' proof (cases \card M = 0\) case True then show ?thesis by (metis Suc.prems card_eq_0_iff sum.empty sum.infinite uniformity_refl) next case False with \card M \ Suc n\ obtain N x where \card N \ n\ and \x \ N\ and \M = insert x N\ by (metis card_Suc_eq less_Suc_eq_0_disj less_Suc_eq_le) from DM have \\m. m\N \ D (f m, f' m)\ using \M = insert x N\ by blast with D3[OF \card N \ n\] have D2_N: \D2 (sum f N, sum f' N)\ using D_def by blast from DM have \D (f x, f' x)\ using \M = insert x N\ by blast then have \D1 (f x, f' x)\ by (simp add: D_def) with D2_N have \E (f x + sum f N, f' x + sum f' N)\ using D1D2E by presburger then show \E (sum f M, sum f' M)\ by (metis False \M = insert x N\ \x \ N\ card.infinite finite_insert sum.insert) qed with \eventually D uniformity\ show ?case by auto qed lemma has_sum_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add,uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "has_sum f (Sigma A B) a" assumes summableB: \\x. x\A \ has_sum (\y. f (x, y)) (B x) (b x)\ shows "has_sum b A a" proof - define F FB FA where \F = finite_subsets_at_top (Sigma A B)\ and \FB x = finite_subsets_at_top (B x)\ and \FA = finite_subsets_at_top A\ for x from summableB have sum_b: \(sum (\y. f (x, y)) \ b x) (FB x)\ if \x \ A\ for x using FB_def[abs_def] has_sum_def that by auto from summableAB have sum_S: \(sum f \ a) F\ using F_def has_sum_def by blast have finite_proj: \finite {b| b. (a,b) \ H}\ if \finite H\ for H :: \('a\'b) set\ and a apply (subst asm_rl[of \{b| b. (a,b) \ H} = snd ` {ab. ab \ H \ fst ab = a}\]) by (auto simp: image_iff that) have \(sum b \ a) FA\ proof (rule tendsto_iff_uniformity[THEN iffD2, rule_format]) fix E :: \('c \ 'c) \ bool\ assume \eventually E uniformity\ then obtain D where D_uni: \eventually D uniformity\ and DDE': \\x y z. D (x, y) \ D (y, z) \ E (x, z)\ by (metis (no_types, lifting) \eventually E uniformity\ uniformity_transE) from sum_S obtain G where \finite G\ and \G \ Sigma A B\ and G_sum: \G \ H \ H \ Sigma A B \ finite H \ D (sum f H, a)\ for H unfolding tendsto_iff_uniformity by (metis (mono_tags, lifting) D_uni F_def eventually_finite_subsets_at_top) have \finite (fst ` G)\ and \fst ` G \ A\ using \finite G\ \G \ Sigma A B\ by auto thm uniformity_prod_def define Ga where \Ga a = {b. (a,b) \ G}\ for a have Ga_fin: \finite (Ga a)\ and Ga_B: \Ga a \ B a\ for a using \finite G\ \G \ Sigma A B\ finite_proj by (auto simp: Ga_def finite_proj) have \E (sum b M, a)\ if \M \ fst ` G\ and \finite M\ and \M \ A\ for M proof - define FMB where \FMB = finite_subsets_at_top (Sigma M B)\ have \eventually (\H. D (\a\M. b a, \(a,b)\H. f (a,b))) FMB\ proof - obtain D' where D'_uni: \eventually D' uniformity\ and \card M' \ card M \ (\m\M'. D' (g m, g' m)) \ D (sum g M', sum g' M')\ for M' :: \'a set\ and g g' apply (rule sum_uniformity[OF plus_cont \eventually D uniformity\, where n=\card M\]) by auto then have D'_sum_D: \(\m\M. D' (g m, g' m)) \ D (sum g M, sum g' M)\ for g g' by auto obtain Ha where \Ha a \ Ga a\ and Ha_fin: \finite (Ha a)\ and Ha_B: \Ha a \ B a\ and D'_sum_Ha: \Ha a \ L \ L \ B a \ finite L \ D' (b a, sum (\b. f (a,b)) L)\ if \a \ A\ for a L proof - from sum_b[unfolded tendsto_iff_uniformity, rule_format, OF _ D'_uni[THEN uniformity_sym]] obtain Ha0 where \finite (Ha0 a)\ and \Ha0 a \ B a\ and \Ha0 a \ L \ L \ B a \ finite L \ D' (b a, sum (\b. f (a,b)) L)\ if \a \ A\ for a L unfolding FB_def eventually_finite_subsets_at_top unfolding prod.case by metis moreover define Ha where \Ha a = Ha0 a \ Ga a\ for a ultimately show ?thesis using that[where Ha=Ha] using Ga_fin Ga_B by auto qed have \D (\a\M. b a, \(a,b)\H. f (a,b))\ if \finite H\ and \H \ Sigma M B\ and \H \ Sigma M Ha\ for H proof - define Ha' where \Ha' a = {b| b. (a,b) \ H}\ for a have [simp]: \finite (Ha' a)\ and [simp]: \Ha' a \ Ha a\ and [simp]: \Ha' a \ B a\ if \a \ M\ for a unfolding Ha'_def using \finite H\ \H \ Sigma M B\ \Sigma M Ha \ H\ that finite_proj by auto have \Sigma M Ha' = H\ using that by (auto simp: Ha'_def) then have *: \(\(a,b)\H. f (a,b)) = (\a\M. \b\Ha' a. f (a,b))\ apply (subst sum.Sigma) using \finite M\ by auto have \D' (b a, sum (\b. f (a,b)) (Ha' a))\ if \a \ M\ for a apply (rule D'_sum_Ha) using that \M \ A\ by auto then have \D (\a\M. b a, \a\M. sum (\b. f (a,b)) (Ha' a))\ by (rule_tac D'_sum_D, auto) with * show ?thesis by auto qed moreover have \Sigma M Ha \ Sigma M B\ using Ha_B \M \ A\ by auto ultimately show ?thesis unfolding FMB_def eventually_finite_subsets_at_top by (intro exI[of _ "Sigma M Ha"]) (use Ha_fin that(2,3) in \fastforce intro!: finite_SigmaI\) qed moreover have \eventually (\H. D (\(a,b)\H. f (a,b), a)) FMB\ unfolding FMB_def eventually_finite_subsets_at_top proof (rule exI[of _ G], safe) fix Y assume Y: "finite Y" "G \ Y" "Y \ Sigma M B" have "Y \ Sigma A B" using Y \M \ A\ by blast thus "D (\(a,b)\Y. f (a, b), a)" using G_sum[of Y] Y by auto qed (use \finite G\ \G \ Sigma A B\ that in auto) ultimately have \\\<^sub>F x in FMB. E (sum b M, a)\ by eventually_elim (use DDE' in auto) then show \E (sum b M, a)\ by (rule eventually_const[THEN iffD1, rotated]) (force simp: FMB_def) qed then show \\\<^sub>F x in FA. E (sum b x, a)\ using \finite (fst ` G)\ and \fst ` G \ A\ by (auto intro!: exI[of _ \fst ` G\] simp add: FA_def eventually_finite_subsets_at_top) qed then show ?thesis by (simp add: FA_def has_sum_def) qed lemma summable_on_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "(\(x,y). f x y) summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (f x) summable_on (B x)\ shows \(\x. infsum (f x) (B x)) summable_on A\ proof - from summableAB obtain a where a: \has_sum (\(x,y). f x y) (Sigma A B) a\ using has_sum_infsum by blast from summableB have b: \\x. x\A \ has_sum (f x) (B x) (infsum (f x) (B x))\ by (auto intro!: has_sum_infsum) show ?thesis using plus_cont a b by (auto intro: has_sum_Sigma[where f=\\(x,y). f x y\, simplified] simp: summable_on_def) qed lemma infsum_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "f summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (\y. f (x, y)) summable_on (B x)\ shows "infsum f (Sigma A B) = infsum (\x. infsum (\y. f (x, y)) (B x)) A" proof - from summableAB have a: \has_sum f (Sigma A B) (infsum f (Sigma A B))\ using has_sum_infsum by blast from summableB have b: \\x. x\A \ has_sum (\y. f (x, y)) (B x) (infsum (\y. f (x, y)) (B x))\ by (auto intro!: has_sum_infsum) show ?thesis using plus_cont a b by (auto intro: infsumI[symmetric] has_sum_Sigma simp: summable_on_def) qed lemma infsum_Sigma': fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "(\(x,y). f x y) summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (f x) summable_on (B x)\ shows \infsum (\x. infsum (f x) (B x)) A = infsum (\(x,y). f x y) (Sigma A B)\ using infsum_Sigma[of \\(x,y). f x y\ A B] using assms by auto text \A special case of @{thm [source] infsum_Sigma} etc. for Banach spaces. It has less premises.\ lemma fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::banach\ assumes [simp]: "(\(x,y). f x y) summable_on (Sigma A B)" shows infsum_Sigma'_banach: \infsum (\x. infsum (f x) (B x)) A = infsum (\(x,y). f x y) (Sigma A B)\ (is ?thesis1) and summable_on_Sigma_banach: \(\x. infsum (f x) (B x)) summable_on A\ (is ?thesis2) proof - have [simp]: \(f x) summable_on (B x)\ if \x \ A\ for x proof - from assms have \(\(x,y). f x y) summable_on (Pair x ` B x)\ by (meson image_subset_iff summable_on_subset_banach mem_Sigma_iff that) then have \((\(x,y). f x y) o Pair x) summable_on (B x)\ apply (rule_tac summable_on_reindex[THEN iffD1]) by (simp add: inj_on_def) then show ?thesis by (auto simp: o_def) qed show ?thesis1 apply (rule infsum_Sigma') by auto show ?thesis2 apply (rule summable_on_Sigma) by auto qed lemma infsum_Sigma_banach: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::banach\ assumes [simp]: "f summable_on (Sigma A B)" shows \infsum (\x. infsum (\y. f (x,y)) (B x)) A = infsum f (Sigma A B)\ using assms by (subst infsum_Sigma'_banach) auto lemma infsum_swap: fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c::{comm_monoid_add,t2_space,uniform_space}" assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes \(\(x, y). f x y) summable_on (A \ B)\ assumes \\a. a\A \ (f a) summable_on B\ assumes \\b. b\B \ (\a. f a b) summable_on A\ shows \infsum (\x. infsum (\y. f x y) B) A = infsum (\y. infsum (\x. f x y) A) B\ proof - have [simp]: \(\(x, y). f y x) summable_on (B \ A)\ apply (subst product_swap[symmetric]) apply (subst summable_on_reindex) using assms by (auto simp: o_def) have \infsum (\x. infsum (\y. f x y) B) A = infsum (\(x,y). f x y) (A \ B)\ apply (subst infsum_Sigma) using assms by auto also have \\ = infsum (\(x,y). f y x) (B \ A)\ apply (subst product_swap[symmetric]) apply (subst infsum_reindex) using assms by (auto simp: o_def) also have \\ = infsum (\y. infsum (\x. f x y) A) B\ apply (subst infsum_Sigma) using assms by auto finally show ?thesis . qed lemma infsum_swap_banach: fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c::banach" assumes \(\(x, y). f x y) summable_on (A \ B)\ shows "infsum (\x. infsum (\y. f x y) B) A = infsum (\y. infsum (\x. f x y) A) B" proof - have [simp]: \(\(x, y). f y x) summable_on (B \ A)\ apply (subst product_swap[symmetric]) apply (subst summable_on_reindex) using assms by (auto simp: o_def) have \infsum (\x. infsum (\y. f x y) B) A = infsum (\(x,y). f x y) (A \ B)\ apply (subst infsum_Sigma'_banach) using assms by auto also have \\ = infsum (\(x,y). f y x) (B \ A)\ apply (subst product_swap[symmetric]) apply (subst infsum_reindex) using assms by (auto simp: o_def) also have \\ = infsum (\y. infsum (\x. f x y) A) B\ apply (subst infsum_Sigma'_banach) using assms by auto finally show ?thesis . qed lemma nonneg_infsum_le_0D: fixes f :: "'a \ 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}" assumes "infsum f A \ 0" and abs_sum: "f summable_on A" and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" proof (rule ccontr) assume \f x \ 0\ have ex: \f summable_on (A-{x})\ by (rule summable_on_cofin_subset) (use assms in auto) have pos: \infsum f (A - {x}) \ 0\ by (rule infsum_nonneg) (use nneg in auto) have [trans]: \x \ y \ y > z \ x > z\ for x y z :: 'b by auto have \infsum f A = infsum f (A-{x}) + infsum f {x}\ by (subst infsum_Un_disjoint[symmetric]) (use assms ex in \auto simp: insert_absorb\) also have \\ \ infsum f {x}\ (is \_ \ \\) using pos by (rule add_increasing) simp also have \\ = f x\ (is \_ = \\) by (subst infsum_finite) auto also have \\ > 0\ using \f x \ 0\ assms(4) nneg by fastforce finally show False using assms by auto qed lemma nonneg_has_sum_le_0D: fixes f :: "'a \ 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}" assumes "has_sum f A a" \a \ 0\ and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" by (metis assms(1) assms(2) assms(4) infsumI nonneg_infsum_le_0D summable_on_def nneg) lemma has_sum_cmult_left: fixes f :: "'a \ 'b :: {topological_semigroup_mult, semiring_0}" assumes \has_sum f A a\ shows "has_sum (\x. f x * c) A (a * c)" proof - from assms have \(sum f \ a) (finite_subsets_at_top A)\ using has_sum_def by blast then have \((\F. sum f F * c) \ a * c) (finite_subsets_at_top A)\ by (simp add: tendsto_mult_right) then have \(sum (\x. f x * c) \ a * c) (finite_subsets_at_top A)\ apply (rule tendsto_cong[THEN iffD1, rotated]) apply (rule eventually_finite_subsets_at_top_weakI) using sum_distrib_right by blast then show ?thesis using infsumI has_sum_def by blast qed lemma infsum_cmult_left: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \c \ 0 \ f summable_on A\ shows "infsum (\x. f x * c) A = infsum f A * c" proof (cases \c=0\) case True then show ?thesis by auto next case False then have \has_sum f A (infsum f A)\ by (simp add: assms) then show ?thesis by (auto intro!: infsumI has_sum_cmult_left) qed lemma summable_on_cmult_left: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \f summable_on A\ shows "(\x. f x * c) summable_on A" using assms summable_on_def has_sum_cmult_left by blast lemma has_sum_cmult_right: fixes f :: "'a \ 'b :: {topological_semigroup_mult, semiring_0}" assumes \has_sum f A a\ shows "has_sum (\x. c * f x) A (c * a)" proof - from assms have \(sum f \ a) (finite_subsets_at_top A)\ using has_sum_def by blast then have \((\F. c * sum f F) \ c * a) (finite_subsets_at_top A)\ by (simp add: tendsto_mult_left) then have \(sum (\x. c * f x) \ c * a) (finite_subsets_at_top A)\ apply (rule tendsto_cong[THEN iffD1, rotated]) apply (rule eventually_finite_subsets_at_top_weakI) using sum_distrib_left by blast then show ?thesis using infsumI has_sum_def by blast qed lemma infsum_cmult_right: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \c \ 0 \ f summable_on A\ shows \infsum (\x. c * f x) A = c * infsum f A\ proof (cases \c=0\) case True then show ?thesis by auto next case False then have \has_sum f A (infsum f A)\ by (simp add: assms) then show ?thesis by (auto intro!: infsumI has_sum_cmult_right) qed lemma summable_on_cmult_right: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \f summable_on A\ shows "(\x. c * f x) summable_on A" using assms summable_on_def has_sum_cmult_right by blast lemma summable_on_cmult_left': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" assumes \c \ 0\ shows "(\x. f x * c) summable_on A \ f summable_on A" proof assume \f summable_on A\ then show \(\x. f x * c) summable_on A\ by (rule summable_on_cmult_left) next assume \(\x. f x * c) summable_on A\ then have \(\x. f x * c * inverse c) summable_on A\ by (rule summable_on_cmult_left) then show \f summable_on A\ by (metis (no_types, lifting) assms summable_on_cong mult.assoc mult.right_neutral right_inverse) qed lemma summable_on_cmult_right': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" assumes \c \ 0\ shows "(\x. c * f x) summable_on A \ f summable_on A" proof assume \f summable_on A\ then show \(\x. c * f x) summable_on A\ by (rule summable_on_cmult_right) next assume \(\x. c * f x) summable_on A\ then have \(\x. inverse c * (c * f x)) summable_on A\ by (rule summable_on_cmult_right) then show \f summable_on A\ by (metis (no_types, lifting) assms summable_on_cong left_inverse mult.assoc mult.left_neutral) qed lemma infsum_cmult_left': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" shows "infsum (\x. f x * c) A = infsum f A * c" proof (cases \c \ 0 \ f summable_on A\) case True then show ?thesis apply (rule_tac infsum_cmult_left) by auto next case False note asm = False then show ?thesis proof (cases \c=0\) case True then show ?thesis by auto next case False with asm have nex: \\ f summable_on A\ by simp moreover have nex': \\ (\x. f x * c) summable_on A\ using asm False apply (subst summable_on_cmult_left') by auto ultimately show ?thesis unfolding infsum_def by simp qed qed lemma infsum_cmult_right': fixes f :: "'a \ 'b :: {t2_space,topological_semigroup_mult,division_ring}" shows "infsum (\x. c * f x) A = c * infsum f A" proof (cases \c \ 0 \ f summable_on A\) case True then show ?thesis apply (rule_tac infsum_cmult_right) by auto next case False note asm = False then show ?thesis proof (cases \c=0\) case True then show ?thesis by auto next case False with asm have nex: \\ f summable_on A\ by simp moreover have nex': \\ (\x. c * f x) summable_on A\ using asm False apply (subst summable_on_cmult_right') by auto ultimately show ?thesis unfolding infsum_def by simp qed qed lemma has_sum_constant[simp]: assumes \finite F\ shows \has_sum (\_. c) F (of_nat (card F) * c)\ by (metis assms has_sum_finite sum_constant) lemma infsum_constant[simp]: assumes \finite F\ shows \infsum (\_. c) F = of_nat (card F) * c\ apply (subst infsum_finite[OF assms]) by simp lemma infsum_diverge_constant: \ \This probably does not really need all of \<^class>\archimedean_field\ but Isabelle/HOL has no type class such as, e.g., "archimedean ring".\ fixes c :: \'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\ assumes \infinite A\ and \c \ 0\ shows \\ (\_. c) summable_on A\ proof (rule notI) assume \(\_. c) summable_on A\ then have \(\_. inverse c * c) summable_on A\ by (rule summable_on_cmult_right) then have [simp]: \(\_. 1::'a) summable_on A\ using assms by auto have \infsum (\_. 1) A \ d\ for d :: 'a proof - obtain n :: nat where \of_nat n \ d\ by (meson real_arch_simple) from assms obtain F where \F \ A\ and \finite F\ and \card F = n\ by (meson infinite_arbitrarily_large) note \d \ of_nat n\ also have \of_nat n = infsum (\_. 1::'a) F\ by (simp add: \card F = n\ \finite F\) also have \\ \ infsum (\_. 1::'a) A\ apply (rule infsum_mono_neutral) using \finite F\ \F \ A\ by auto finally show ?thesis . qed then show False by (meson linordered_field_no_ub not_less) qed lemma has_sum_constant_archimedean[simp]: \ \This probably does not really need all of \<^class>\archimedean_field\ but Isabelle/HOL has no type class such as, e.g., "archimedean ring".\ fixes c :: \'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\ shows \infsum (\_. c) A = of_nat (card A) * c\ apply (cases \finite A\) apply simp apply (cases \c = 0\) apply simp by (simp add: infsum_diverge_constant infsum_not_exists) lemma has_sum_uminus: fixes f :: \'a \ 'b::topological_ab_group_add\ shows \has_sum (\x. - f x) A a \ has_sum f A (- a)\ by (auto simp add: sum_negf[abs_def] tendsto_minus_cancel_left has_sum_def) lemma summable_on_uminus: fixes f :: \'a \ 'b::topological_ab_group_add\ shows\(\x. - f x) summable_on A \ f summable_on A\ by (metis summable_on_def has_sum_uminus verit_minus_simplify(4)) lemma infsum_uminus: fixes f :: \'a \ 'b::{topological_ab_group_add, t2_space}\ shows \infsum (\x. - f x) A = - infsum f A\ by (metis (full_types) add.inverse_inverse add.inverse_neutral infsumI infsum_def has_sum_infsum has_sum_uminus) lemma has_sum_le_finite_sums: fixes a :: \'a::{comm_monoid_add,topological_space,linorder_topology}\ assumes \has_sum f A a\ assumes \\F. finite F \ F \ A \ sum f F \ b\ shows \a \ b\ proof - from assms(1) have 1: \(sum f \ a) (finite_subsets_at_top A)\ unfolding has_sum_def . from assms(2) have 2: \\\<^sub>F F in finite_subsets_at_top A. sum f F \ b\ by (rule_tac eventually_finite_subsets_at_top_weakI) show \a \ b\ using _ _ 1 2 apply (rule tendsto_le[where f=\\_. b\]) by auto qed lemma infsum_le_finite_sums: fixes b :: \'a::{comm_monoid_add,topological_space,linorder_topology}\ assumes \f summable_on A\ assumes \\F. finite F \ F \ A \ sum f F \ b\ shows \infsum f A \ b\ by (meson assms(1) assms(2) has_sum_infsum has_sum_le_finite_sums) lemma summable_on_scaleR_left [intro]: fixes c :: \'a :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. f x *\<^sub>R c) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>R c) = (\y. y *\<^sub>R c) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleR_left.additive_axioms) lemma summable_on_scaleR_right [intro]: fixes f :: \'a \ 'b :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. c *\<^sub>R f x) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. c *\<^sub>R f x) = (\y. c *\<^sub>R y) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleR_right.additive_axioms) lemma infsum_scaleR_left: fixes c :: \'a :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "infsum (\x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>R c) = (\y. y *\<^sub>R c) o f\], simp add: o_def) apply (rule infsum_comm_additive) using assms by (auto simp add: scaleR_left.additive_axioms) lemma infsum_scaleR_right: fixes f :: \'a \ 'b :: real_normed_vector\ shows "infsum (\x. c *\<^sub>R f x) A = c *\<^sub>R infsum f A" proof - consider (summable) \f summable_on A\ | (c0) \c = 0\ | (not_summable) \\ f summable_on A\ \c \ 0\ by auto then show ?thesis proof cases case summable then show ?thesis apply (subst asm_rl[of \(\x. c *\<^sub>R f x) = (\y. c *\<^sub>R y) o f\], simp add: o_def) apply (rule infsum_comm_additive) using summable by (auto simp add: scaleR_right.additive_axioms) next case c0 then show ?thesis by auto next case not_summable have \\ (\x. c *\<^sub>R f x) summable_on A\ proof (rule notI) assume \(\x. c *\<^sub>R f x) summable_on A\ then have \(\x. inverse c *\<^sub>R c *\<^sub>R f x) summable_on A\ using summable_on_scaleR_right by blast then have \f summable_on A\ using not_summable by auto with not_summable show False by simp qed then show ?thesis by (simp add: infsum_not_exists not_summable(1)) qed qed lemma infsum_Un_Int: fixes f :: "'a \ 'b::{topological_ab_group_add, t2_space}" assumes [simp]: "f summable_on A - B" "f summable_on B - A" \f summable_on A \ B\ shows "infsum f (A \ B) = infsum f A + infsum f B - infsum f (A \ B)" proof - have [simp]: \f summable_on A\ apply (subst asm_rl[of \A = (A-B) \ (A\B)\]) apply auto[1] apply (rule summable_on_Un_disjoint) by auto have \infsum f (A \ B) = infsum f A + infsum f (B - A)\ apply (subst infsum_Un_disjoint[symmetric]) by auto moreover have \infsum f (B - A \ A \ B) = infsum f (B - A) + infsum f (A \ B)\ by (rule infsum_Un_disjoint) auto moreover have "B - A \ A \ B = B" by blast ultimately show ?thesis by auto qed lemma inj_combinator': assumes "x \ F" shows \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B \ B x)\ proof - have "inj_on ((\(y, g). g(x := y)) \ prod.swap) (Pi\<^sub>E F B \ B x)" using inj_combinator[of x F B] assms by (intro comp_inj_on) (auto simp: product_swap) thus ?thesis by (simp add: o_def) qed lemma infsum_prod_PiE: \ \See also \infsum_prod_PiE_abs\ below with incomparable premises.\ fixes f :: "'a \ 'b \ 'c :: {comm_monoid_mult, topological_semigroup_mult, division_ring, banach}" assumes finite: "finite A" assumes "\x. x \ A \ f x summable_on B x" assumes "(\g. \x\A. f x (g x)) summable_on (PiE A B)" shows "infsum (\g. \x\A. f x (g x)) (PiE A B) = (\x\A. infsum (f x) (B x))" proof (use finite assms(2-) in induction) case empty then show ?case by auto next case (insert x F) have pi: \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ unfolding PiE_insert_eq by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) have prod: \(\x'\F. f x' ((p(x:=y)) x')) = (\x'\F. f x' (p x'))\ for p y by (rule prod.cong) (use insert.hyps in auto) have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B \ B x)\ using \x \ F\ by (rule inj_combinator') have summable1: \(\g. \x\insert x F. f x (g x)) summable_on Pi\<^sub>E (insert x F) B\ using insert.prems(2) . also have \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ by (simp only: pi) also have "(\g. \x\insert x F. f x (g x)) summable_on \ \ ((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) summable_on (Pi\<^sub>E F B \ B x)" using inj by (rule summable_on_reindex) also have "(\z\F. f z ((g(x := y)) z)) = (\z\F. f z (g z))" for g y using insert.hyps by (intro prod.cong) auto hence "((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) = (\(p, y). f x y * (\x'\F. f x' (p x')))" using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) finally have summable2: \(\(p, y). f x y * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B \ B x\ . then have \(\p. \\<^sub>\y\B x. f x y * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ by (rule summable_on_Sigma_banach) then have \(\p. (\\<^sub>\y\B x. f x y) * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ apply (subst infsum_cmult_left[symmetric]) using insert.prems(1) by blast then have summable3: \(\p. (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ if \(\\<^sub>\y\B x. f x y) \ 0\ apply (subst (asm) summable_on_cmult_right') using that by auto have \(\\<^sub>\g\Pi\<^sub>E (insert x F) B. \x\insert x F. f x (g x)) = (\\<^sub>\(p,y)\Pi\<^sub>E F B \ B x. \x'\insert x F. f x' ((p(x:=y)) x'))\ apply (subst pi) apply (subst infsum_reindex) using inj by (auto simp: o_def case_prod_unfold) also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' ((p(x:=y)) x')))\ apply (subst prod.insert) using insert by auto also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' (p x')))\ apply (subst prod) by rule also have \\ = (\\<^sub>\p\Pi\<^sub>E F B. \\<^sub>\y\B x. f x y * (\x'\F. f x' (p x')))\ apply (subst infsum_Sigma_banach[symmetric]) using summable2 apply blast by fastforce also have \\ = (\\<^sub>\y\B x. f x y) * (\\<^sub>\p\Pi\<^sub>E F B. \x'\F. f x' (p x'))\ apply (subst infsum_cmult_left') apply (subst infsum_cmult_right') by (rule refl) also have \\ = (\x\insert x F. infsum (f x) (B x))\ apply (subst prod.insert) using \finite F\ \x \ F\ apply auto[2] apply (cases \infsum (f x) (B x) = 0\) apply simp apply (subst insert.IH) apply (simp add: insert.prems(1)) apply (rule summable3) by auto finally show ?case by simp qed lemma infsum_prod_PiE_abs: \ \See also @{thm [source] infsum_prod_PiE} above with incomparable premises.\ fixes f :: "'a \ 'b \ 'c :: {banach, real_normed_div_algebra, comm_semiring_1}" assumes finite: "finite A" assumes "\x. x \ A \ f x abs_summable_on B x" shows "infsum (\g. \x\A. f x (g x)) (PiE A B) = (\x\A. infsum (f x) (B x))" proof (use finite assms(2) in induction) case empty then show ?case by auto next case (insert x F) have pi: \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ for x F and B :: "'a \ 'b set" unfolding PiE_insert_eq by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) have prod: \(\x'\F. f x' ((p(x:=y)) x')) = (\x'\F. f x' (p x'))\ for p y by (rule prod.cong) (use insert.hyps in auto) have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B \ B x)\ using \x \ F\ by (rule inj_combinator') define s where \s x = infsum (\y. norm (f x y)) (B x)\ for x have *: \(\p\P. norm (\x\F. f x (p x))) \ prod s F\ if P: \P \ Pi\<^sub>E F B\ and [simp]: \finite P\ \finite F\ and sum: \\x. x \ F \ f x abs_summable_on B x\ for P F proof - define B' where \B' x = {p x| p. p\P}\ for x have [simp]: \finite (B' x)\ for x using that by (auto simp: B'_def) have [simp]: \finite (Pi\<^sub>E F B')\ by (simp add: finite_PiE) have [simp]: \P \ Pi\<^sub>E F B'\ using that by (auto simp: B'_def) have B'B: \B' x \ B x\ if \x \ F\ for x unfolding B'_def using P that by auto have s_bound: \(\y\B' x. norm (f x y)) \ s x\ if \x \ F\ for x apply (simp_all add: s_def flip: infsum_finite) apply (rule infsum_mono_neutral) using that sum B'B by auto have \(\p\P. norm (\x\F. f x (p x))) \ (\p\Pi\<^sub>E F B'. norm (\x\F. f x (p x)))\ apply (rule sum_mono2) by auto also have \\ = (\p\Pi\<^sub>E F B'. \x\F. norm (f x (p x)))\ apply (subst prod_norm[symmetric]) by simp also have \\ = (\x\F. \y\B' x. norm (f x y))\ proof (use \finite F\ in induction) case empty then show ?case by simp next case (insert x F) have aux: \a = b \ c * a = c * b\ for a b c :: real by auto have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B' \ B' x)\ by (rule inj_combinator') (use insert.hyps in auto) have \(\p\Pi\<^sub>E (insert x F) B'. \x\insert x F. norm (f x (p x))) = (\(p,y)\Pi\<^sub>E F B' \ B' x. \x'\insert x F. norm (f x' ((p(x := y)) x')))\ apply (subst pi) apply (subst sum.reindex) using inj by (auto simp: case_prod_unfold) also have \\ = (\(p,y)\Pi\<^sub>E F B' \ B' x. norm (f x y) * (\x'\F. norm (f x' ((p(x := y)) x'))))\ apply (subst prod.insert) using insert.hyps by (auto simp: case_prod_unfold) also have \\ = (\(p, y)\Pi\<^sub>E F B' \ B' x. norm (f x y) * (\x'\F. norm (f x' (p x'))))\ apply (rule sum.cong) apply blast unfolding case_prod_unfold apply (rule aux) apply (rule prod.cong) using insert.hyps(2) by auto also have \\ = (\y\B' x. norm (f x y)) * (\p\Pi\<^sub>E F B'. \x'\F. norm (f x' (p x')))\ apply (subst sum_product) apply (subst sum.swap) apply (subst sum.cartesian_product) by simp also have \\ = (\y\B' x. norm (f x y)) * (\x\F. \y\B' x. norm (f x y))\ by (simp add: insert.IH) also have \\ = (\x\insert x F. \y\B' x. norm (f x y))\ using insert.hyps(1) insert.hyps(2) by force finally show ?case . qed also have \\ = (\x\F. \\<^sub>\y\B' x. norm (f x y))\ by auto also have \\ \ (\x\F. s x)\ apply (rule prod_mono) apply auto apply (simp add: sum_nonneg) using s_bound by presburger finally show ?thesis . qed have \(\g. \x\insert x F. f x (g x)) abs_summable_on Pi\<^sub>E (insert x F) B\ apply (rule nonneg_bdd_above_summable_on) apply (simp; fail) apply (rule bdd_aboveI[where M=\\x'\insert x F. s x'\]) using * insert.hyps insert.prems by blast also have \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ by (simp only: pi) also have "(\g. \x\insert x F. f x (g x)) abs_summable_on \ \ ((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) abs_summable_on (Pi\<^sub>E F B \ B x)" using inj by (subst summable_on_reindex) (auto simp: o_def) also have "(\z\F. f z ((g(x := y)) z)) = (\z\F. f z (g z))" for g y using insert.hyps by (intro prod.cong) auto hence "((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) = (\(p, y). f x y * (\x'\F. f x' (p x')))" using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) finally have summable2: \(\(p, y). f x y * (\x'\F. f x' (p x'))) abs_summable_on Pi\<^sub>E F B \ B x\ . have \(\\<^sub>\g\Pi\<^sub>E (insert x F) B. \x\insert x F. f x (g x)) = (\\<^sub>\(p,y)\Pi\<^sub>E F B \ B x. \x'\insert x F. f x' ((p(x:=y)) x'))\ apply (subst pi) apply (subst infsum_reindex) using inj by (auto simp: o_def case_prod_unfold) also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' ((p(x:=y)) x')))\ apply (subst prod.insert) using insert by auto also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' (p x')))\ apply (subst prod) by rule also have \\ = (\\<^sub>\p\Pi\<^sub>E F B. \\<^sub>\y\B x. f x y * (\x'\F. f x' (p x')))\ apply (subst infsum_Sigma_banach[symmetric]) using summable2 abs_summable_summable apply blast by fastforce also have \\ = (\\<^sub>\y\B x. f x y) * (\\<^sub>\p\Pi\<^sub>E F B. \x'\F. f x' (p x'))\ apply (subst infsum_cmult_left') apply (subst infsum_cmult_right') by (rule refl) also have \\ = (\x\insert x F. infsum (f x) (B x))\ apply (subst prod.insert) using \finite F\ \x \ F\ apply auto[2] apply (cases \infsum (f x) (B x) = 0\) apply (simp; fail) apply (subst insert.IH) apply (auto simp add: insert.prems(1)) done finally show ?case by simp qed subsection \Absolute convergence\ lemma abs_summable_countable: assumes \f abs_summable_on A\ shows \countable {x\A. f x \ 0}\ proof - have fin: \finite {x\A. norm (f x) \ t}\ if \t > 0\ for t proof (rule ccontr) assume *: \infinite {x \ A. t \ norm (f x)}\ have \infsum (\x. norm (f x)) A \ b\ for b proof - obtain b' where b': \of_nat b' \ b / t\ by (meson real_arch_simple) from * obtain F where cardF: \card F \ b'\ and \finite F\ and F: \F \ {x \ A. t \ norm (f x)}\ by (meson finite_if_finite_subsets_card_bdd nle_le) have \b \ of_nat b' * t\ using b' \t > 0\ by (simp add: field_simps split: if_splits) also have \\ \ of_nat (card F) * t\ by (simp add: cardF that) also have \\ = sum (\x. t) F\ by simp also have \\ \ sum (\x. norm (f x)) F\ by (metis (mono_tags, lifting) F in_mono mem_Collect_eq sum_mono) also have \\ = infsum (\x. norm (f x)) F\ using \finite F\ by (rule infsum_finite[symmetric]) also have \\ \ infsum (\x. norm (f x)) A\ by (rule infsum_mono_neutral) (use \finite F\ assms F in auto) finally show ?thesis . qed then show False by (meson gt_ex linorder_not_less) qed have \countable (\i\{1..}. {x\A. norm (f x) \ 1/of_nat i})\ by (rule countable_UN) (use fin in \auto intro!: countable_finite\) also have \\ = {x\A. f x \ 0}\ proof safe fix x assume x: "x \ A" "f x \ 0" define i where "i = max 1 (nat (ceiling (1 / norm (f x))))" have "i \ 1" by (simp add: i_def) moreover have "real i \ 1 / norm (f x)" unfolding i_def by linarith hence "1 / real i \ norm (f x)" using \f x \ 0\ by (auto simp: divide_simps mult_ac) ultimately show "x \ (\i\{1..}. {x \ A. 1 / real i \ norm (f x)})" using \x \ A\ by auto qed auto finally show ?thesis . qed (* Logically belongs in the section about reals, but needed as a dependency here *) lemma summable_on_iff_abs_summable_on_real: fixes f :: \'a \ real\ shows \f summable_on A \ f abs_summable_on A\ proof (rule iffI) assume \f summable_on A\ define n A\<^sub>p A\<^sub>n where \n x = norm (f x)\ and \A\<^sub>p = {x\A. f x \ 0}\ and \A\<^sub>n = {x\A. f x < 0}\ for x have [simp]: \A\<^sub>p \ A\<^sub>n = A\ \A\<^sub>p \ A\<^sub>n = {}\ by (auto simp: A\<^sub>p_def A\<^sub>n_def) from \f summable_on A\ have [simp]: \f summable_on A\<^sub>p\ \f summable_on A\<^sub>n\ using A\<^sub>p_def A\<^sub>n_def summable_on_subset_banach by fastforce+ then have [simp]: \n summable_on A\<^sub>p\ apply (subst summable_on_cong[where g=f]) by (simp_all add: A\<^sub>p_def n_def) moreover have [simp]: \n summable_on A\<^sub>n\ apply (subst summable_on_cong[where g=\\x. - f x\]) apply (simp add: A\<^sub>n_def n_def[abs_def]) by (simp add: summable_on_uminus) ultimately have [simp]: \n summable_on (A\<^sub>p \ A\<^sub>n)\ apply (rule summable_on_Un_disjoint) by simp then show \n summable_on A\ by simp next show \f abs_summable_on A \ f summable_on A\ using abs_summable_summable by blast qed lemma abs_summable_on_Sigma_iff: shows "f abs_summable_on Sigma A B \ (\x\A. (\y. f (x, y)) abs_summable_on B x) \ ((\x. infsum (\y. norm (f (x, y))) (B x)) abs_summable_on A)" proof (intro iffI conjI ballI) assume asm: \f abs_summable_on Sigma A B\ then have \(\x. infsum (\y. norm (f (x,y))) (B x)) summable_on A\ apply (rule_tac summable_on_Sigma_banach) by (auto simp: case_prod_unfold) then show \(\x. \\<^sub>\y\B x. norm (f (x, y))) abs_summable_on A\ using summable_on_iff_abs_summable_on_real by force show \(\y. f (x, y)) abs_summable_on B x\ if \x \ A\ for x proof - from asm have \f abs_summable_on Pair x ` B x\ apply (rule summable_on_subset_banach) using that by auto then show ?thesis apply (subst (asm) summable_on_reindex) by (auto simp: o_def inj_on_def) qed next assume asm: \(\x\A. (\xa. f (x, xa)) abs_summable_on B x) \ (\x. \\<^sub>\y\B x. norm (f (x, y))) abs_summable_on A\ have \(\xy\F. norm (f xy)) \ (\\<^sub>\x\A. \\<^sub>\y\B x. norm (f (x, y)))\ if \F \ Sigma A B\ and [simp]: \finite F\ for F proof - have [simp]: \(SIGMA x:fst ` F. {y. (x, y) \ F}) = F\ by (auto intro!: set_eqI simp add: Domain.DomainI fst_eq_Domain) have [simp]: \finite {y. (x, y) \ F}\ for x by (metis \finite F\ Range.intros finite_Range finite_subset mem_Collect_eq subsetI) have \(\xy\F. norm (f xy)) = (\x\fst ` F. \y\{y. (x,y)\F}. norm (f (x,y)))\ apply (subst sum.Sigma) by auto also have \\ = (\\<^sub>\x\fst ` F. \\<^sub>\y\{y. (x,y)\F}. norm (f (x,y)))\ apply (subst infsum_finite) by auto also have \\ \ (\\<^sub>\x\fst ` F. \\<^sub>\y\B x. norm (f (x,y)))\ apply (rule infsum_mono) apply (simp; fail) apply (simp; fail) apply (rule infsum_mono_neutral) using asm that(1) by auto also have \\ \ (\\<^sub>\x\A. \\<^sub>\y\B x. norm (f (x,y)))\ by (rule infsum_mono_neutral) (use asm that(1) in \auto simp add: infsum_nonneg\) finally show ?thesis . qed then show \f abs_summable_on Sigma A B\ by (intro nonneg_bdd_above_summable_on) (auto simp: bdd_above_def) qed lemma abs_summable_on_comparison_test: assumes "g abs_summable_on A" assumes "\x. x \ A \ norm (f x) \ norm (g x)" shows "f abs_summable_on A" proof (rule nonneg_bdd_above_summable_on) show "bdd_above (sum (\x. norm (f x)) ` {F. F \ A \ finite F})" proof (rule bdd_aboveI2) fix F assume F: "F \ {F. F \ A \ finite F}" have \sum (\x. norm (f x)) F \ sum (\x. norm (g x)) F\ using assms F by (intro sum_mono) auto also have \\ = infsum (\x. norm (g x)) F\ using F by simp also have \\ \ infsum (\x. norm (g x)) A\ proof (rule infsum_mono_neutral) show "g abs_summable_on F" by (rule summable_on_subset_banach[OF assms(1)]) (use F in auto) qed (use F assms in auto) finally show "(\x\F. norm (f x)) \ (\\<^sub>\x\A. norm (g x))" . qed qed auto lemma abs_summable_iff_bdd_above: fixes f :: \'a \ 'b::real_normed_vector\ shows \f abs_summable_on A \ bdd_above (sum (\x. norm (f x)) ` {F. F\A \ finite F})\ proof (rule iffI) assume \f abs_summable_on A\ show \bdd_above (sum (\x. norm (f x)) ` {F. F \ A \ finite F})\ proof (rule bdd_aboveI2) fix F assume F: "F \ {F. F \ A \ finite F}" show "(\x\F. norm (f x)) \ (\\<^sub>\x\A. norm (f x))" by (rule finite_sum_le_infsum) (use \f abs_summable_on A\ F in auto) qed next assume \bdd_above (sum (\x. norm (f x)) ` {F. F\A \ finite F})\ then show \f abs_summable_on A\ by (simp add: nonneg_bdd_above_summable_on) qed lemma abs_summable_product: fixes x :: "'a \ 'b::{real_normed_div_algebra,banach,second_countable_topology}" assumes x2_sum: "(\i. (x i) * (x i)) abs_summable_on A" and y2_sum: "(\i. (y i) * (y i)) abs_summable_on A" shows "(\i. x i * y i) abs_summable_on A" proof (rule nonneg_bdd_above_summable_on) show "bdd_above (sum (\xa. norm (x xa * y xa)) ` {F. F \ A \ finite F})" proof (rule bdd_aboveI2) fix F assume F: \F \ {F. F \ A \ finite F}\ then have r1: "finite F" and b4: "F \ A" by auto have a1: "(\\<^sub>\i\F. norm (x i * x i)) \ (\\<^sub>\i\A. norm (x i * x i))" apply (rule infsum_mono_neutral) using b4 r1 x2_sum by auto have "norm (x i * y i) \ norm (x i * x i) + norm (y i * y i)" for i unfolding norm_mult by (smt mult_left_mono mult_nonneg_nonneg mult_right_mono norm_ge_zero) hence "(\i\F. norm (x i * y i)) \ (\i\F. norm (x i * x i) + norm (y i * y i))" by (simp add: sum_mono) also have "\ = (\i\F. norm (x i * x i)) + (\i\F. norm (y i * y i))" by (simp add: sum.distrib) also have "\ = (\\<^sub>\i\F. norm (x i * x i)) + (\\<^sub>\i\F. norm (y i * y i))" by (simp add: \finite F\) also have "\ \ (\\<^sub>\i\A. norm (x i * x i)) + (\\<^sub>\i\A. norm (y i * y i))" using F assms by (intro add_mono infsum_mono2) auto finally show \(\xa\F. norm (x xa * y xa)) \ (\\<^sub>\i\A. norm (x i * x i)) + (\\<^sub>\i\A. norm (y i * y i))\ by simp qed qed auto subsection \Extended reals and nats\ lemma summable_on_ennreal[simp]: \(f::_ \ ennreal) summable_on S\ by (rule nonneg_summable_on_complete) simp lemma summable_on_enat[simp]: \(f::_ \ enat) summable_on S\ by (rule nonneg_summable_on_complete) simp lemma has_sum_superconst_infinite_ennreal: fixes f :: \'a \ ennreal\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "has_sum f S \" proof - have \(sum f \ \) (finite_subsets_at_top S)\ proof (rule order_tendstoI[rotated], simp) fix y :: ennreal assume \y < \\ then have \y / b < \\ by (metis b ennreal_divide_eq_top_iff gr_implies_not_zero infinity_ennreal_def top.not_eq_extremum) then obtain F where \finite F\ and \F \ S\ and cardF: \card F > y / b\ using \infinite S\ by (metis ennreal_Ex_less_of_nat infinite_arbitrarily_large infinity_ennreal_def) moreover have \sum f Y > y\ if \finite Y\ and \F \ Y\ and \Y \ S\ for Y proof - have \y < b * card F\ by (metis \y < \\ b cardF divide_less_ennreal ennreal_mult_eq_top_iff gr_implies_not_zero infinity_ennreal_def mult.commute top.not_eq_extremum) also have \\ \ b * card Y\ by (meson b card_mono less_imp_le mult_left_mono of_nat_le_iff that(1) that(2)) also have \\ = sum (\_. b) Y\ by (simp add: mult.commute) also have \\ \ sum f Y\ using geqb by (meson subset_eq sum_mono that(3)) finally show ?thesis . qed ultimately show \\\<^sub>F x in finite_subsets_at_top S. y < sum f x\ unfolding eventually_finite_subsets_at_top by auto qed then show ?thesis by (simp add: has_sum_def) qed lemma infsum_superconst_infinite_ennreal: fixes f :: \'a \ ennreal\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "infsum f S = \" using assms infsumI has_sum_superconst_infinite_ennreal by blast lemma infsum_superconst_infinite_ereal: fixes f :: \'a \ ereal\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "infsum f S = \" proof - obtain b' where b': \e2ennreal b' = b\ and \b' > 0\ using b by blast have "0 < e2ennreal b" using b' b by (metis dual_order.refl enn2ereal_e2ennreal gr_zeroI order_less_le zero_ennreal.abs_eq) hence *: \infsum (e2ennreal o f) S = \\ using assms b' by (intro infsum_superconst_infinite_ennreal[where b=b']) (auto intro!: e2ennreal_mono) have \infsum f S = infsum (enn2ereal o (e2ennreal o f)) S\ using geqb b by (intro infsum_cong) (fastforce simp: enn2ereal_e2ennreal) also have \\ = enn2ereal \\ apply (subst infsum_comm_additive_general) using * by (auto simp: continuous_at_enn2ereal) also have \\ = \\ by simp finally show ?thesis . qed lemma has_sum_superconst_infinite_ereal: fixes f :: \'a \ ereal\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "has_sum f S \" by (metis Infty_neq_0(1) assms infsum_def has_sum_infsum infsum_superconst_infinite_ereal) lemma infsum_superconst_infinite_enat: fixes f :: \'a \ enat\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "infsum f S = \" proof - have \ennreal_of_enat (infsum f S) = infsum (ennreal_of_enat o f) S\ apply (rule infsum_comm_additive_general[symmetric]) by auto also have \\ = \\ by (metis assms(3) b comp_apply ennreal_of_enat_0 ennreal_of_enat_inj ennreal_of_enat_le_iff geqb infsum_superconst_infinite_ennreal not_gr_zero) also have \\ = ennreal_of_enat \\ by simp finally show ?thesis by (rule ennreal_of_enat_inj[THEN iffD1]) qed lemma has_sum_superconst_infinite_enat: fixes f :: \'a \ enat\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "has_sum f S \" by (metis assms i0_lb has_sum_infsum infsum_superconst_infinite_enat nonneg_summable_on_complete) text \This lemma helps to relate a real-valued infsum to a supremum over extended nonnegative reals.\ lemma infsum_nonneg_is_SUPREMUM_ennreal: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ennreal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))" proof - have \ennreal (infsum f A) = infsum (ennreal o f) A\ apply (rule infsum_comm_additive_general[symmetric]) apply (subst sum_ennreal[symmetric]) using assms by auto also have \\ = (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))\ apply (subst nonneg_infsum_complete, simp) apply (rule SUP_cong, blast) apply (subst sum_ennreal[symmetric]) using fnn by auto finally show ?thesis . qed text \This lemma helps to related a real-valued infsum to a supremum over extended reals.\ lemma infsum_nonneg_is_SUPREMUM_ereal: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ereal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" proof - have \ereal (infsum f A) = infsum (ereal o f) A\ apply (rule infsum_comm_additive_general[symmetric]) using assms by auto also have \\ = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))\ by (subst nonneg_infsum_complete) (simp_all add: assms) finally show ?thesis . qed subsection \Real numbers\ text \Most lemmas in the general property section already apply to real numbers. A few ones that are specific to reals are given here.\ lemma infsum_nonneg_is_SUPREMUM_real: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "infsum f A = (SUP F\{F. finite F \ F \ A}. (sum f F))" proof - have "ereal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" using assms by (rule infsum_nonneg_is_SUPREMUM_ereal) also have "\ = ereal (SUP F\{F. finite F \ F \ A}. (sum f F))" proof (subst ereal_SUP) show "\SUP a\{F. finite F \ F \ A}. ereal (sum f a)\ \ \" using calculation by fastforce show "(SUP F\{F. finite F \ F \ A}. ereal (sum f F)) = (SUP a\{F. finite F \ F \ A}. ereal (sum f a))" by simp qed finally show ?thesis by simp qed lemma has_sum_nonneg_SUPREMUM_real: fixes f :: "'a \ real" assumes "f summable_on A" and "\x. x\A \ f x \ 0" shows "has_sum f A (SUP F\{F. finite F \ F \ A}. (sum f F))" by (metis (mono_tags, lifting) assms has_sum_infsum infsum_nonneg_is_SUPREMUM_real) lemma summable_countable_real: fixes f :: \'a \ real\ assumes \f summable_on A\ shows \countable {x\A. f x \ 0}\ using abs_summable_countable assms summable_on_iff_abs_summable_on_real by blast subsection \Complex numbers\ lemma has_sum_cnj_iff[simp]: fixes f :: \'a \ complex\ shows \has_sum (\x. cnj (f x)) M (cnj a) \ has_sum f M a\ by (simp add: has_sum_def lim_cnj del: cnj_sum add: cnj_sum[symmetric, abs_def, of f]) lemma summable_on_cnj_iff[simp]: "(\i. cnj (f i)) summable_on A \ f summable_on A" by (metis complex_cnj_cnj summable_on_def has_sum_cnj_iff) lemma infsum_cnj[simp]: \infsum (\x. cnj (f x)) M = cnj (infsum f M)\ by (metis complex_cnj_zero infsumI has_sum_cnj_iff infsum_def summable_on_cnj_iff has_sum_infsum) lemma infsum_Re: assumes "f summable_on M" shows "infsum (\x. Re (f x)) M = Re (infsum f M)" apply (rule infsum_comm_additive[where f=Re, unfolded o_def]) using assms by (auto intro!: additive.intro) lemma has_sum_Re: assumes "has_sum f M a" shows "has_sum (\x. Re (f x)) M (Re a)" apply (rule has_sum_comm_additive[where f=Re, unfolded o_def]) using assms by (auto intro!: additive.intro tendsto_Re) lemma summable_on_Re: assumes "f summable_on M" shows "(\x. Re (f x)) summable_on M" apply (rule summable_on_comm_additive[where f=Re, unfolded o_def]) using assms by (auto intro!: additive.intro) lemma infsum_Im: assumes "f summable_on M" shows "infsum (\x. Im (f x)) M = Im (infsum f M)" apply (rule infsum_comm_additive[where f=Im, unfolded o_def]) using assms by (auto intro!: additive.intro) lemma has_sum_Im: assumes "has_sum f M a" shows "has_sum (\x. Im (f x)) M (Im a)" apply (rule has_sum_comm_additive[where f=Im, unfolded o_def]) using assms by (auto intro!: additive.intro tendsto_Im) lemma summable_on_Im: assumes "f summable_on M" shows "(\x. Im (f x)) summable_on M" apply (rule summable_on_comm_additive[where f=Im, unfolded o_def]) using assms by (auto intro!: additive.intro) lemma nonneg_infsum_le_0D_complex: fixes f :: "'a \ complex" assumes "infsum f A \ 0" and abs_sum: "f summable_on A" and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" proof - have \Im (f x) = 0\ apply (rule nonneg_infsum_le_0D[where A=A]) using assms by (auto simp add: infsum_Im summable_on_Im less_eq_complex_def) moreover have \Re (f x) = 0\ apply (rule nonneg_infsum_le_0D[where A=A]) using assms by (auto simp add: summable_on_Re infsum_Re less_eq_complex_def) ultimately show ?thesis by (simp add: complex_eqI) qed lemma nonneg_has_sum_le_0D_complex: fixes f :: "'a \ complex" assumes "has_sum f A a" and \a \ 0\ and "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" by (metis assms infsumI nonneg_infsum_le_0D_complex summable_on_def) text \The lemma @{thm [source] infsum_mono_neutral} above applies to various linear ordered monoids such as the reals but not to the complex numbers. Thus we have a separate corollary for those:\ lemma infsum_mono_neutral_complex: fixes f :: "'a \ complex" assumes [simp]: "f summable_on A" and [simp]: "g summable_on B" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows \infsum f A \ infsum g B\ proof - have \infsum (\x. Re (f x)) A \ infsum (\x. Re (g x)) B\ apply (rule infsum_mono_neutral) using assms(3-5) by (auto simp add: summable_on_Re less_eq_complex_def) then have Re: \Re (infsum f A) \ Re (infsum g B)\ by (metis assms(1-2) infsum_Re) have \infsum (\x. Im (f x)) A = infsum (\x. Im (g x)) B\ apply (rule infsum_cong_neutral) using assms(3-5) by (auto simp add: summable_on_Re less_eq_complex_def) then have Im: \Im (infsum f A) = Im (infsum g B)\ by (metis assms(1-2) infsum_Im) from Re Im show ?thesis by (auto simp: less_eq_complex_def) qed lemma infsum_mono_complex: \ \For \<^typ>\real\, @{thm [source] infsum_mono} can be used. But \<^typ>\complex\ does not have the right typeclass.\ fixes f g :: "'a \ complex" assumes f_sum: "f summable_on A" and g_sum: "g summable_on A" assumes leq: "\x. x \ A \ f x \ g x" shows "infsum f A \ infsum g A" by (metis DiffE IntD1 f_sum g_sum infsum_mono_neutral_complex leq) lemma infsum_nonneg_complex: fixes f :: "'a \ complex" assumes "f summable_on M" and "\x. x \ M \ 0 \ f x" shows "infsum f M \ 0" (is "?lhs \ _") by (metis assms(1) assms(2) infsum_0_simp summable_on_0_simp infsum_mono_complex) lemma infsum_cmod: assumes "f summable_on M" and fnn: "\x. x \ M \ 0 \ f x" shows "infsum (\x. cmod (f x)) M = cmod (infsum f M)" proof - have \complex_of_real (infsum (\x. cmod (f x)) M) = infsum (\x. complex_of_real (cmod (f x))) M\ proof (rule infsum_comm_additive[symmetric, unfolded o_def]) have "(\z. Re (f z)) summable_on M" using assms summable_on_Re by blast also have "?this \ f abs_summable_on M" using fnn by (intro summable_on_cong) (auto simp: less_eq_complex_def cmod_def) finally show \ . qed (auto simp: additive_def) also have \\ = infsum f M\ apply (rule infsum_cong) using fnn cmod_eq_Re complex_is_Real_iff less_eq_complex_def by force finally show ?thesis by (metis abs_of_nonneg infsum_def le_less_trans norm_ge_zero norm_infsum_bound norm_of_real not_le order_refl) qed lemma summable_on_iff_abs_summable_on_complex: fixes f :: \'a \ complex\ shows \f summable_on A \ f abs_summable_on A\ proof (rule iffI) assume \f summable_on A\ define i r ni nr n where \i x = Im (f x)\ and \r x = Re (f x)\ and \ni x = norm (i x)\ and \nr x = norm (r x)\ and \n x = norm (f x)\ for x from \f summable_on A\ have \i summable_on A\ by (simp add: i_def[abs_def] summable_on_Im) then have [simp]: \ni summable_on A\ using ni_def[abs_def] summable_on_iff_abs_summable_on_real by force from \f summable_on A\ have \r summable_on A\ by (simp add: r_def[abs_def] summable_on_Re) then have [simp]: \nr summable_on A\ by (metis nr_def summable_on_cong summable_on_iff_abs_summable_on_real) have n_sum: \n x \ nr x + ni x\ for x by (simp add: n_def nr_def ni_def r_def i_def cmod_le) have *: \(\x. nr x + ni x) summable_on A\ apply (rule summable_on_add) by auto show \n summable_on A\ apply (rule nonneg_bdd_above_summable_on) apply (simp add: n_def; fail) apply (rule bdd_aboveI[where M=\infsum (\x. nr x + ni x) A\]) using * n_sum by (auto simp flip: infsum_finite simp: ni_def[abs_def] nr_def[abs_def] intro!: infsum_mono_neutral) next show \f abs_summable_on A \ f summable_on A\ using abs_summable_summable by blast qed lemma summable_countable_complex: fixes f :: \'a \ complex\ assumes \f summable_on A\ shows \countable {x\A. f x \ 0}\ using abs_summable_countable assms summable_on_iff_abs_summable_on_complex by blast end diff --git a/src/HOL/Analysis/Interval_Integral.thy b/src/HOL/Analysis/Interval_Integral.thy --- a/src/HOL/Analysis/Interval_Integral.thy +++ b/src/HOL/Analysis/Interval_Integral.thy @@ -1,1129 +1,1129 @@ (* Title: HOL/Analysis/Interval_Integral.thy Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU) Lebesgue integral over an interval (with endpoints possibly +-\) *) theory Interval_Integral (*FIX ME rename? Lebesgue *) imports Equivalence_Lebesgue_Henstock_Integration begin definition "einterval a b = {x. a < ereal x \ ereal x < b}" lemma einterval_eq[simp]: shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}" and einterval_eq_Ici: "einterval (ereal a) \ = {a <..}" and einterval_eq_Iic: "einterval (- \) (ereal b) = {..< b}" and einterval_eq_UNIV: "einterval (- \) \ = UNIV" by (auto simp: einterval_def) lemma einterval_same: "einterval a a = {}" by (auto simp: einterval_def) lemma einterval_iff: "x \ einterval a b \ a < ereal x \ ereal x < b" by (simp add: einterval_def) lemma einterval_nonempty: "a < b \ \c. c \ einterval a b" by (cases a b rule: ereal2_cases, auto simp: einterval_def intro!: dense gt_ex lt_ex) lemma open_einterval[simp]: "open (einterval a b)" by (cases a b rule: ereal2_cases) (auto simp: einterval_def intro!: open_Collect_conj open_Collect_less continuous_intros) lemma borel_einterval[measurable]: "einterval a b \ sets borel" unfolding einterval_def by measurable subsection \Approximating a (possibly infinite) interval\ lemma filterlim_sup1: "(LIM x F. f x :> G1) \ (LIM x F. f x :> (sup G1 G2))" unfolding filterlim_def by (auto intro: le_supI1) lemma ereal_incseq_approx: fixes a b :: ereal assumes "a < b" obtains X :: "nat \ real" where "incseq X" "\i. a < X i" "\i. X i < b" "X \ b" proof (cases b) case PInf with \a < b\ have "a = -\ \ (\r. a = ereal r)" by (cases a) auto moreover have "(\x. ereal (real (Suc x))) \ \" by (simp add: Lim_PInfty filterlim_sequentially_Suc) (metis le_SucI of_nat_Suc of_nat_mono order_trans real_arch_simple) moreover have "\r. (\x. ereal (r + real (Suc x))) \ \" by (simp add: filterlim_sequentially_Suc Lim_PInfty) (metis add.commute diff_le_eq nat_ceiling_le_eq) ultimately show thesis by (intro that[of "\i. real_of_ereal a + Suc i"]) (auto simp: incseq_def PInf) next case (real b') define d where "d = b' - (if a = -\ then b' - 1 else real_of_ereal a)" with \a < b\ have a': "0 < d" by (cases a) (auto simp: real) moreover have "\i r. r < b' \ (b' - r) * 1 < (b' - r) * real (Suc (Suc i))" by (intro mult_strict_left_mono) auto with \a < b\ a' have "\i. a < ereal (b' - d / real (Suc (Suc i)))" by (cases a) (auto simp: real d_def field_simps) moreover have "(\i. b' - d / real i) \ b'" by (force intro: tendsto_eq_intros tendsto_divide_0[OF tendsto_const] filterlim_sup1 simp: at_infinity_eq_at_top_bot filterlim_real_sequentially) then have "(\i. b' - d / Suc (Suc i)) \ b'" by (blast intro: dest: filterlim_sequentially_Suc [THEN iffD2]) ultimately show thesis by (intro that[of "\i. b' - d / Suc (Suc i)"]) (auto simp: real incseq_def intro!: divide_left_mono) qed (use \a < b\ in auto) lemma ereal_decseq_approx: fixes a b :: ereal assumes "a < b" obtains X :: "nat \ real" where "decseq X" "\i. a < X i" "\i. X i < b" "X \ a" proof - have "-b < -a" using \a < b\ by simp from ereal_incseq_approx[OF this] obtain X where "incseq X" "\i. - b < ereal (X i)" "\i. ereal (X i) < - a" "(\x. ereal (X x)) \ - a" by auto then show thesis apply (intro that[of "\i. - X i"]) apply (auto simp: decseq_def incseq_def simp flip: uminus_ereal.simps) apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+ done qed proposition einterval_Icc_approximation: fixes a b :: ereal assumes "a < b" obtains u l :: "nat \ real" where "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" proof - from dense[OF \a < b\] obtain c where "a < c" "c < b" by safe from ereal_incseq_approx[OF \c < b\] obtain u where u: "incseq u" "\i. c < ereal (u i)" "\i. ereal (u i) < b" "(\x. ereal (u x)) \ b" by auto from ereal_decseq_approx[OF \a < c\] obtain l where l: "decseq l" "\i. a < ereal (l i)" "\i. ereal (l i) < c" "(\x. ereal (l x)) \ a" by auto { fix i from less_trans[OF \l i < c\ \c < u i\] have "l i < u i" by simp } have "einterval a b = (\i. {l i .. u i})" proof (auto simp: einterval_iff) fix x assume "a < ereal x" "ereal x < b" have "eventually (\i. ereal (l i) < ereal x) sequentially" using l(4) \a < ereal x\ by (rule order_tendstoD) moreover have "eventually (\i. ereal x < ereal (u i)) sequentially" using u(4) \ereal x< b\ by (rule order_tendstoD) ultimately have "eventually (\i. l i < x \ x < u i) sequentially" by eventually_elim auto then show "\i. l i \ x \ x \ u i" by (auto intro: less_imp_le simp: eventually_sequentially) next fix x i assume "l i \ x" "x \ u i" with \a < ereal (l i)\ \ereal (u i) < b\ show "a < ereal x" "ereal x < b" by (auto simp flip: ereal_less_eq(3)) qed show thesis by (intro that) fact+ qed (* TODO: in this definition, it would be more natural if einterval a b included a and b when they are real. *) definition\<^marker>\tag important\ interval_lebesgue_integral :: "real measure \ ereal \ ereal \ (real \ 'a) \ 'a::{banach, second_countable_topology}" where "interval_lebesgue_integral M a b f = (if a \ b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))" syntax "_ascii_interval_lebesgue_integral" :: "pttrn \ real \ real \ real measure \ real \ real" ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60) translations "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\x. f)" definition\<^marker>\tag important\ interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ 'a::{banach, second_countable_topology}) \ bool" where "interval_lebesgue_integrable M a b f = (if a \ b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)" syntax "_ascii_interval_lebesgue_borel_integral" :: "pttrn \ real \ real \ real \ real" ("(4LBINT _=_.._. _)" [0,60,60,61] 60) translations "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\x. f)" subsection\Basic properties of integration over an interval\ lemma interval_lebesgue_integral_cong: "a \ b \ (\x. x \ einterval a b \ f x = g x) \ einterval a b \ sets M \ interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" by (auto intro: set_lebesgue_integral_cong simp: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ a \ b \ AE x \ einterval a b in M. f x = g x \ einterval a b \ sets M \ interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" by (auto intro: set_lebesgue_integral_cong_AE simp: interval_lebesgue_integral_def) lemma interval_integrable_mirror: shows "interval_lebesgue_integrable lborel a b (\x. f (-x)) \ interval_lebesgue_integrable lborel (-b) (-a) f" proof - have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)" for a b :: ereal and x :: real by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator) show ?thesis unfolding interval_lebesgue_integrable_def using lborel_integrable_real_affine_iff[symmetric, of "-1" "\x. indicator (einterval _ _) x *\<^sub>R f x" 0] by (simp add: * set_integrable_def) qed lemma interval_lebesgue_integral_add [intro, simp]: fixes M a b f assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g" shows "interval_lebesgue_integrable M a b (\x. f x + g x)" and "interval_lebesgue_integral M a b (\x. f x + g x) = interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g" using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def field_simps) lemma interval_lebesgue_integral_diff [intro, simp]: fixes M a b f assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g" shows "interval_lebesgue_integrable M a b (\x. f x - g x)" and "interval_lebesgue_integral M a b (\x. f x - g x) = interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g" using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def field_simps) lemma interval_lebesgue_integrable_mult_right [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. c * f x)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integrable_mult_left [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. f x * c)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integrable_divide [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. f x / c)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integral_mult_right [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. c * f x) = c * interval_lebesgue_integral M a b f" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_mult_left [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. f x * c) = interval_lebesgue_integral M a b f * c" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_divide [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. f x / c) = interval_lebesgue_integral M a b f / c" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_uminus: "interval_lebesgue_integral M a b (\x. - f x) = - interval_lebesgue_integral M a b f" by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def) lemma interval_lebesgue_integral_of_real: "interval_lebesgue_integral M a b (\x. complex_of_real (f x)) = of_real (interval_lebesgue_integral M a b f)" unfolding interval_lebesgue_integral_def by (auto simp: interval_lebesgue_integral_def set_integral_complex_of_real) lemma interval_lebesgue_integral_le_eq: fixes a b f assumes "a \ b" shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)" using assms by (auto simp: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_gt_eq: fixes a b f assumes "a > b" shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)" using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def) lemma interval_lebesgue_integral_gt_eq': fixes a b f assumes "a > b" shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f" using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def) lemma interval_integral_endpoints_same [simp]: "(LBINT x=a..a. f x) = 0" by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same) lemma interval_integral_endpoints_reverse: "(LBINT x=a..b. f x) = -(LBINT x=b..a. f x)" by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same) lemma interval_integrable_endpoints_reverse: "interval_lebesgue_integrable lborel a b f \ interval_lebesgue_integrable lborel b a f" by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same) lemma interval_integral_reflect: "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))" proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (auto simp: interval_lebesgue_integral_def interval_integrable_endpoints_reverse split: if_split_asm) next case (le a b) have "LBINT x:{x. - x \ einterval a b}. f (- x) = LBINT x:einterval (- b) (- a). f (- x)" unfolding interval_lebesgue_integrable_def set_lebesgue_integral_def apply (rule Bochner_Integration.integral_cong [OF refl]) by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less simp flip: uminus_ereal.simps split: split_indicator) then show ?case unfolding interval_lebesgue_integral_def by (subst set_integral_reflect) (simp add: le) qed lemma interval_lebesgue_integral_0_infty: "interval_lebesgue_integrable M 0 \ f \ set_integrable M {0<..} f" "interval_lebesgue_integral M 0 \ f = (LINT x:{0<..}|M. f x)" unfolding zero_ereal_def by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def) lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\ | M. f x) = (LINT x : {a<..} | M. f x)" unfolding interval_lebesgue_integral_def by auto proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \ f) = (set_integrable M {a<..} f)" unfolding interval_lebesgue_integrable_def by auto subsection\Basic properties of integration over an interval wrt lebesgue measure\ lemma interval_integral_zero [simp]: fixes a b :: ereal shows "LBINT x=a..b. 0 = 0" unfolding interval_lebesgue_integral_def set_lebesgue_integral_def einterval_eq by simp lemma interval_integral_const [intro, simp]: fixes a b c :: real shows "interval_lebesgue_integrable lborel a b (\x. c)" and "LBINT x=a..b. c = c * (b - a)" unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq by (auto simp: less_imp_le field_simps measure_def set_integrable_def set_lebesgue_integral_def) lemma interval_integral_cong_AE: assumes [measurable]: "f \ borel_measurable borel" "g \ borel_measurable borel" assumes "AE x \ einterval (min a b) (max a b) in lborel. f x = g x" shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g" using assms proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b]) next case (le a b) then show ?case by (auto simp: interval_lebesgue_integral_def max_def min_def intro!: set_lebesgue_integral_cong_AE) qed lemma interval_integral_cong: assumes "\x. x \ einterval (min a b) (max a b) \ f x = g x" shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g" using assms proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b]) next case (le a b) then show ?case by (auto simp: interval_lebesgue_integral_def max_def min_def intro!: set_lebesgue_integral_cong) qed lemma interval_lebesgue_integrable_cong_AE: "f \ borel_measurable lborel \ g \ borel_measurable lborel \ AE x \ einterval (min a b) (max a b) in lborel. f x = g x \ interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g" apply (simp add: interval_lebesgue_integrable_def) apply (intro conjI impI set_integrable_cong_AE) apply (auto simp: min_def max_def) done lemma interval_integrable_abs_iff: fixes f :: "real \ real" shows "f \ borel_measurable lborel \ interval_lebesgue_integrable lborel a b (\x. \f x\) = interval_lebesgue_integrable lborel a b f" unfolding interval_lebesgue_integrable_def by (subst (1 2) set_integrable_abs_iff') simp_all lemma interval_integral_Icc: fixes a b :: real shows "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a..b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{a, b}"] simp add: interval_lebesgue_integral_def) lemma interval_integral_Icc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a \ ereal x \ ereal x \ b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ioc: "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a<..b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{a, b}"] simp add: interval_lebesgue_integral_def einterval_iff) (* TODO: other versions as well? *) (* Yes: I need the Icc' version. *) lemma interval_integral_Ioc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \ ereal x \ b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ico: "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a..a\ < \ \ (LBINT x=a..\. f x) = (LBINT x : {real_of_ereal a <..}. f x)" by (auto simp: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ioo: "a \ b \ \a\ < \ ==> \b\ < \ \ (LBINT x=a..b. f x) = (LBINT x : {real_of_ereal a <..< real_of_ereal b}. f x)" by (auto simp: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_discrete_difference: fixes f :: "real \ 'b::{banach, second_countable_topology}" and a b :: ereal assumes "countable X" and eq: "\x. a \ b \ a < x \ x < b \ x \ X \ f x = g x" and anti_eq: "\x. b \ a \ b < x \ x < a \ x \ X \ f x = g x" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" unfolding interval_lebesgue_integral_def set_lebesgue_integral_def apply (intro if_cong refl arg_cong[where f="\x. - x"] integral_discrete_difference[of X] assms) apply (auto simp: eq anti_eq einterval_iff split: split_indicator) done lemma interval_integral_sum: fixes a b c :: ereal assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f" shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)" proof - let ?I = "\a b. LBINT x=a..b. f x" { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \ b" "b \ c" then have ord: "a \ b" "b \ c" "a \ c" and f': "set_integrable lborel (einterval a c) f" by (auto simp: interval_lebesgue_integrable_def) then have f: "set_borel_measurable borel (einterval a c) f" unfolding set_integrable_def set_borel_measurable_def by (drule_tac borel_measurable_integrable) simp have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \ einterval b c. f x)" proof (rule set_integral_cong_set) show "AE x in lborel. (x \ einterval a b \ einterval b c) = (x \ einterval a c)" using AE_lborel_singleton[of "real_of_ereal b"] ord by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff) show "set_borel_measurable lborel (einterval a c) f" "set_borel_measurable lborel (einterval a b \ einterval b c) f" unfolding set_borel_measurable_def using ord by (auto simp: einterval_iff intro!: set_borel_measurable_subset[OF f, unfolded set_borel_measurable_def]) qed also have "\ = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)" using ord by (intro set_integral_Un_AE) (auto intro!: set_integrable_subset[OF f'] simp: einterval_iff not_less) finally have "?I a b + ?I b c = ?I a c" using ord by (simp add: interval_lebesgue_integral_def) } note 1 = this { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \ b" "b \ c" from 1[OF this] have "?I b c + ?I a b = ?I a c" by (metis add.commute) } note 2 = this have 3: "\a b. b \ a \ (LBINT x=a..b. f x) = - (LBINT x=b..a. f x)" by (rule interval_integral_endpoints_reverse) show ?thesis using integrable apply (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases]) apply simp_all (* remove some looping cases *) by (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3) qed lemma interval_integrable_isCont: fixes a b and f :: "real \ 'a::{banach, second_countable_topology}" shows "(\x. min a b \ x \ x \ max a b \ isCont f x) \ interval_lebesgue_integrable lborel a b f" proof (induct a b rule: linorder_wlog) case (le a b) then show ?case unfolding interval_lebesgue_integrable_def set_integrable_def by (auto simp: interval_lebesgue_integrable_def intro!: set_integrable_subset[unfolded set_integrable_def, OF borel_integrable_compact[of "{a .. b}"]] continuous_at_imp_continuous_on) qed (auto intro: interval_integrable_endpoints_reverse[THEN iffD1]) lemma interval_integrable_continuous_on: fixes a b :: real and f assumes "a \ b" and "continuous_on {a..b} f" shows "interval_lebesgue_integrable lborel a b f" using assms unfolding interval_lebesgue_integrable_def apply simp by (rule set_integrable_subset, rule borel_integrable_atLeastAtMost' [of a b], auto) lemma interval_integral_eq_integral: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ set_integrable lborel {a..b} f \ LBINT x=a..b. f x = integral {a..b} f" by (subst interval_integral_Icc, simp) (rule set_borel_integral_eq_integral) lemma interval_integral_eq_integral': fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ set_integrable lborel (einterval a b) f \ LBINT x=a..b. f x = integral (einterval a b) f" by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral) subsection\General limit approximation arguments\ proposition interval_integral_Icc_approx_nonneg: fixes a b :: ereal assumes "a < b" fixes u l :: "nat \ real" assumes approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" fixes f :: "real \ real" assumes f_integrable: "\i. set_integrable lborel {l i..u i} f" assumes f_nonneg: "AE x in lborel. a < ereal x \ ereal x < b \ 0 \ f x" assumes f_measurable: "set_borel_measurable lborel (einterval a b) f" assumes lbint_lim: "(\i. LBINT x=l i.. u i. f x) \ C" shows "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = C" proof - have 1 [unfolded set_integrable_def]: "\i. set_integrable lborel {l i..u i} f" by (rule f_integrable) have 2: "AE x in lborel. mono (\n. indicator {l n..u n} x *\<^sub>R f x)" proof - from f_nonneg have "AE x in lborel. \i. l i \ x \ x \ u i \ 0 \ f x" by eventually_elim (metis approx(5) approx(6) dual_order.strict_trans1 ereal_less_eq(3) le_less_trans) then show ?thesis apply eventually_elim apply (auto simp: mono_def split: split_indicator) apply (metis approx(3) decseqD order_trans) apply (metis approx(2) incseqD order_trans) done qed have 3: "AE x in lborel. (\i. indicator {l i..u i} x *\<^sub>R f x) \ indicator (einterval a b) x *\<^sub>R f x" proof - { fix x i assume "l i \ x" "x \ u i" then have "eventually (\i. l i \ x \ x \ u i) sequentially" apply (auto simp: eventually_sequentially intro!: exI[of _ i]) apply (metis approx(3) decseqD order_trans) apply (metis approx(2) incseqD order_trans) done then have "eventually (\i. f x * indicator {l i..u i} x = f x) sequentially" by eventually_elim auto } then show ?thesis unfolding approx(1) by (auto intro!: AE_I2 tendsto_eventually split: split_indicator) qed have 4: "(\i. \ x. indicator {l i..u i} x *\<^sub>R f x \lborel) \ C" using lbint_lim by (simp add: interval_integral_Icc [unfolded set_lebesgue_integral_def] approx less_imp_le) have 5: "(\x. indicat_real (einterval a b) x *\<^sub>R f x) \ borel_measurable lborel" using f_measurable set_borel_measurable_def by blast have "(LBINT x=a..b. f x) = lebesgue_integral lborel (\x. indicator (einterval a b) x *\<^sub>R f x)" using assms by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def less_imp_le) also have "\ = C" by (rule integral_monotone_convergence [OF 1 2 3 4 5]) finally show "(LBINT x=a..b. f x) = C" . show "set_integrable lborel (einterval a b) f" unfolding set_integrable_def by (rule integrable_monotone_convergence[OF 1 2 3 4 5]) qed proposition interval_integral_Icc_approx_integrable: fixes u l :: "nat \ real" and a b :: ereal fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes "a < b" assumes approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" assumes f_integrable: "set_integrable lborel (einterval a b) f" shows "(\i. LBINT x=l i.. u i. f x) \ (LBINT x=a..b. f x)" proof - have "(\i. LBINT x:{l i.. u i}. f x) \ (LBINT x:einterval a b. f x)" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable lborel (\x. norm (indicator (einterval a b) x *\<^sub>R f x))" using f_integrable integrable_norm set_integrable_def by blast show "(\x. indicat_real (einterval a b) x *\<^sub>R f x) \ borel_measurable lborel" using f_integrable by (simp add: set_integrable_def) then show "\i. (\x. indicat_real {l i..u i} x *\<^sub>R f x) \ borel_measurable lborel" by (rule set_borel_measurable_subset [unfolded set_borel_measurable_def]) (auto simp: approx) show "\i. AE x in lborel. norm (indicator {l i..u i} x *\<^sub>R f x) \ norm (indicator (einterval a b) x *\<^sub>R f x)" by (intro AE_I2) (auto simp: approx split: split_indicator) show "AE x in lborel. (\i. indicator {l i..u i} x *\<^sub>R f x) \ indicator (einterval a b) x *\<^sub>R f x" proof (intro AE_I2 tendsto_intros tendsto_eventually) fix x { fix i assume "l i \ x" "x \ u i" with \incseq u\[THEN incseqD, of i] \decseq l\[THEN decseqD, of i] have "eventually (\i. l i \ x \ x \ u i) sequentially" by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) } then show "eventually (\xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially" using approx order_tendstoD(2)[OF \l \ a\, of x] order_tendstoD(1)[OF \u \ b\, of x] by (auto split: split_indicator) qed qed with \a < b\ \\i. l i < u i\ show ?thesis by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le) qed subsection\A slightly stronger Fundamental Theorem of Calculus\ text\Three versions: first, for finite intervals, and then two versions for arbitrary intervals.\ (* TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.) *) lemma interval_integral_FTC_finite: fixes f F :: "real \ 'a::euclidean_space" and a b :: real assumes f: "continuous_on {min a b..max a b} f" assumes F: "\x. min a b \ x \ x \ max a b \ (F has_vector_derivative (f x)) (at x within {min a b..max a b})" shows "(LBINT x=a..b. f x) = F b - F a" proof (cases "a \ b") case True have "(LBINT x=a..b. f x) = (LBINT x. indicat_real {a..b} x *\<^sub>R f x)" by (simp add: True interval_integral_Icc set_lebesgue_integral_def) also have "\ = F b - F a" proof (rule integral_FTC_atLeastAtMost [OF True]) show "continuous_on {a..b} f" using True f by linarith show "\x. \a \ x; x \ b\ \ (F has_vector_derivative f x) (at x within {a..b})" by (metis F True max.commute max_absorb1 min_def) qed finally show ?thesis . next case False then have "b \ a" by simp have "- interval_lebesgue_integral lborel (ereal b) (ereal a) f = - (LBINT x. indicat_real {b..a} x *\<^sub>R f x)" by (simp add: \b \ a\ interval_integral_Icc set_lebesgue_integral_def) also have "\ = F b - F a" proof (subst integral_FTC_atLeastAtMost [OF \b \ a\]) show "continuous_on {b..a} f" using False f by linarith show "\x. \b \ x; x \ a\ \ (F has_vector_derivative f x) (at x within {b..a})" by (metis F False max_def min_def) qed auto finally show ?thesis by (metis interval_integral_endpoints_reverse) qed lemma interval_integral_FTC_nonneg: fixes f F :: "real \ real" and a b :: ereal assumes "a < b" assumes F: "\x. a < ereal x \ ereal x < b \ DERIV F x :> f x" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_nonneg: "AE x in lborel. a < ereal x \ ereal x < b \ 0 \ f x" assumes A: "((F \ real_of_ereal) \ A) (at_right a)" assumes B: "((F \ real_of_ereal) \ B) (at_left b)" shows "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = B - A" proof - obtain u l where approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have FTCi: "\i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)" using assms approx apply (intro interval_integral_FTC_finite) apply (auto simp: less_imp_le min_def max_def - has_field_derivative_iff_has_vector_derivative[symmetric]) + has_real_derivative_iff_has_vector_derivative[symmetric]) apply (rule continuous_at_imp_continuous_on, auto intro!: f) by (rule DERIV_subset [OF F], auto) have 1: "\i. set_integrable lborel {l i..u i} f" proof - fix i show "set_integrable lborel {l i .. u i} f" using \a < l i\ \u i < b\ unfolding set_integrable_def by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI) (auto simp flip: ereal_less_eq) qed have 2: "set_borel_measurable lborel (einterval a b) f" unfolding set_borel_measurable_def by (auto simp del: real_scaleR_def intro!: borel_measurable_continuous_on_indicator simp: continuous_on_eq_continuous_at einterval_iff f) have 3: "(\i. LBINT x=l i..u i. f x) \ B - A" apply (subst FTCi) apply (intro tendsto_intros) using B approx unfolding tendsto_at_iff_sequentially comp_def using tendsto_at_iff_sequentially[where 'a=real] apply (elim allE[of _ "\i. ereal (u i)"], auto) using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) show "(LBINT x=a..b. f x) = B - A" by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) show "set_integrable lborel (einterval a b) f" by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) qed theorem interval_integral_FTC_integrable: fixes f F :: "real \ 'a::euclidean_space" and a b :: ereal assumes "a < b" assumes F: "\x. a < ereal x \ ereal x < b \ (F has_vector_derivative f x) (at x)" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_integrable: "set_integrable lborel (einterval a b) f" assumes A: "((F \ real_of_ereal) \ A) (at_right a)" assumes B: "((F \ real_of_ereal) \ B) (at_left b)" shows "(LBINT x=a..b. f x) = B - A" proof - obtain u l where approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have FTCi: "\i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)" using assms approx by (auto simp: less_imp_le min_def max_def intro!: f continuous_at_imp_continuous_on interval_integral_FTC_finite intro: has_vector_derivative_at_within) have "(\i. LBINT x=l i..u i. f x) \ B - A" unfolding FTCi proof (intro tendsto_intros) show "(\x. F (l x)) \ A" using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) show "(\x. F (u x)) \ B" using B approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (u i)"], auto) qed moreover have "(\i. LBINT x=l i..u i. f x) \ (LBINT x=a..b. f x)" by (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx f_integrable]) ultimately show ?thesis by (elim LIMSEQ_unique) qed (* The second Fundamental Theorem of Calculus and existence of antiderivatives on an einterval. *) theorem interval_integral_FTC2: fixes a b c :: real and f :: "real \ 'a::euclidean_space" assumes "a \ c" "c \ b" and contf: "continuous_on {a..b} f" fixes x :: real assumes "a \ x" and "x \ b" shows "((\u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})" proof - let ?F = "(\u. LBINT y=a..u. f y)" have intf: "set_integrable lborel {a..b} f" by (rule borel_integrable_atLeastAtMost', rule contf) have "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" using \a \ x\ \x \ b\ by (auto intro: integral_has_vector_derivative continuous_on_subset [OF contf]) then have "((\u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})" by simp then have "(?F has_vector_derivative (f x)) (at x within {a..b})" by (rule has_vector_derivative_weaken) (auto intro!: assms interval_integral_eq_integral[symmetric] set_integrable_subset [OF intf]) then have "((\x. (LBINT y=c..a. f y) + ?F x) has_vector_derivative (f x)) (at x within {a..b})" by (auto intro!: derivative_eq_intros) then show ?thesis proof (rule has_vector_derivative_weaken) fix u assume "u \ {a .. b}" then show "(LBINT y=c..a. f y) + (LBINT y=a..u. f y) = (LBINT y=c..u. f y)" using assms apply (intro interval_integral_sum) apply (auto simp: interval_lebesgue_integrable_def simp del: real_scaleR_def) by (rule set_integrable_subset [OF intf], auto simp: min_def max_def) qed (insert assms, auto) qed proposition einterval_antiderivative: fixes a b :: ereal and f :: "real \ 'a::euclidean_space" assumes "a < b" and contf: "\x :: real. a < x \ x < b \ isCont f x" shows "\F. \x :: real. a < x \ x < b \ (F has_vector_derivative f x) (at x)" proof - from einterval_nonempty [OF \a < b\] obtain c :: real where [simp]: "a < c" "c < b" by (auto simp: einterval_def) let ?F = "(\u. LBINT y=c..u. f y)" show ?thesis proof (rule exI, clarsimp) fix x :: real assume [simp]: "a < x" "x < b" have 1: "a < min c x" by simp from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x" by (auto simp: einterval_def) have 2: "max c x < b" by simp from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b" by (auto simp: einterval_def) have "(?F has_vector_derivative f x) (at x within {d<..x. \d \ x; x \ e\ \ a < ereal x" using \a < ereal d\ ereal_less_ereal_Ex by auto show "\x. \d \ x; x \ e\ \ ereal x < b" using \ereal e < b\ ereal_less_eq(3) le_less_trans by blast qed then show "(?F has_vector_derivative f x) (at x within {d..e})" by (intro interval_integral_FTC2) (use \d < c\ \c < e\ \d < x\ \x < e\ in \linarith+\) qed auto then show "(?F has_vector_derivative f x) (at x)" by (force simp: has_vector_derivative_within_open [of _ "{d<..The substitution theorem\ text\Once again, three versions: first, for finite intervals, and then two versions for arbitrary intervals.\ theorem interval_integral_substitution_finite: fixes a b :: real and f :: "real \ 'a::euclidean_space" assumes "a \ b" and derivg: "\x. a \ x \ x \ b \ (g has_real_derivative (g' x)) (at x within {a..b})" and contf : "continuous_on (g ` {a..b}) f" and contg': "continuous_on {a..b} g'" shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y" proof- have v_derivg: "\x. a \ x \ x \ b \ (g has_vector_derivative (g' x)) (at x within {a..b})" - using derivg unfolding has_field_derivative_iff_has_vector_derivative . + using derivg unfolding has_real_derivative_iff_has_vector_derivative . then have contg [simp]: "continuous_on {a..b} g" by (rule continuous_on_vector_derivative) auto have 1: "\x\{a..b}. u = g x" if "min (g a) (g b) \ u" "u \ max (g a) (g b)" for u by (cases "g a \ g b") (use that assms IVT' [of g a u b] IVT2' [of g b u a] in \auto simp: min_def max_def\) obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \ d" by (metis continuous_image_closed_interval contg \a \ b\) obtain F where derivF: "\x. \a \ x; x \ b\ \ (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" using continuous_on_subset [OF contf] g_im by (metis antiderivative_continuous atLeastAtMost_iff image_subset_iff set_eq_subset) have contfg: "continuous_on {a..b} (\x. f (g x))" by (blast intro: continuous_on_compose2 contf contg) have "LBINT x. indicat_real {a..b} x *\<^sub>R g' x *\<^sub>R f (g x) = F (g b) - F (g a)" apply (rule integral_FTC_atLeastAtMost [OF \a \ b\ vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def]]) apply (auto intro!: continuous_on_scaleR contg' contfg) done then have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)" by (simp add: assms interval_integral_Icc set_lebesgue_integral_def) moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)" proof (rule interval_integral_FTC_finite) show "continuous_on {min (g a) (g b)..max (g a) (g b)} f" by (rule continuous_on_subset [OF contf]) (auto simp: image_def 1) show "(F has_vector_derivative f y) (at y within {min (g a) (g b)..max (g a) (g b)})" if y: "min (g a) (g b) \ y" "y \ max (g a) (g b)" for y proof - obtain x where "a \ x" "x \ b" "y = g x" using 1 y by force then show ?thesis by (auto simp: image_def intro!: 1 has_vector_derivative_within_subset [OF derivF]) qed qed ultimately show ?thesis by simp qed (* TODO: is it possible to lift the assumption here that g' is nonnegative? *) theorem interval_integral_substitution_integrable: fixes f :: "real \ 'a::euclidean_space" and a b u v :: ereal assumes "a < b" and deriv_g: "\x. a < ereal x \ ereal x < b \ DERIV g x :> g' x" and contf: "\x. a < ereal x \ ereal x < b \ isCont f (g x)" and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" and A: "((ereal \ g \ real_of_ereal) \ A) (at_right a)" and B: "((ereal \ g \ real_of_ereal) \ B) (at_left b)" and integrable: "set_integrable lborel (einterval a b) (\x. g' x *\<^sub>R f (g x))" and integrable2: "set_integrable lborel (einterval A B) (\x. f x)" shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" proof - obtain u l where approx [simp]: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) note less_imp_le [simp] have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) then have lessb[simp]: "\i. l i < b" using approx(4) less_eq_real_def by blast have [simp]: "\i. a < u i" by (rule order_less_trans, rule approx, auto, rule approx) have lle[simp]: "\i j. i \ j \ l j \ l i" by (rule decseqD, rule approx) have [simp]: "\i j. i \ j \ u i \ u j" by (rule incseqD, rule approx) have g_nondec [simp]: "g x \ g y" if "a < x" "x \ y" "y < b" for x y proof (rule DERIV_nonneg_imp_nondecreasing [OF \x \ y\], intro exI conjI) show "\u. x \ u \ u \ y \ (g has_real_derivative g' u) (at u)" by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that) show "\u. x \ u \ u \ y \ 0 \ g' u" by (meson assms(5) dual_order.trans le_ereal_le less_imp_le order_refl that) qed have "A \ B" and un: "einterval A B = (\i. {g(l i)<..i. g (l i)) \ A" using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (l i)" in spec, auto) hence A3: "\i. g (l i) \ A" by (intro decseq_ge, auto simp: decseq_def) have B2: "(\i. g (u i)) \ B" using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (u i)" in spec, auto) hence B3: "\i. g (u i) \ B" by (intro incseq_le, auto simp: incseq_def) have "ereal (g (l 0)) \ ereal (g (u 0))" by auto then show "A \ B" by (meson A3 B3 order.trans) { fix x :: real assume "A < x" and "x < B" then have "eventually (\i. ereal (g (l i)) < x \ x < ereal (g (u i))) sequentially" by (fast intro: eventually_conj order_tendstoD A2 B2) hence "\i. g (l i) < x \ x < g (u i)" by (simp add: eventually_sequentially, auto) } note AB = this show "einterval A B = (\i. {g(l i)<.. (\i. {g(l i)<..i. {g(l i)<.. einterval A B" proof (clarsimp simp add: einterval_def, intro conjI) show "\x i. \g (l i) < x; x < g (u i)\ \ A < ereal x" using A3 le_ereal_less by blast show "\x i. \g (l i) < x; x < g (u i)\ \ ereal x < B" using B3 ereal_le_less by blast qed qed qed (* finally, the main argument *) have eq1: "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)" for i apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) - unfolding has_field_derivative_iff_has_vector_derivative[symmetric] + unfolding has_real_derivative_iff_has_vector_derivative[symmetric] apply (auto intro!: continuous_at_imp_continuous_on contf contg') done have "(\i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) \ (LBINT x=a..b. g' x *\<^sub>R f (g x))" apply (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx]) by (rule assms) hence 2: "(\i. (LBINT y=g (l i)..g (u i). f y)) \ (LBINT x=a..b. g' x *\<^sub>R f (g x))" by (simp add: eq1) have incseq: "incseq (\i. {g (l i)<..i. set_lebesgue_integral lborel {g (l i)<.. set_lebesgue_integral lborel (einterval A B) f" unfolding un by (rule set_integral_cont_up) (use incseq integrable2 un in auto) then have "(\i. (LBINT y=g (l i)..g (u i). f y)) \ (LBINT x = A..B. f x)" by (simp add: interval_lebesgue_integral_le_eq \A \ B\) thus ?thesis by (intro LIMSEQ_unique [OF _ 2]) qed (* TODO: the last two proofs are only slightly different. Factor out common part? An alternative: make the second one the main one, and then have another lemma that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *) theorem interval_integral_substitution_nonneg: fixes f g g':: "real \ real" and a b u v :: ereal assumes "a < b" and deriv_g: "\x. a < ereal x \ ereal x < b \ DERIV g x :> g' x" and contf: "\x. a < ereal x \ ereal x < b \ isCont f (g x)" and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and f_nonneg: "\x. a < ereal x \ ereal x < b \ 0 \ f (g x)" (* TODO: make this AE? *) and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" and A: "((ereal \ g \ real_of_ereal) \ A) (at_right a)" and B: "((ereal \ g \ real_of_ereal) \ B) (at_left b)" and integrable_fg: "set_integrable lborel (einterval a b) (\x. f (g x) * g' x)" shows "set_integrable lborel (einterval A B) f" "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))" proof - from einterval_Icc_approximation[OF \a < b\] obtain u l where approx [simp]: "einterval a b = (\i. {l i..u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < ereal (l i)" "\i. ereal (u i) < b" "(\x. ereal (l x)) \ a" "(\x. ereal (u x)) \ b" by this auto have aless[simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have lessb[simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have llb[simp]: "\i. l i < b" using lessb approx(4) less_eq_real_def by blast have alu[simp]: "\i. a < u i" by (rule order_less_trans, rule approx, auto, rule approx) have [simp]: "\i j. i \ j \ l j \ l i" by (rule decseqD, rule approx) have uleu[simp]: "\i j. i \ j \ u i \ u j" by (rule incseqD, rule approx) have g_nondec [simp]: "g x \ g y" if "a < x" "x \ y" "y < b" for x y proof (rule DERIV_nonneg_imp_nondecreasing [OF \x \ y\], intro exI conjI) show "\u. x \ u \ u \ y \ (g has_real_derivative g' u) (at u)" by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that) show "\u. x \ u \ u \ y \ 0 \ g' u" by (meson g'_nonneg less_ereal.simps(1) less_trans not_less that) qed have "A \ B" and un: "einterval A B = (\i. {g(l i)<..i. g (l i)) \ A" using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (l i)" in spec, auto) hence A3: "\i. g (l i) \ A" by (intro decseq_ge, auto simp: decseq_def) have B2: "(\i. g (u i)) \ B" using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (u i)" in spec, auto) hence B3: "\i. g (u i) \ B" by (intro incseq_le, auto simp: incseq_def) have "ereal (g (l 0)) \ ereal (g (u 0))" by (auto simp: less_imp_le) then show "A \ B" by (meson A3 B3 order.trans) { fix x :: real assume "A < x" and "x < B" then have "eventually (\i. ereal (g (l i)) < x \ x < ereal (g (u i))) sequentially" by (fast intro: eventually_conj order_tendstoD A2 B2) hence "\i. g (l i) < x \ x < g (u i)" by (simp add: eventually_sequentially, auto) } note AB = this show "einterval A B = (\i. {g(l i)<.. (\i. {g (l i)<..i. {g (l i)<.. einterval A B" apply (clarsimp simp: einterval_def, intro conjI) using A3 le_ereal_less apply blast using B3 ereal_le_less by blast qed qed (* finally, the main argument *) have eq1: "(LBINT x=l i.. u i. (f (g x) * g' x)) = (LBINT y=g (l i)..g (u i). f y)" for i proof - have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)" apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) - unfolding has_field_derivative_iff_has_vector_derivative[symmetric] + unfolding has_real_derivative_iff_has_vector_derivative[symmetric] apply (auto simp: less_imp_le intro!: continuous_at_imp_continuous_on contf contg') done then show ?thesis by (simp add: ac_simps) qed have incseq: "incseq (\i. {g (l i)<..c \ l i. c \ u i \ x = g c" if "g (l i) \ x" "x \ g (u i)" for x i proof - have "continuous_on {l i..u i} g" by (force intro!: DERIV_isCont deriv_g continuous_at_imp_continuous_on) with that show ?thesis using IVT' [of g] approx(4) dual_order.strict_implies_order by blast qed have "continuous_on {g (l i)..g (u i)} f" for i apply (intro continuous_intros continuous_at_imp_continuous_on) using contf img by force then have int_f: "\i. set_integrable lborel {g (l i)<..i. {g (l i)<..i. LBINT x=l i..u i. f (g x) * g' x) \ ?l" by (intro assms interval_integral_Icc_approx_integrable [OF \a < b\ approx]) hence "(\i. (LBINT y=g (l i)..g (u i). f y)) \ ?l" by (simp add: eq1) then show "(\i. set_lebesgue_integral lborel {g (l i)<.. ?l" unfolding interval_lebesgue_integral_def by (auto simp: less_imp_le) have "\x i. g (l i) \ x \ x \ g (u i) \ 0 \ f x" using aless f_nonneg img lessb by blast then show "\x i. x \ {g (l i)<.. 0 \ f x" using less_eq_real_def by auto qed (auto simp: greaterThanLessThan_borel) thus "set_integrable lborel (einterval A B) f" by (simp add: un) have "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" proof (rule interval_integral_substitution_integrable) show "set_integrable lborel (einterval a b) (\x. g' x *\<^sub>R f (g x))" using integrable_fg by (simp add: ac_simps) qed fact+ then show "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))" by (simp add: ac_simps) qed syntax "_complex_lebesgue_borel_integral" :: "pttrn \ real \ complex" ("(2CLBINT _. _)" [0,60] 60) translations "CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\x. f)" syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ complex" ("(3CLBINT _:_. _)" [0,60,61] 60) translations "CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\x. f)" abbreviation complex_interval_lebesgue_integral :: "real measure \ ereal \ ereal \ (real \ complex) \ complex" where "complex_interval_lebesgue_integral M a b f \ interval_lebesgue_integral M a b f" abbreviation complex_interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ complex) \ bool" where "complex_interval_lebesgue_integrable M a b f \ interval_lebesgue_integrable M a b f" syntax "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \ ereal \ ereal \ real \ complex" ("(4CLBINT _=_.._. _)" [0,60,60,61] 60) translations "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\x. f)" proposition interval_integral_norm: fixes f :: "real \ 'a :: {banach, second_countable_topology}" shows "interval_lebesgue_integrable lborel a b f \ a \ b \ norm (LBINT t=a..b. f t) \ LBINT t=a..b. norm (f t)" using integral_norm_bound[of lborel "\x. indicator (einterval a b) x *\<^sub>R f x"] by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def) proposition interval_integral_norm2: "interval_lebesgue_integrable lborel a b f \ norm (LBINT t=a..b. f t) \ \LBINT t=a..b. norm (f t)\" proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b]) next case (le a b) then have "\LBINT t=a..b. norm (f t)\ = LBINT t=a..b. norm (f t)" using integrable_norm[of lborel "\x. indicator (einterval a b) x *\<^sub>R f x"] by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def intro!: integral_nonneg_AE abs_of_nonneg) then show ?case using le by (simp add: interval_integral_norm) qed (* TODO: should we have a library of facts like these? *) lemma integral_cos: "t \ 0 \ LBINT x=a..b. cos (t * x) = sin (t * b) / t - sin (t * a) / t" apply (intro interval_integral_FTC_finite continuous_intros) - by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric]) + by (auto intro!: derivative_eq_intros simp: has_real_derivative_iff_has_vector_derivative[symmetric]) end diff --git a/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy b/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy --- a/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy +++ b/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy @@ -1,406 +1,406 @@ (* Title: HOL/Analysis/Lebesgue_Integral_Substitution.thy Author: Manuel Eberl Provides lemmas for integration by substitution for the basic integral types. Note that the substitution function must have a nonnegative derivative. This could probably be weakened somehow. *) section \Integration by Substition for the Lebesgue Integral\ theory Lebesgue_Integral_Substitution imports Interval_Integral begin lemma nn_integral_substitution_aux: fixes f :: "real \ ennreal" assumes Mf: "f \ borel_measurable borel" assumes nonnegf: "\x. f x \ 0" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a < b" shows "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof- from \a < b\ have [simp]: "a \ b" by simp from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and Mg': "set_borel_measurable borel {a..b} g'" by (simp_all only: set_measurable_continuous_on_ivl) from derivg have derivg': "\x. x \ {a..b} \ (g has_vector_derivative g' x) (at x)" - by (simp only: has_field_derivative_iff_has_vector_derivative) + by (simp only: has_real_derivative_iff_has_vector_derivative) have real_ind[simp]: "\A x. enn2real (indicator A x) = indicator A x" by (auto split: split_indicator) have ennreal_ind[simp]: "\A x. ennreal (indicator A x) = indicator A x" by (auto split: split_indicator) have [simp]: "\x A. indicator A (g x) = indicator (g -` A) x" by (auto split: split_indicator) from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all with monog have [simp]: "g a \ g b" by (auto intro: mono_onD) show ?thesis proof (induction rule: borel_measurable_induct[OF Mf, case_names cong set mult add sup]) case (cong f1 f2) from cong.hyps(3) have "f1 = f2" by auto with cong show ?case by simp next case (set A) from set.hyps show ?case proof (induction rule: borel_set_induct) case empty thus ?case by simp next case (interval c d) { fix u v :: real assume asm: "{u..v} \ {g a..g b}" "u \ v" obtain u' v' where u'v': "{a..b} \ g-`{u..v} = {u'..v'}" "u' \ v'" "g u' = u" "g v' = v" using asm by (rule_tac continuous_interval_vimage_Int[OF contg monog, of u v]) simp_all hence "{u'..v'} \ {a..b}" "{u'..v'} \ g -` {u..v}" by blast+ with u'v'(2) have "u' \ g -` {u..v}" "v' \ g -` {u..v}" by auto from u'v'(1) have [simp]: "{a..b} \ g -` {u..v} \ sets borel" by simp have A: "continuous_on {min u' v'..max u' v'} g'" by (simp only: u'v' max_absorb2 min_absorb1) (intro continuous_on_subset[OF contg'], insert u'v', auto) have "\x. x \ {u'..v'} \ (g has_real_derivative g' x) (at x within {u'..v'})" using asm by (intro has_field_derivative_subset[OF derivg] subsetD[OF \{u'..v'} \ {a..b}\]) auto hence B: "\x. min u' v' \ x \ x \ max u' v' \ (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})" by (simp only: u'v' max_absorb2 min_absorb1) - (auto simp: has_field_derivative_iff_has_vector_derivative) + (auto simp: has_real_derivative_iff_has_vector_derivative) have "integrable lborel (\x. indicator ({a..b} \ g -` {u..v}) x *\<^sub>R g' x)" using set_integrable_subset borel_integrable_atLeastAtMost'[OF contg'] by (metis \{u'..v'} \ {a..b}\ eucl_ivals(5) set_integrable_def sets_lborel u'v'(1)) hence "(\\<^sup>+x. ennreal (g' x) * indicator ({a..b} \ g-` {u..v}) x \lborel) = LBINT x:{a..b} \ g-`{u..v}. g' x" unfolding set_lebesgue_integral_def by (subst nn_integral_eq_integral[symmetric]) (auto intro!: derivg_nonneg nn_integral_cong split: split_indicator) also from interval_integral_FTC_finite[OF A B] have "LBINT x:{a..b} \ g-`{u..v}. g' x = v - u" by (simp add: u'v' interval_integral_Icc \u \ v\) finally have "(\\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \ g -` {u..v}) x \lborel) = ennreal (v - u)" . } note A = this have "(\\<^sup>+x. indicator {c..d} (g x) * ennreal (g' x) * indicator {a..b} x \lborel) = (\\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \ g -` {c..d}) x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "{a..b} \ g-`{c..d} = {a..b} \ g-`{max (g a) c..min (g b) d}" using \a \ b\ \c \ d\ by (auto intro!: monog intro: order.trans) also have "(\\<^sup>+ x. ennreal (g' x) * indicator ... x \lborel) = (if max (g a) c \ min (g b) d then min (g b) d - max (g a) c else 0)" using \c \ d\ by (simp add: A) also have "... = (\\<^sup>+ x. indicator ({g a..g b} \ {c..d}) x \lborel)" by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:) also have "... = (\\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (auto split: split_indicator) finally show ?case .. next case (compl A) note \A \ sets borel\[measurable] from emeasure_mono[of "A \ {g a..g b}" "{g a..g b}" lborel] have [simp]: "emeasure lborel (A \ {g a..g b}) \ top" by (auto simp: top_unique) have [simp]: "g -` A \ {a..b} \ sets borel" by (rule set_borel_measurable_sets[OF Mg]) auto have [simp]: "g -` (-A) \ {a..b} \ sets borel" by (rule set_borel_measurable_sets[OF Mg]) auto have "(\\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. indicator (-A \ {g a..g b}) x \lborel)" by (rule nn_integral_cong) (simp split: split_indicator) also from compl have "... = emeasure lborel ({g a..g b} - A)" using derivg_nonneg by (simp add: vimage_Compl diff_eq Int_commute[of "-A"]) also have "{g a..g b} - A = {g a..g b} - A \ {g a..g b}" by blast also have "emeasure lborel ... = g b - g a - emeasure lborel (A \ {g a..g b})" using \A \ sets borel\ by (subst emeasure_Diff) auto also have "emeasure lborel (A \ {g a..g b}) = \\<^sup>+x. indicator A x * indicator {g a..g b} x \lborel" using \A \ sets borel\ by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong) (simp split: split_indicator) also have "... = \\<^sup>+ x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel" (is "_ = ?I") by (subst compl.IH, intro nn_integral_cong) (simp split: split_indicator) also have "g b - g a = LBINT x:{a..b}. g' x" using derivg' unfolding set_lebesgue_integral_def by (intro integral_FTC_atLeastAtMost[symmetric]) (auto intro: continuous_on_subset[OF contg'] has_field_derivative_subset[OF derivg] has_vector_derivative_at_within) also have "ennreal ... = \\<^sup>+ x. g' x * indicator {a..b} x \lborel" using borel_integrable_atLeastAtMost'[OF contg'] unfolding set_lebesgue_integral_def by (subst nn_integral_eq_integral) (simp_all add: mult.commute derivg_nonneg set_integrable_def split: split_indicator) also have Mg'': "(\x. indicator (g -` A \ {a..b}) x * ennreal (g' x * indicator {a..b} x)) \ borel_measurable borel" using Mg' by (intro borel_measurable_times_ennreal borel_measurable_indicator) (simp_all add: mult.commute set_borel_measurable_def) have le: "(\\<^sup>+x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel) \ (\\<^sup>+x. ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_mono) (simp split: split_indicator add: derivg_nonneg) note integrable = borel_integrable_atLeastAtMost'[OF contg'] with le have notinf: "(\\<^sup>+x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel) \ top" by (auto simp: real_integrable_def nn_integral_set_ennreal mult.commute top_unique set_integrable_def) have "(\\<^sup>+ x. g' x * indicator {a..b} x \lborel) - ?I = \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) - indicator (g -` A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel" apply (intro nn_integral_diff[symmetric]) apply (insert Mg', simp add: mult.commute set_borel_measurable_def) [] apply (insert Mg'', simp) [] apply (simp split: split_indicator add: derivg_nonneg) apply (rule notinf) apply (simp split: split_indicator add: derivg_nonneg) done also have "... = \\<^sup>+ x. indicator (-A) (g x) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (simp split: split_indicator) finally show ?case . next case (union f) then have [simp]: "\i. {a..b} \ g -` f i \ sets borel" by (subst Int_commute, intro set_borel_measurable_sets[OF Mg]) auto have "g -` (\i. f i) \ {a..b} = (\i. {a..b} \ g -` f i)" by auto hence "g -` (\i. f i) \ {a..b} \ sets borel" by (auto simp del: UN_simps) have "(\\<^sup>+x. indicator (\i. f i) x * indicator {g a..g b} x \lborel) = \\<^sup>+x. indicator (\i. {g a..g b} \ f i) x \lborel" by (intro nn_integral_cong) (simp split: split_indicator) also from union have "... = emeasure lborel (\i. {g a..g b} \ f i)" by simp also from union have "... = (\i. emeasure lborel ({g a..g b} \ f i))" by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def) also from union have "... = (\i. \\<^sup>+x. indicator ({g a..g b} \ f i) x \lborel)" by simp also have "(\i. \\<^sup>+x. indicator ({g a..g b} \ f i) x \lborel) = (\i. \\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \lborel)" by (intro ext nn_integral_cong) (simp split: split_indicator) also from union.IH have "(\i. \\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \lborel) = (\i. \\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by simp also have "(\i. \\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel) = (\i. \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x \lborel)" by (intro ext nn_integral_cong) (simp split: split_indicator) also have "(\i. ... i) = \\<^sup>+ x. (\i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x) \lborel" using Mg' apply (intro nn_integral_suminf[symmetric]) apply (rule borel_measurable_times_ennreal, simp add: mult.commute set_borel_measurable_def) apply (rule borel_measurable_indicator, subst sets_lborel) apply (simp_all split: split_indicator add: derivg_nonneg) done also have "(\x i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x) = (\x i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x)" by (intro ext) (simp split: split_indicator) also have "(\\<^sup>+ x. (\i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \lborel) = \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * (\i. indicator (g -` f i) x) \lborel" by (intro nn_integral_cong) (auto split: split_indicator simp: derivg_nonneg) also from union have "(\x. \i. indicator (g -` f i) x :: ennreal) = (\x. indicator (\i. g -` f i) x)" by (intro ext suminf_indicator) (auto simp: disjoint_family_on_def) also have "(\\<^sup>+x. ennreal (g' x * indicator {a..b} x) * ... x \lborel) = (\\<^sup>+x. indicator (\i. f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) finally show ?case . qed next case (mult f c) note Mf[measurable] = \f \ borel_measurable borel\ let ?I = "indicator {a..b}" have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. f (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have Mf': "(\x. f (g x) * ennreal (g' x) * ?I x) \ borel_measurable borel" . with mult show ?case by (subst (1 2 3) mult_ac, subst (1 2) nn_integral_cmult) (simp_all add: mult_ac) next case (add f2 f1) let ?I = "indicator {a..b}" { fix f :: "real \ ennreal" assume Mf: "f \ borel_measurable borel" have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. f (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have "(\x. f (g x) * ennreal (g' x) * ?I x) \ borel_measurable borel" . } note Mf' = this[OF \f1 \ borel_measurable borel\] this[OF \f2 \ borel_measurable borel\] have "(\\<^sup>+ x. (f1 x + f2 x) * indicator {g a..g b} x \lborel) = (\\<^sup>+ x. f1 x * indicator {g a..g b} x + f2 x * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from add have "... = (\\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x \lborel) + (\\<^sup>+ x. f2 (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (simp_all add: nn_integral_add) also from add have "... = (\\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x + f2 (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_add[symmetric]) (auto simp add: Mf' derivg_nonneg split: split_indicator) also have "... = \\<^sup>+ x. (f1 (g x) + f2 (g x)) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (simp split: split_indicator add: distrib_right) finally show ?case . next case (sup F) { fix i let ?I = "indicator {a..b}" have "(\x. F i (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (rule_tac borel_measurable_times_ennreal, rule_tac measurable_compose[OF _ sup.hyps(1)]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. F i (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. F i (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have "... \ borel_measurable borel" . } note Mf' = this have "(\\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \lborel) = \\<^sup>+x. (SUP i. F i x* indicator {g a..g b} x) \lborel" by (intro nn_integral_cong) (simp split: split_indicator) also from sup have "... = (SUP i. \\<^sup>+x. F i x* indicator {g a..g b} x \lborel)" by (intro nn_integral_monotone_convergence_SUP) (auto simp: incseq_def le_fun_def split: split_indicator) also from sup have "... = (SUP i. \\<^sup>+x. F i (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by simp also from sup have "... = \\<^sup>+x. (SUP i. F i (g x) * ennreal (g' x) * indicator {a..b} x) \lborel" by (intro nn_integral_monotone_convergence_SUP[symmetric]) (auto simp: incseq_def le_fun_def derivg_nonneg Mf' split: split_indicator intro!: mult_right_mono) also from sup have "... = \\<^sup>+x. (SUP i. F i (g x)) * ennreal (g' x) * indicator {a..b} x \lborel" by (subst mult.assoc, subst mult.commute, subst SUP_mult_left_ennreal) (auto split: split_indicator simp: derivg_nonneg mult_ac) finally show ?case by (simp add: image_comp) qed qed theorem nn_integral_substitution: fixes f :: "real \ real" assumes Mf[measurable]: "set_borel_measurable borel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof (cases "a = b") assume "a \ b" with \a \ b\ have "a < b" by auto let ?f' = "\x. f x * indicator {g a..g b} x" from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all have bounds: "\x. x \ a \ x \ b \ g x \ g a" "\x. x \ a \ x \ b \ g x \ g b" by (auto intro: monog) from derivg_nonneg have nonneg: "\f x. x \ a \ x \ b \ g' x \ 0 \ f x * ennreal (g' x) \ 0 \ f x \ 0" by (force simp: field_simps) have nonneg': "\x. a \ x \ x \ b \ \ 0 \ f (g x) \ 0 \ f (g x) * g' x \ g' x = 0" by (metis atLeastAtMost_iff derivg_nonneg eq_iff mult_eq_0_iff mult_le_0_iff) have "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. ennreal (?f' x) * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (auto split: split_indicator split_max simp: zero_ennreal.rep_eq ennreal_neg) also have "... = \\<^sup>+ x. ?f' (g x) * ennreal (g' x) * indicator {a..b} x \lborel" using Mf by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \a < b\]) (auto simp add: mult.commute set_borel_measurable_def) also have "... = \\<^sup>+ x. f (g x) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (auto split: split_indicator simp: max_def dest: bounds) also have "... = \\<^sup>+x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel" by (intro nn_integral_cong) (auto simp: mult.commute derivg_nonneg ennreal_mult' split: split_indicator) finally show ?thesis . qed auto theorem integral_substitution: assumes integrable: "set_integrable lborel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "set_integrable lborel {a..b} (\x. f (g x) * g' x)" and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" proof- from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) with contg' have Mg: "set_borel_measurable borel {a..b} g" and Mg': "set_borel_measurable borel {a..b} g'" by (simp_all only: set_measurable_continuous_on_ivl) from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all have "(\x. ennreal (f x) * indicator {g a..g b} x) = (\x. ennreal (f x * indicator {g a..g b} x))" by (intro ext) (simp split: split_indicator) with integrable have M1: "(\x. f x * indicator {g a..g b} x) \ borel_measurable borel" by (force simp: mult.commute set_integrable_def) from integrable have M2: "(\x. -f x * indicator {g a..g b} x) \ borel_measurable borel" by (force simp: mult.commute set_integrable_def) have "LBINT x. (f x :: real) * indicator {g a..g b} x = enn2real (\\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \lborel) - enn2real (\\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \lborel)" using integrable unfolding set_integrable_def by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute) also have *: "(\\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \lborel) = (\\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from M1 * have A: "(\\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel)" by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ennreal mult.commute set_borel_measurable_def) also have **: "(\\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \lborel) = (\\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from M2 ** have B: "(\\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \lborel)" by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ennreal mult.commute set_borel_measurable_def) also { from integrable have Mf: "set_borel_measurable borel {g a..g b} f" unfolding set_borel_measurable_def set_integrable_def by simp from measurable_compose Mg Mf Mg' borel_measurable_times have "(\x. f (g x * indicator {a..b} x) * indicator {g a..g b} (g x * indicator {a..b} x) * (g' x * indicator {a..b} x)) \ borel_measurable borel" (is "?f \ _") by (simp add: mult.commute set_borel_measurable_def) also have "?f = (\x. f (g x) * g' x * indicator {a..b} x)" using monog by (intro ext) (auto split: split_indicator) finally show "set_integrable lborel {a..b} (\x. f (g x) * g' x)" using A B integrable unfolding real_integrable_def set_integrable_def by (simp_all add: nn_integral_set_ennreal mult.commute) } note integrable' = this have "enn2real (\\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel) - enn2real (\\<^sup>+ x. ennreal (-f (g x) * g' x * indicator {a..b} x) \lborel) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable' unfolding set_integrable_def by (subst real_lebesgue_integral_def) (simp_all add: field_simps) finally show "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" . qed theorem interval_integral_substitution: assumes integrable: "set_integrable lborel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "set_integrable lborel {a..b} (\x. f (g x) * g' x)" and "(LBINT x=g a..g b. f x) = (LBINT x=a..b. f (g x) * g' x)" apply (rule integral_substitution[OF assms], simp, simp) apply (subst (1 2) interval_integral_Icc, fact) apply (rule deriv_nonneg_imp_mono[OF derivg derivg_nonneg], simp, simp, fact) using integral_substitution(2)[OF assms] apply (simp add: mult.commute set_lebesgue_integral_def) done lemma set_borel_integrable_singleton[simp]: "set_integrable lborel {x} (f :: real \ real)" unfolding set_integrable_def by (subst integrable_discrete_difference[where X="{x}" and g="\_. 0"]) auto end diff --git a/src/HOL/Analysis/Topology_Euclidean_Space.thy b/src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy @@ -1,2579 +1,2590 @@ (* 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 \Vector Analysis\ theory Topology_Euclidean_Space imports Elementary_Normed_Spaces Linear_Algebra Norm_Arith begin section \Elementary Topology in Euclidean Space\ lemma euclidean_dist_l2: fixes x y :: "'a :: euclidean_space" shows "dist x y = L2_set (\i. dist (x \ i) (y \ i)) Basis" unfolding dist_norm norm_eq_sqrt_inner L2_set_def by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) lemma norm_nth_le: "norm (x \ i) \ norm x" if "i \ Basis" proof - have "(x \ i)\<^sup>2 = (\i\{i}. (x \ i)\<^sup>2)" by simp also have "\ \ (\i\Basis. (x \ i)\<^sup>2)" by (intro sum_mono2) (auto simp: that) finally show ?thesis unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def by (auto intro!: real_le_rsqrt) qed subsection\<^marker>\tag unimportant\ \Continuity of the representation WRT an orthogonal basis\ lemma orthogonal_Basis: "pairwise orthogonal Basis" by (simp add: inner_not_same_Basis orthogonal_def pairwise_def) lemma representation_bound: fixes B :: "'N::real_inner set" assumes "finite B" "independent B" "b \ B" and orth: "pairwise orthogonal B" obtains m where "m > 0" "\x. x \ span B \ \representation B x b\ \ m * norm x" proof fix x assume x: "x \ span B" have "b \ 0" using \independent B\ \b \ B\ dependent_zero by blast have [simp]: "b \ b' = (if b' = b then (norm b)\<^sup>2 else 0)" if "b \ B" "b' \ B" for b b' using orth by (simp add: orthogonal_def pairwise_def norm_eq_sqrt_inner that) have "norm x = norm (\b\B. representation B x b *\<^sub>R b)" using real_vector.sum_representation_eq [OF \independent B\ x \finite B\ order_refl] by simp also have "\ = sqrt ((\b\B. representation B x b *\<^sub>R b) \ (\b\B. representation B x b *\<^sub>R b))" by (simp add: norm_eq_sqrt_inner) also have "\ = sqrt (\b\B. (representation B x b *\<^sub>R b) \ (representation B x b *\<^sub>R b))" using \finite B\ by (simp add: inner_sum_left inner_sum_right if_distrib [of "\x. _ * x"] cong: if_cong sum.cong_simp) also have "\ = sqrt (\b\B. (norm (representation B x b *\<^sub>R b))\<^sup>2)" by (simp add: mult.commute mult.left_commute power2_eq_square) also have "\ = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" by (simp add: norm_mult power_mult_distrib) finally have "norm x = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" . moreover have "sqrt ((representation B x b)\<^sup>2 * (norm b)\<^sup>2) \ sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ B\ \finite B\ by (auto intro: member_le_sum) then have "\representation B x b\ \ (1 / norm b) * sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ 0\ by (simp add: field_split_simps real_sqrt_mult del: real_sqrt_le_iff) ultimately show "\representation B x b\ \ (1 / norm b) * norm x" by simp next show "0 < 1 / norm b" using \independent B\ \b \ B\ dependent_zero by auto qed lemma continuous_on_representation: fixes B :: "'N::euclidean_space set" assumes "finite B" "independent B" "b \ B" "pairwise orthogonal B" shows "continuous_on (span B) (\x. representation B x b)" proof show "\d>0. \x'\span B. dist x' x < d \ dist (representation B x' b) (representation B x b) \ e" if "e > 0" "x \ span B" for x e proof - obtain m where "m > 0" and m: "\x. x \ span B \ \representation B x b\ \ m * norm x" using assms representation_bound by blast show ?thesis unfolding dist_norm proof (intro exI conjI ballI impI) show "e/m > 0" by (simp add: \e > 0\ \m > 0\) show "norm (representation B x' b - representation B x b) \ e" if x': "x' \ span B" and less: "norm (x'-x) < e/m" for x' proof - have "\representation B (x'-x) b\ \ m * norm (x'-x)" using m [of "x'-x"] \x \ span B\ span_diff x' by blast also have "\ < e" by (metis \m > 0\ less mult.commute pos_less_divide_eq) finally have "\representation B (x'-x) b\ \ e" by simp then show ?thesis by (simp add: \x \ span B\ \independent B\ representation_diff x') qed qed qed qed subsection\<^marker>\tag unimportant\\Balls in Euclidean Space\ lemma cball_subset_cball_iff: fixes a :: "'a :: euclidean_space" shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r \ r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also from \a \ a'\ have "... = \- norm (a - a') - r\" by simp (simp add: field_simps) finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" by linarith from \a \ a'\ show ?thesis using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] by (simp add: dist_norm scaleR_add_left) qed then show ?rhs by (simp add: dist_norm) qed qed metric lemma cball_subset_ball_iff: "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" (is "?lhs \ ?rhs") for a :: "'a::euclidean_space" proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r < r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have False if "norm (a - a') + r \ r'" proof - from that have "\r' - norm (a - a')\ \ r" by (simp split: abs_split) (metis \0 \ r\ \?lhs\ centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) then show ?thesis using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ apply (simp add: dist_norm) apply (simp add: scaleR_left_diff_distrib) apply (simp add: field_simps) done qed then show ?thesis by force qed then show ?rhs by (simp add: dist_norm) qed next assume ?rhs then show ?lhs by metric qed lemma ball_subset_cball_iff: "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") for a :: "'a::euclidean_space" proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "(cball a r \ cball a' r')" by (metis False closed_cball closure_ball closure_closed closure_mono not_less) with False show ?rhs by (fastforce iff: cball_subset_cball_iff) next assume ?rhs with False show ?lhs by metric qed qed lemma ball_subset_ball_iff: fixes a :: "'a :: euclidean_space" shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "0 < r'" using False by metric then have "(cball a r \ cball a' r')" by (metis False\?lhs\ closure_ball closure_mono not_less) then show ?rhs using False cball_subset_cball_iff by fastforce qed metric qed lemma ball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d \ 0 \ e \ 0") case True with \?lhs\ show ?rhs by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset ball_subset_ball_iff) qed lemma cball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d < 0 \ e < 0") case True with \?lhs\ show ?rhs by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset cball_subset_cball_iff) qed lemma ball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ done next assume ?rhs then show ?lhs by auto qed lemma cball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = ball y e \ d < 0 \ e \ 0" using ball_eq_cball_iff by blast lemma finite_ball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\ball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "0 < e1" and e1_b:"ball p e1 \ S" using open_contains_ball_eq[OF \open S\] assms by auto obtain e2 where "0 < e2" and "\x\X. x \ p \ e2 \ dist p x" using finite_set_avoid[OF \finite X\,of p] by auto hence "\w\ball p (min e1 e2). w\S \ (w\p \ w\X)" using e1_b by auto thus "\e>0. \w\ball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ \e1>0\ apply (rule_tac x="min e1 e2" in exI) by auto qed lemma finite_cball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\cball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "e1>0" and e1: "\w\ball p e1. w\S \ (w\p \ w\X)" using finite_ball_avoid[OF assms] by auto define e2 where "e2 \ e1/2" have "e2>0" and "e2 < e1" unfolding e2_def using \e1>0\ by auto then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) then show "\e>0. \w\cball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ e1 by auto qed lemma dim_cball: assumes "e > 0" shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)" proof - { fix x :: "'n::euclidean_space" define y where "y = (e / norm x) *\<^sub>R x" then have "y \ cball 0 e" using assms by auto moreover have *: "x = (norm x / e) *\<^sub>R y" using y_def assms by simp moreover from * have "x = (norm x/e) *\<^sub>R y" by auto ultimately have "x \ span (cball 0 e)" using span_scale[of y "cball 0 e" "norm x/e"] span_superset[of "cball 0 e"] by (simp add: span_base) } then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" by auto then show ?thesis using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto) qed subsection \Boxes\ abbreviation\<^marker>\tag important\ One :: "'a::euclidean_space" where "One \ \Basis" lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False proof - have "dependent (Basis :: 'a set)" apply (simp add: dependent_finite) apply (rule_tac x="\i. 1" in exI) using SOME_Basis apply (auto simp: assms) done with independent_Basis show False by force qed corollary\<^marker>\tag unimportant\ One_neq_0[iff]: "One \ 0" by (metis One_non_0) corollary\<^marker>\tag unimportant\ Zero_neq_One[iff]: "0 \ One" by (metis One_non_0) definition\<^marker>\tag important\ (in euclidean_space) eucl_less (infix " (\i\Basis. a \ i < b \ i)" definition\<^marker>\tag important\ box_eucl_less: "box a b = {x. a x \tag important\ "cbox a b = {x. \i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" lemma box_def: "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" and in_box_eucl_less: "x \ box a b \ a x box a b \ (\i\Basis. a \ i < x \ i \ x \ i < b \ i)" "x \ cbox a b \ (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i)" by (auto simp: box_eucl_less eucl_less_def cbox_def) lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \ cbox c d" by (force simp: cbox_def Basis_prod_def) lemma cbox_Pair_iff [iff]: "(x, y) \ cbox (a, c) (b, d) \ x \ cbox a b \ y \ cbox c d" by (force simp: cbox_Pair_eq) lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (cbox a b \ cbox c d)" apply (auto simp: cbox_def Basis_complex_def) apply (rule_tac x = "(Re x, Im x)" in image_eqI) using complex_eq by auto lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \ cbox a b = {} \ cbox c d = {}" by (force simp: cbox_Pair_eq) lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" by auto lemma mem_box_real[simp]: "(x::real) \ box a b \ a < x \ x < b" "(x::real) \ cbox a b \ a \ x \ x \ b" by (auto simp: mem_box) lemma box_real[simp]: fixes a b:: real shows "box a b = {a <..< b}" "cbox a b = {a .. b}" by auto lemma box_Int_box: fixes a :: "'a::euclidean_space" shows "box a b \ box c d = box (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma rational_boxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ box a b \ box a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by (auto) have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\xa. a xa \ \ \ a xa < x \ xa \ x \ xa - a xa < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\xa. b xa \ \ \ x \ xa < b xa \ b xa - x \ xa < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ box ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i < y\i \ y\i < b i" using * i by (auto simp: box_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) qed (insert a b, auto simp: box_def) qed lemma open_UNION_box: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. box (a' f) (b' f) \ M}" shows "M = (\f\I. box (a' f) (b' f))" proof - have "x \ (\f\I. box (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ box a b" "\i \ Basis. a \ i \ \" "\i\Basis. b \ i \ \" "box a b \ ball x e" using rational_boxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_box: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = box a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. box (?a f) (?b f) \ S}" let ?\ = "(\f. box (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_box [OF assms] by metis qed auto qed lemma rational_cboxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ cbox a b \ cbox a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by auto have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\u. a u \ \ \ a u < x \ u \ x \ u - a u < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\u. b u \ \ \ x \ u < b u \ b u - x \ u < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ cbox ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i \ y\i \ y\i \ b i" using * i by (auto simp: cbox_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) next show "x \ cbox (\i\Basis. a i *\<^sub>R i) (\i\Basis. b i *\<^sub>R i)" using a b less_imp_le by (auto simp: cbox_def) qed (use a b cbox_def in auto) qed lemma open_UNION_cbox: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. cbox (a' f) (b' f) \ M}" shows "M = (\f\I. cbox (a' f) (b' f))" proof - have "x \ (\f\I. cbox (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ cbox a b" "\i \ Basis. a \ i \ \" "\i \ Basis. b \ i \ \" "cbox a b \ ball x e" using rational_cboxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_cbox: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = cbox a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. cbox (?a f) (?b f) \ S}" let ?\ = "(\f. cbox (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_cbox [OF assms] by metis qed auto qed lemma box_eq_empty: fixes a :: "'a::euclidean_space" shows "(box a b = {} \ (\i\Basis. b\i \ a\i))" (is ?th1) and "(cbox a b = {} \ (\i\Basis. b\i < a\i))" (is ?th2) proof - { fix i x assume i: "i\Basis" and as:"b\i \ a\i" and x:"x\box a b" then have "a \ i < x \ i \ x \ i < b \ i" unfolding mem_box by (auto simp: box_def) then have "a\i < b\i" by auto then have False using as by auto } moreover { assume as: "\i\Basis. \ (b\i \ a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i: "i \ Basis" have "a\i < b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i < ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i < b\i" by (auto simp: inner_add_left) } then have "box a b \ {}" using mem_box(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x assume i: "i \ Basis" and as:"b\i < a\i" and x:"x\cbox a b" then have "a \ i \ x \ i \ x \ i \ b \ i" unfolding mem_box by auto then have "a\i \ b\i" by auto then have False using as by auto } moreover { assume as:"\i\Basis. \ (b\i < a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i:"i \ Basis" have "a\i \ b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i \ ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i \ b\i" by (auto simp: inner_add_left) } then have "cbox a b \ {}" using mem_box(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma box_ne_empty: fixes a :: "'a::euclidean_space" shows "cbox a b \ {} \ (\i\Basis. a\i \ b\i)" and "box a b \ {} \ (\i\Basis. a\i < b\i)" unfolding box_eq_empty[of a b] by fastforce+ lemma fixes a :: "'a::euclidean_space" shows cbox_sing [simp]: "cbox a a = {a}" and box_sing [simp]: "box a a = {}" unfolding set_eq_iff mem_box eq_iff [symmetric] by (auto intro!: euclidean_eqI[where 'a='a]) (metis all_not_in_conv nonempty_Basis) lemma subset_box_imp: fixes a :: "'a::euclidean_space" shows "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ cbox c d \ cbox a b" and "(\i\Basis. a\i < c\i \ d\i < b\i) \ cbox c d \ box a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ cbox a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ box a b" unfolding subset_eq[unfolded Ball_def] unfolding mem_box by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ lemma box_subset_cbox: fixes a :: "'a::euclidean_space" shows "box a b \ cbox a b" unfolding subset_eq [unfolded Ball_def] mem_box by (fast intro: less_imp_le) lemma subset_box: fixes a :: "'a::euclidean_space" shows "cbox c d \ cbox a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) and "cbox c d \ box a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) and "box c d \ cbox a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) and "box c d \ box a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th4) proof - let ?lesscd = "\i\Basis. c\i < d\i" let ?lerhs = "\i\Basis. a\i \ c\i \ d\i \ b\i" show ?th1 ?th2 by (fastforce simp: mem_box)+ have acdb: "a\i \ c\i \ d\i \ b\i" if i: "i \ Basis" and box: "box c d \ cbox a b" and cd: "\i. i \ Basis \ c\i < d\i" for i proof - have "box c d \ {}" using that unfolding box_eq_empty by force { let ?x = "(\j\Basis. (if j=i then ((min (a\j) (d\j))+c\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "a\i > c\i" then have "c \ j < ?x \ j \ ?x \ j < d \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "a\i \ c\i" by force moreover { let ?x = "(\j\Basis. (if j=i then ((max (b\j) (c\j))+d\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "b\i < d\i" then have "d \ j > ?x \ j \ ?x \ j > c \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "b\i \ d\i" by (rule ccontr) auto ultimately show ?thesis by auto qed show ?th3 using acdb by (fastforce simp add: mem_box) have acdb': "a\i \ c\i \ d\i \ b\i" if "i \ Basis" "box c d \ box a b" "\i. i \ Basis \ c\i < d\i" for i using box_subset_cbox[of a b] that acdb by auto show ?th4 using acdb' by (fastforce simp add: mem_box) qed lemma eq_cbox: "cbox a b = cbox c d \ cbox a b = {} \ cbox c d = {} \ a = c \ b = d" (is "?lhs = ?rhs") proof assume ?lhs then have "cbox a b \ cbox c d" "cbox c d \ cbox a b" by auto then show ?rhs by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) next assume ?rhs then show ?lhs by force qed lemma eq_cbox_box [simp]: "cbox a b = box c d \ cbox a b = {} \ box c d = {}" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "cbox a b \ box c d" "box c d \ cbox a b" by auto then show ?rhs apply (simp add: subset_box) using L box_ne_empty box_sing apply (fastforce simp add:) done qed force lemma eq_box_cbox [simp]: "box a b = cbox c d \ box a b = {} \ cbox c d = {}" by (metis eq_cbox_box) lemma eq_box: "box a b = box c d \ box a b = {} \ box c d = {} \ a = c \ b = d" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "box a b \ box c d" "box c d \ box a b" by auto then show ?rhs apply (simp add: subset_box) using box_ne_empty(2) L apply auto apply (meson euclidean_eqI less_eq_real_def not_less)+ done qed force lemma subset_box_complex: "cbox a b \ cbox c d \ (Re a \ Re b \ Im a \ Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "cbox a b \ box c d \ (Re a \ Re b \ Im a \ Im b) \ Re a > Re c \ Im a > Im c \ Re b < Re d \ Im b < Im d" "box a b \ cbox c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "box a b \ box c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" by (subst subset_box; force simp: Basis_complex_def)+ lemma in_cbox_complex_iff: "x \ cbox a b \ Re x \ {Re a..Re b} \ Im x \ {Im a..Im b}" by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq) +lemma cbox_complex_of_real: "cbox (complex_of_real x) (complex_of_real y) = complex_of_real ` {x..y}" +proof - + have "(x \ Re z \ Re z \ y \ Im z = 0) = (z \ complex_of_real ` {x..y})" for z + by (cases z) (simp add: complex_eq_cancel_iff2 image_iff) + then show ?thesis + by (auto simp: in_cbox_complex_iff) +qed + lemma box_Complex_eq: "box (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (box a b \ box c d)" by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff) lemma in_box_complex_iff: "x \ box a b \ Re x \ {Re a<.. Im x \ {Im a<.. cbox c d = cbox (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma disjoint_interval: fixes a::"'a::euclidean_space" shows "cbox a b \ cbox c d = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) and "cbox a b \ box c d = {} \ (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) and "box a b \ cbox c d = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) and "box a b \ box c d = {} \ (\i\Basis. (b\i \ a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th4) proof - let ?z = "(\i\Basis. (((max (a\i) (c\i)) + (min (b\i) (d\i))) / 2) *\<^sub>R i)::'a" have **: "\P Q. (\i :: 'a. i \ Basis \ Q ?z i \ P i) \ (\i x :: 'a. i \ Basis \ P i \ Q x i) \ (\x. \i\Basis. Q x i) \ (\i\Basis. P i)" by blast note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) show ?th1 unfolding * by (intro **) auto show ?th2 unfolding * by (intro **) auto show ?th3 unfolding * by (intro **) auto show ?th4 unfolding * by (intro **) auto qed lemma UN_box_eq_UNIV: "(\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" proof - have "\x \ b\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" if [simp]: "b \ Basis" for x b :: 'a proof - have "\x \ b\ \ real_of_int \\x \ b\\" by (rule le_of_int_ceiling) also have "\ \ real_of_int \Max ((\b. \x \ b\)`Basis)\" by (auto intro!: ceiling_mono) also have "\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" by simp finally show ?thesis . qed then have "\n::nat. \b\Basis. \x \ b\ < real n" for x :: 'a by (metis order.strict_trans reals_Archimedean2) moreover have "\x b::'a. \n::nat. \x \ b\ < real n \ - real n < x \ b \ x \ b < real n" by auto ultimately show ?thesis by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) qed lemma image_affinity_cbox: fixes m::real fixes a b c :: "'a::euclidean_space" shows "(\x. m *\<^sub>R x + c) ` cbox a b = (if cbox a b = {} then {} else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" proof (cases "m = 0") case True { fix x assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" then have "x = c" by (simp add: dual_order.antisym euclidean_eqI) } moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" unfolding True by (auto) ultimately show ?thesis using True by (auto simp: cbox_def) next case False { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" by (auto simp: inner_distrib) } moreover { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" by (auto simp: mult_left_mono_neg inner_distrib) } moreover { fix y assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) done } moreover { fix y assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) done } ultimately show ?thesis using False by (auto simp: cbox_def) qed lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" using image_affinity_cbox[of m 0 a b] by auto lemma swap_continuous: assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "continuous_on (cbox (c,a) (d,b)) (\(x, y). f y x)" proof - have "(\(x, y). f y x) = (\(x, y). f x y) \ prod.swap" by auto then show ?thesis apply (rule ssubst) apply (rule continuous_on_compose) apply (simp add: split_def) apply (rule continuous_intros | simp add: assms)+ done qed subsection \General Intervals\ definition\<^marker>\tag important\ "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" lemma is_interval_1: "is_interval (s::real set) \ (\a\s. \b\s. \ x. a \ x \ x \ b \ x \ s)" unfolding is_interval_def by auto lemma is_interval_Int: "is_interval X \ is_interval Y \ is_interval (X \ Y)" unfolding is_interval_def by blast lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff by (meson order_trans le_less_trans less_le_trans less_trans)+ lemma is_interval_empty [iff]: "is_interval {}" unfolding is_interval_def by simp lemma is_interval_univ [iff]: "is_interval UNIV" unfolding is_interval_def by simp lemma mem_is_intervalI: assumes "is_interval s" and "a \ s" "b \ s" and "\i. i \ Basis \ a \ i \ x \ i \ x \ i \ b \ i \ b \ i \ x \ i \ x \ i \ a \ i" shows "x \ s" by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) lemma interval_subst: fixes S::"'a::euclidean_space set" assumes "is_interval S" and "x \ S" "y j \ S" and "j \ Basis" shows "(\i\Basis. (if i = j then y i \ i else x \ i) *\<^sub>R i) \ S" by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) lemma mem_box_componentwiseI: fixes S::"'a::euclidean_space set" assumes "is_interval S" assumes "\i. i \ Basis \ x \ i \ ((\x. x \ i) ` S)" shows "x \ S" proof - from assms have "\i \ Basis. \s \ S. x \ i = s \ i" by auto with finite_Basis obtain s and bs::"'a list" where s: "\i. i \ Basis \ x \ i = s i \ i" "\i. i \ Basis \ s i \ S" and bs: "set bs = Basis" "distinct bs" by (metis finite_distinct_list) from nonempty_Basis s obtain j where j: "j \ Basis" "s j \ S" by blast define y where "y = rec_list (s j) (\j _ Y. (\i\Basis. (if i = j then s i \ i else Y \ i) *\<^sub>R i))" have "x = (\i\Basis. (if i \ set bs then s i \ i else s j \ i) *\<^sub>R i)" using bs by (auto simp: s(1)[symmetric] euclidean_representation) also have [symmetric]: "y bs = \" using bs(2) bs(1)[THEN equalityD1] by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) also have "y bs \ S" using bs(1)[THEN equalityD1] apply (induct bs) apply (auto simp: y_def j) apply (rule interval_subst[OF assms(1)]) apply (auto simp: s) done finally show ?thesis . qed lemma cbox01_nonempty [simp]: "cbox 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) lemma box01_nonempty [simp]: "box 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left) lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast lemma interval_subset_is_interval: assumes "is_interval S" shows "cbox a b \ S \ cbox a b = {} \ a \ S \ b \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce next assume ?rhs have "cbox a b \ S" if "a \ S" "b \ S" using assms unfolding is_interval_def apply (clarsimp simp add: mem_box) using that by blast with \?rhs\ show ?lhs by blast qed lemma is_real_interval_union: "is_interval (X \ Y)" if X: "is_interval X" and Y: "is_interval Y" and I: "(X \ {} \ Y \ {} \ X \ Y \ {})" for X Y::"real set" proof - consider "X \ {}" "Y \ {}" | "X = {}" | "Y = {}" by blast then show ?thesis proof cases case 1 then obtain r where "r \ X \ X \ Y = {}" "r \ Y \ X \ Y = {}" by blast then show ?thesis using I 1 X Y unfolding is_interval_1 by (metis (full_types) Un_iff le_cases) qed (use that in auto) qed lemma is_interval_translationI: assumes "is_interval X" shows "is_interval ((+) x ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (x + b) \ i \ e \ i \ e \ i \ (x + d) \ i \ (x + d) \ i \ e \ i \ e \ i \ (x + b) \ i" hence "e - x \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "e - x"]) (auto simp: algebra_simps) thus "e \ (+) x ` X" by force qed lemma is_interval_uminusI: assumes "is_interval X" shows "is_interval (uminus ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (- b) \ i \ e \ i \ e \ i \ (- d) \ i \ (- d) \ i \ e \ i \ e \ i \ (- b) \ i" hence "- e \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "- e"]) (auto simp: algebra_simps) thus "e \ uminus ` X" by force qed lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] by (auto simp: image_image) lemma is_interval_neg_translationI: assumes "is_interval X" shows "is_interval ((-) x ` X)" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \" by (metis is_interval_uminusI is_interval_translationI assms) finally show ?thesis . qed lemma is_interval_translation[simp]: "is_interval ((+) x ` X) = is_interval X" using is_interval_neg_translationI[of "(+) x ` X" x] by (auto intro!: is_interval_translationI simp: image_image) lemma is_interval_minus_translation[simp]: shows "is_interval ((-) x ` X) = is_interval X" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \ = is_interval X" by simp finally show ?thesis . qed lemma is_interval_minus_translation'[simp]: shows "is_interval ((\x. x - c) ` X) = is_interval X" using is_interval_translation[of "-c" X] by (metis image_cong uminus_add_conv_diff) lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real by (simp add: cball_eq_atLeastAtMost is_interval_def) lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real by (simp add: ball_eq_greaterThanLessThan is_interval_def) subsection\<^marker>\tag unimportant\ \Bounded Projections\ lemma bounded_inner_imp_bdd_above: assumes "bounded s" shows "bdd_above ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) lemma bounded_inner_imp_bdd_below: assumes "bounded s" shows "bdd_below ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) subsection\<^marker>\tag unimportant\ \Structural rules for pointwise continuity\ lemma continuous_infnorm[continuous_intros]: "continuous F f \ continuous F (\x. infnorm (f x))" unfolding continuous_def by (rule tendsto_infnorm) lemma continuous_inner[continuous_intros]: assumes "continuous F f" and "continuous F g" shows "continuous F (\x. inner (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_inner) subsection\<^marker>\tag unimportant\ \Structural rules for setwise continuity\ lemma continuous_on_infnorm[continuous_intros]: "continuous_on s f \ continuous_on s (\x. infnorm (f x))" unfolding continuous_on by (fast intro: tendsto_infnorm) lemma continuous_on_inner[continuous_intros]: fixes g :: "'a::topological_space \ 'b::real_inner" assumes "continuous_on s f" and "continuous_on s g" shows "continuous_on s (\x. inner (f x) (g x))" using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) subsection\<^marker>\tag unimportant\ \Openness of halfspaces.\ lemma open_halfspace_lt: "open {x. inner a x < b}" by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_gt: "open {x. inner a x > b}" by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\i < a}" by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\i > a}" by (simp add: open_Collect_less continuous_on_inner) lemma eucl_less_eq_halfspaces: fixes a :: "'a::euclidean_space" shows "{x. x i\Basis. {x. x \ i < a \ i})" "{x. a i\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less_def) lemma open_Collect_eucl_less[simp, intro]: fixes a :: "'a::euclidean_space" shows "open {x. x \tag unimportant\ \Closure and Interior of halfspaces and hyperplanes\ lemma continuous_at_inner: "continuous (at x) (inner a)" unfolding continuous_at by (intro tendsto_intros) lemma closed_halfspace_le: "closed {x. inner a x \ b}" by (simp add: closed_Collect_le continuous_on_inner) lemma closed_halfspace_ge: "closed {x. inner a x \ b}" by (simp add: closed_Collect_le continuous_on_inner) lemma closed_hyperplane: "closed {x. inner a x = b}" by (simp add: closed_Collect_eq continuous_on_inner) lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le continuous_on_inner) lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le continuous_on_inner) lemma closed_interval_left: fixes b :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. x\i \ b\i}" by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) lemma closed_interval_right: fixes a :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. a\i \ x\i}" by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) lemma interior_halfspace_le [simp]: assumes "a \ 0" shows "interior {x. a \ x \ b} = {x. a \ x < b}" proof - have *: "a \ x < b" if x: "x \ S" and S: "S \ {x. a \ x \ b}" and "open S" for S x proof - obtain e where "e>0" and e: "cball x e \ S" using \open S\ open_contains_cball x by blast then have "x + (e / norm a) *\<^sub>R a \ cball x e" by (simp add: dist_norm) then have "x + (e / norm a) *\<^sub>R a \ S" using e by blast then have "x + (e / norm a) *\<^sub>R a \ {x. a \ x \ b}" using S by blast moreover have "e * (a \ a) / norm a > 0" by (simp add: \0 < e\ assms) ultimately show ?thesis by (simp add: algebra_simps) qed show ?thesis by (rule interior_unique) (auto simp: open_halfspace_lt *) qed lemma interior_halfspace_ge [simp]: "a \ 0 \ interior {x. a \ x \ b} = {x. a \ x > b}" using interior_halfspace_le [of "-a" "-b"] by simp lemma closure_halfspace_lt [simp]: assumes "a \ 0" shows "closure {x. a \ x < b} = {x. a \ x \ b}" proof - have [simp]: "-{x. a \ x < b} = {x. a \ x \ b}" by (force simp:) then show ?thesis using interior_halfspace_ge [of a b] assms by (force simp: closure_interior) qed lemma closure_halfspace_gt [simp]: "a \ 0 \ closure {x. a \ x > b} = {x. a \ x \ b}" using closure_halfspace_lt [of "-a" "-b"] by simp lemma interior_hyperplane [simp]: assumes "a \ 0" shows "interior {x. a \ x = b} = {}" proof - have [simp]: "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by (force simp:) then show ?thesis by (auto simp: assms) qed lemma frontier_halfspace_le: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_le) qed lemma frontier_halfspace_ge: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_ge) qed lemma frontier_halfspace_lt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x < b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_lt) qed lemma frontier_halfspace_gt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x > b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_gt) qed subsection\<^marker>\tag unimportant\\Some more convenient intermediate-value theorem formulations\ lemma connected_ivt_hyperplane: assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" shows "\z \ S. inner a z = b" proof (rule ccontr) assume as:"\ (\z\S. inner a z = b)" let ?A = "{x. inner a x < b}" let ?B = "{x. inner a x > b}" have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto moreover have "?A \ ?B = {}" by auto moreover have "S \ ?A \ ?B" using as by auto ultimately show False using \connected S\[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] using xy b by auto qed lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows "connected S \ x \ S \ y \ S \ x\k \ a \ a \ y\k \ (\z\S. z\k = a)" using connected_ivt_hyperplane[of S x y "k::'a" a] by (auto simp: inner_commute) subsection \Limit Component Bounds\ lemma Lim_component_le: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. f(x)\i \ b) net" shows "l\i \ b" by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) lemma Lim_component_ge: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. b \ (f x)\i) net" shows "b \ l\i" by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) lemma Lim_component_eq: fixes f :: "'a \ 'b::euclidean_space" assumes net: "(f \ l) net" "\ trivial_limit net" and ev:"eventually (\x. f(x)\i = b) net" shows "l\i = b" using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] using Lim_component_le[OF net, of i b] by auto lemma open_box[intro]: "open (box a b)" proof - have "open (\i\Basis. ((\) i) -` {a \ i <..< b \ i})" by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) also have "(\i\Basis. ((\) i) -` {a \ i <..< b \ i}) = box a b" by (auto simp: box_def inner_commute) finally show ?thesis . qed lemma closed_cbox[intro]: fixes a b :: "'a::euclidean_space" shows "closed (cbox a b)" proof - have "closed (\i\Basis. (\x. x\i) -` {a\i .. b\i})" by (intro closed_INT ballI continuous_closed_vimage allI linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = cbox a b" by (auto simp: cbox_def) finally show "closed (cbox a b)" . qed lemma interior_cbox [simp]: fixes a b :: "'a::euclidean_space" shows "interior (cbox a b) = box a b" (is "?L = ?R") proof(rule subset_antisym) show "?R \ ?L" using box_subset_cbox open_box by (rule interior_maximal) { fix x assume "x \ interior (cbox a b)" then obtain s where s: "open s" "x \ s" "s \ cbox a b" .. then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ cbox a b" unfolding open_dist and subset_eq by auto { fix i :: 'a assume i: "i \ Basis" have "dist (x - (e / 2) *\<^sub>R i) x < e" and "dist (x + (e / 2) *\<^sub>R i) x < e" unfolding dist_norm apply auto unfolding norm_minus_cancel using norm_Basis[OF i] \e>0\ apply auto done then have "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" and "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] unfolding mem_box using i by blast+ then have "a \ i < x \ i" and "x \ i < b \ i" using \e>0\ i by (auto simp: inner_diff_left inner_Basis inner_add_left) } then have "x \ box a b" unfolding mem_box by auto } then show "?L \ ?R" .. qed lemma bounded_cbox [simp]: fixes a :: "'a::euclidean_space" shows "bounded (cbox a b)" proof - let ?b = "\i\Basis. \a\i\ + \b\i\" { fix x :: "'a" assume "\i. i\Basis \ a \ i \ x \ i \ x \ i \ b \ i" then have "(\i\Basis. \x \ i\) \ ?b" by (force simp: intro!: sum_mono) then have "norm x \ ?b" using norm_le_l1[of x] by auto } then show ?thesis unfolding cbox_def bounded_iff by force qed lemma bounded_box [simp]: fixes a :: "'a::euclidean_space" shows "bounded (box a b)" using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] by simp lemma not_interval_UNIV [simp]: fixes a :: "'a::euclidean_space" shows "cbox a b \ UNIV" "box a b \ UNIV" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma not_interval_UNIV2 [simp]: fixes a :: "'a::euclidean_space" shows "UNIV \ cbox a b" "UNIV \ box a b" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma box_midpoint: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "((1/2) *\<^sub>R (a + b)) \ box a b" proof - have "a \ i < ((1 / 2) *\<^sub>R (a + b)) \ i \ ((1 / 2) *\<^sub>R (a + b)) \ i < b \ i" if "i \ Basis" for i using assms that by (auto simp: inner_add_left box_ne_empty) then show ?thesis unfolding mem_box by auto qed lemma open_cbox_convex: fixes x :: "'a::euclidean_space" assumes x: "x \ box a b" and y: "y \ cbox a b" and e: "0 < e" "e \ 1" shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ box a b" proof - { fix i :: 'a assume i: "i \ Basis" have "a \ i = e * (a \ i) + (1 - e) * (a \ i)" unfolding left_diff_distrib by simp also have "\ < e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (a \ i) < e * (x \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (a \ i) \ (1 - e) * (y \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" unfolding inner_simps by auto moreover { have "b \ i = e * (b\i) + (1 - e) * (b\i)" unfolding left_diff_distrib by simp also have "\ > e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (x \ i) < e * (b \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (y \ i) \ (1 - e) * (b \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" unfolding inner_simps by auto } ultimately have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" by auto } then show ?thesis unfolding mem_box by auto qed lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" by (simp add: closed_cbox) lemma closure_box [simp]: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "closure (box a b) = cbox a b" proof - have ab: "a cbox a b" define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n { fix n assume fn: "f n a f n = x" and xc: "x \ ?c" have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \ 1" unfolding inverse_le_1_iff by auto have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" by (auto simp: algebra_simps) then have "f n (f \ x) sequentially" { fix e :: real assume "e > 0" then obtain N :: nat where N: "inverse (real (N + 1)) < e" using reals_Archimedean by auto have "inverse (real n + 1) < e" if "N \ n" for n by (auto intro!: that le_less_trans [OF _ N]) then have "\N::nat. \n\N. inverse (real n + 1) < e" by auto } then have "((\n. inverse (real n + 1)) \ 0) sequentially" unfolding lim_sequentially by(auto simp: dist_norm) then have "(f \ x) sequentially" unfolding f_def using tendsto_add[OF tendsto_const, of "\n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] using tendsto_scaleR [OF _ tendsto_const, of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto } ultimately have "x \ closure (box a b)" using as box_midpoint[OF assms] unfolding closure_def islimpt_sequential by (cases "x=?c") (auto simp: in_box_eucl_less) } then show ?thesis using closure_minimal[OF box_subset_cbox, of a b] by blast qed lemma bounded_subset_box_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ box (-a) a" proof - obtain b where "b>0" and b: "\x\S. norm x \ b" using assms[unfolded bounded_pos] by auto define a :: 'a where "a = (\i\Basis. (b + 1) *\<^sub>R i)" have "(-a)\i < x\i" and "x\i < a\i" if "x \ S" and i: "i \ Basis" for x i using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) then have "S \ box (-a) a" by (auto simp: simp add: box_def) then show ?thesis .. qed lemma bounded_subset_cbox_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ cbox (-a) a" proof - obtain a where "S \ box (-a) a" using bounded_subset_box_symmetric[OF assms] by auto then show ?thesis by (meson box_subset_cbox dual_order.trans that) qed lemma frontier_cbox: fixes a b :: "'a::euclidean_space" shows "frontier (cbox a b) = cbox a b - box a b" unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. lemma frontier_box: fixes a b :: "'a::euclidean_space" shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" proof (cases "box a b = {}") case True then show ?thesis using frontier_empty by auto next case False then show ?thesis unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] by auto qed lemma Int_interval_mixed_eq_empty: fixes a :: "'a::euclidean_space" assumes "box c d \ {}" shows "box a b \ cbox c d = {} \ box a b \ box c d = {}" unfolding closure_box[OF assms, symmetric] unfolding open_Int_closure_eq_empty[OF open_box] .. subsection \Class Instances\ lemma compact_lemma: fixes f :: "nat \ 'a::euclidean_space" assumes "bounded (range f)" shows "\d\Basis. \l::'a. \ r. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) \ i) (l \ i) < e) sequentially)" by (rule compact_lemma_general[where unproj="\e. \i\Basis. e i *\<^sub>R i"]) (auto intro!: assms bounded_linear_inner_left bounded_linear_image simp: euclidean_representation) instance\<^marker>\tag important\ euclidean_space \ heine_borel proof fix f :: "nat \ 'a" assume f: "bounded (range f)" then obtain l::'a and r where r: "strict_mono r" and l: "\e>0. eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e) sequentially" using compact_lemma [OF f] by blast { fix e::real assume "e > 0" hence "e / real_of_nat DIM('a) > 0" by (simp) with l have "eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))) sequentially" by simp moreover { fix n assume n: "\i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))" have "dist (f (r n)) l \ (\i\Basis. dist (f (r n) \ i) (l \ i))" apply (subst euclidean_dist_l2) using zero_le_dist apply (rule L2_set_le_sum) done also have "\ < (\i\(Basis::'a set). e / (real_of_nat DIM('a)))" apply (rule sum_strict_mono) using n apply auto done finally have "dist (f (r n)) l < e" by auto } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_mono) } then have *: "((f \ r) \ l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" by auto qed instance\<^marker>\tag important\ euclidean_space \ banach .. instance euclidean_space \ second_countable_topology proof define a where "a f = (\i\Basis. fst (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have a: "\f. (\i\Basis. fst (f i) *\<^sub>R i) = a f" by simp define b where "b f = (\i\Basis. snd (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have b: "\f. (\i\Basis. snd (f i) *\<^sub>R i) = b f" by simp define B where "B = (\f. box (a f) (b f)) ` (Basis \\<^sub>E (\ \ \))" have "Ball B open" by (simp add: B_def open_box) moreover have "(\A. open A \ (\B'\B. \B' = A))" proof safe fix A::"'a set" assume "open A" show "\B'\B. \B' = A" apply (rule exI[of _ "{b\B. b \ A}"]) apply (subst (3) open_UNION_box[OF \open A\]) apply (auto simp: a b B_def) done qed ultimately have "topological_basis B" unfolding topological_basis_def by blast moreover have "countable B" unfolding B_def by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) ultimately show "\B::'a set set. countable B \ open = generate_topology B" by (blast intro: topological_basis_imp_subbasis) qed instance euclidean_space \ polish_space .. subsection \Compact Boxes\ lemma compact_cbox [simp]: fixes a :: "'a::euclidean_space" shows "compact (cbox a b)" using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] by (auto simp: compact_eq_seq_compact_metric) proposition is_interval_compact: "is_interval S \ compact S \ (\a b. S = cbox a b)" (is "?lhs = ?rhs") proof (cases "S = {}") case True with empty_as_interval show ?thesis by auto next case False show ?thesis proof assume L: ?lhs then have "is_interval S" "compact S" by auto define a where "a \ \i\Basis. (INF x\S. x \ i) *\<^sub>R i" define b where "b \ \i\Basis. (SUP x\S. x \ i) *\<^sub>R i" have 1: "\x i. \x \ S; i \ Basis\ \ (INF x\S. x \ i) \ x \ i" by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) have 2: "\x i. \x \ S; i \ Basis\ \ x \ i \ (SUP x\S. x \ i)" by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) have 3: "x \ S" if inf: "\i. i \ Basis \ (INF x\S. x \ i) \ x \ i" and sup: "\i. i \ Basis \ x \ i \ (SUP x\S. x \ i)" for x proof (rule mem_box_componentwiseI [OF \is_interval S\]) fix i::'a assume i: "i \ Basis" have cont: "continuous_on S (\x. x \ i)" by (intro continuous_intros) obtain a where "a \ S" and a: "\y. y\S \ a \ i \ y \ i" using continuous_attains_inf [OF \compact S\ False cont] by blast obtain b where "b \ S" and b: "\y. y\S \ y \ i \ b \ i" using continuous_attains_sup [OF \compact S\ False cont] by blast have "a \ i \ (INF x\S. x \ i)" by (simp add: False a cINF_greatest) also have "\ \ x \ i" by (simp add: i inf) finally have ai: "a \ i \ x \ i" . have "x \ i \ (SUP x\S. x \ i)" by (simp add: i sup) also have "(SUP x\S. x \ i) \ b \ i" by (simp add: False b cSUP_least) finally have bi: "x \ i \ b \ i" . show "x \ i \ (\x. x \ i) ` S" apply (rule_tac x="\j\Basis. (if j = i then x \ i else a \ j) *\<^sub>R j" in image_eqI) apply (simp add: i) apply (rule mem_is_intervalI [OF \is_interval S\ \a \ S\ \b \ S\]) using i ai bi apply force done qed have "S = cbox a b" by (auto simp: a_def b_def mem_box intro: 1 2 3) then show ?rhs by blast next assume R: ?rhs then show ?lhs using compact_cbox is_interval_cbox by blast qed qed subsection\<^marker>\tag unimportant\\Componentwise limits and continuity\ text\But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\ lemma Euclidean_dist_upper: "i \ Basis \ dist (x \ i) (y \ i) \ dist x y" by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) text\But is the premise \<^term>\i \ Basis\ really necessary?\ lemma open_preimage_inner: assumes "open S" "i \ Basis" shows "open {x. x \ i \ S}" proof (rule openI, simp) fix x assume x: "x \ i \ S" with assms obtain e where "0 < e" and e: "ball (x \ i) e \ S" by (auto simp: open_contains_ball_eq) have "\e>0. ball (y \ i) e \ S" if dxy: "dist x y < e / 2" for y proof (intro exI conjI) have "dist (x \ i) (y \ i) < e / 2" by (meson \i \ Basis\ dual_order.trans Euclidean_dist_upper not_le that) then have "dist (x \ i) z < e" if "dist (y \ i) z < e / 2" for z by (metis dist_commute dist_triangle_half_l that) then have "ball (y \ i) (e / 2) \ ball (x \ i) e" using mem_ball by blast with e show "ball (y \ i) (e / 2) \ S" by (metis order_trans) qed (simp add: \0 < e\) then show "\e>0. ball x e \ {s. s \ i \ S}" by (metis (no_types, lifting) \0 < e\ \open S\ half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) qed proposition tendsto_componentwise_iff: fixes f :: "_ \ 'b::euclidean_space" shows "(f \ l) F \ (\i \ Basis. ((\x. (f x \ i)) \ (l \ i)) F)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding tendsto_def apply clarify apply (drule_tac x="{s. s \ i \ S}" in spec) apply (auto simp: open_preimage_inner) done next assume R: ?rhs then have "\e. e > 0 \ \i\Basis. \\<^sub>F x in F. dist (f x \ i) (l \ i) < e" unfolding tendsto_iff by blast then have R': "\e. e > 0 \ \\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e" by (simp add: eventually_ball_finite_distrib [symmetric]) show ?lhs unfolding tendsto_iff proof clarify fix e::real assume "0 < e" have *: "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" if "\i\Basis. dist (f x \ i) (l \ i) < e / real DIM('b)" for x proof - have "L2_set (\i. dist (f x \ i) (l \ i)) Basis \ sum (\i. dist (f x \ i) (l \ i)) Basis" by (simp add: L2_set_le_sum) also have "... < DIM('b) * (e / real DIM('b))" apply (rule sum_bounded_above_strict) using that by auto also have "... = e" by (simp add: field_simps) finally show "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" . qed have "\\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e / DIM('b)" apply (rule R') using \0 < e\ by simp then show "\\<^sub>F x in F. dist (f x) l < e" apply (rule eventually_mono) apply (subst euclidean_dist_l2) using * by blast qed qed corollary continuous_componentwise: "continuous F f \ (\i \ Basis. continuous F (\x. (f x \ i)))" by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) corollary continuous_on_componentwise: fixes S :: "'a :: t2_space set" shows "continuous_on S f \ (\i \ Basis. continuous_on S (\x. (f x \ i)))" apply (simp add: continuous_on_eq_continuous_within) using continuous_componentwise by blast lemma linear_componentwise_iff: "(linear f') \ (\i\Basis. linear (\x. f' x \ i))" apply (auto simp: linear_iff inner_left_distrib) apply (metis inner_left_distrib euclidean_eq_iff) by (metis euclidean_eqI inner_scaleR_left) lemma bounded_linear_componentwise_iff: "(bounded_linear f') \ (\i\Basis. bounded_linear (\x. f' x \ i))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: bounded_linear_inner_left_comp) next assume ?rhs then have "(\i\Basis. \K. \x. \f' x \ i\ \ norm x * K)" "linear f'" by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) then obtain F where F: "\i x. i \ Basis \ \f' x \ i\ \ norm x * F i" by metis have "norm (f' x) \ norm x * sum F Basis" for x proof - have "norm (f' x) \ (\i\Basis. \f' x \ i\)" by (rule norm_le_l1) also have "... \ (\i\Basis. norm x * F i)" by (metis F sum_mono) also have "... = norm x * sum F Basis" by (simp add: sum_distrib_left) finally show ?thesis . qed then show ?lhs by (force simp: bounded_linear_def bounded_linear_axioms_def \linear f'\) qed subsection\<^marker>\tag unimportant\ \Continuous Extension\ definition clamp :: "'a::euclidean_space \ 'a \ 'a \ 'a" where "clamp a b x = (if (\i\Basis. a \ i \ b \ i) then (\i\Basis. (if x\i < a\i then a\i else if x\i \ b\i then x\i else b\i) *\<^sub>R i) else a)" lemma clamp_in_interval[simp]: assumes "\i. i \ Basis \ a \ i \ b \ i" shows "clamp a b x \ cbox a b" unfolding clamp_def using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) lemma clamp_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "clamp a b x = x" using assms by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) lemma clamp_empty_interval: assumes "i \ Basis" "a \ i > b \ i" shows "clamp a b = (\_. a)" using assms by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) lemma dist_clamps_le_dist_args: fixes x :: "'a::euclidean_space" shows "dist (clamp a b y) (clamp a b x) \ dist y x" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" then have "(\i\Basis. (dist (clamp a b y \ i) (clamp a b x \ i))\<^sup>2) \ (\i\Basis. (dist (y \ i) (x \ i))\<^sup>2)" by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) then show ?thesis by (auto intro: real_sqrt_le_mono simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) qed (auto simp: clamp_def) lemma clamp_continuous_at: fixes f :: "'a::euclidean_space \ 'b::metric_space" and x :: 'a assumes f_cont: "continuous_on (cbox a b) f" shows "continuous (at x) (\x. f (clamp a b x))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" show ?thesis unfolding continuous_at_eps_delta proof safe fix x :: 'a fix e :: real assume "e > 0" moreover have "clamp a b x \ cbox a b" by (simp add: le) moreover note f_cont[simplified continuous_on_iff] ultimately obtain d where d: "0 < d" "\x'. x' \ cbox a b \ dist x' (clamp a b x) < d \ dist (f x') (f (clamp a b x)) < e" by force show "\d>0. \x'. dist x' x < d \ dist (f (clamp a b x')) (f (clamp a b x)) < e" using le by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) qed qed (auto simp: clamp_empty_interval) lemma clamp_continuous_on: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes f_cont: "continuous_on (cbox a b) f" shows "continuous_on S (\x. f (clamp a b x))" using assms by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) lemma clamp_bounded: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes bounded: "bounded (f ` (cbox a b))" shows "bounded (range (\x. f (clamp a b x)))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" from bounded obtain c where f_bound: "\x\f ` cbox a b. dist undefined x \ c" by (auto simp: bounded_any_center[where a=undefined]) then show ?thesis by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] simp: bounded_any_center[where a=undefined]) qed (auto simp: clamp_empty_interval image_def) definition ext_cont :: "('a::euclidean_space \ 'b::metric_space) \ 'a \ 'a \ 'a \ 'b" where "ext_cont f a b = (\x. f (clamp a b x))" lemma ext_cont_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "ext_cont f a b x = f x" using assms unfolding ext_cont_def by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) lemma continuous_on_ext_cont[continuous_intros]: "continuous_on (cbox a b) f \ continuous_on S (ext_cont f a b)" by (auto intro!: clamp_continuous_on simp: ext_cont_def) subsection \Separability\ lemma univ_second_countable_sequence: obtains B :: "nat \ 'a::euclidean_space set" where "inj B" "\n. open(B n)" "\S. open S \ \k. S = \{B n |n. n \ k}" proof - obtain \ :: "'a set set" where "countable \" and opn: "\C. C \ \ \ open C" and Un: "\S. open S \ \U. U \ \ \ S = \U" using univ_second_countable by blast have *: "infinite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule Infinite_Set.range_inj_infinite) apply (simp add: inj_on_def ball_eq_ball_iff) done have "infinite \" proof assume "finite \" then have "finite (Union ` (Pow \))" by simp then have "finite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule rev_finite_subset) by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) with * show False by simp qed obtain f :: "nat \ 'a set" where "\ = range f" "inj f" by (blast intro: countable_as_injective_image [OF \countable \\ \infinite \\]) have *: "\k. S = \{f n |n. n \ k}" if "open S" for S using Un [OF that] apply clarify apply (rule_tac x="f-`U" in exI) using \inj f\ \\ = range f\ apply force done show ?thesis apply (rule that [OF \inj f\ _ *]) apply (auto simp: \\ = range f\ opn) done qed proposition separable: fixes S :: "'a::{metric_space, second_countable_topology} set" obtains T where "countable T" "T \ S" "S \ closure T" proof - obtain \ :: "'a set set" where "countable \" and "{} \ \" and ope: "\C. C \ \ \ openin(top_of_set S) C" and if_ope: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" by (meson subset_second_countable) then obtain f where f: "\C. C \ \ \ f C \ C" by (metis equals0I) show ?thesis proof show "countable (f ` \)" by (simp add: \countable \\) show "f ` \ \ S" using ope f openin_imp_subset by blast show "S \ closure (f ` \)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" have "openin (top_of_set S) (S \ ball x e)" by (simp add: openin_Int_open) with if_ope obtain \ where \: "\ \ \" "S \ ball x e = \\" by meson show "\C \ \. dist (f C) x < e" proof (cases "\ = {}") case True then show ?thesis using \0 < e\ \ \x \ S\ by auto next case False then obtain C where "C \ \" by blast show ?thesis proof show "dist (f C) x < e" by (metis Int_iff Union_iff \ \C \ \\ dist_commute f mem_ball subsetCE) show "C \ \" using \\ \ \\ \C \ \\ by blast qed qed qed qed qed subsection\<^marker>\tag unimportant\ \Diameter\ lemma diameter_cball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(cball a r) = 2*r" if "r \ 0" proof (rule order_antisym) show "diameter (cball a r) \ 2*r" proof (rule diameter_le) fix x y assume "x \ cball a r" "y \ cball a r" then have "norm (x - a) \ r" "norm (a - y) \ r" by (auto simp: dist_norm norm_minus_commute) then have "norm (x - y) \ r+r" using norm_diff_triangle_le by blast then show "norm (x - y) \ 2*r" by simp qed (simp add: that) have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" apply (simp add: dist_norm) by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) also have "... \ diameter (cball a r)" apply (rule diameter_bounded_bound) using that by (auto simp: dist_norm) finally show "2*r \ diameter (cball a r)" . qed then show ?thesis by simp qed lemma diameter_ball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(ball a r) = 2*r" if "r > 0" by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) then show ?thesis by (simp add: diameter_def) qed lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" proof - have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) then show ?thesis by simp qed lemma diameter_open_interval [simp]: "diameter {a<..i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) subsection\<^marker>\tag unimportant\\Relating linear images to open/closed/interior/closure/connected\ proposition open_surjective_linear_image: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open A" "linear f" "surj f" shows "open(f ` A)" unfolding open_dist proof clarify fix x assume "x \ A" have "bounded (inv f ` Basis)" by (simp add: finite_imp_bounded) with bounded_pos obtain B where "B > 0" and B: "\x. x \ inv f ` Basis \ norm x \ B" by metis obtain e where "e > 0" and e: "\z. dist z x < e \ z \ A" by (metis open_dist \x \ A\ \open A\) define \ where "\ \ e / B / DIM('b)" show "\e>0. \y. dist y (f x) < e \ y \ f ` A" proof (intro exI conjI) show "\ > 0" using \e > 0\ \B > 0\ by (simp add: \_def field_split_simps) have "y \ f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y proof - define u where "u \ y - f x" show ?thesis proof (rule image_eqI) show "y = f (x + (\i\Basis. (u \ i) *\<^sub>R inv f i))" apply (simp add: linear_add linear_sum linear.scaleR \linear f\ surj_f_inv_f \surj f\) apply (simp add: euclidean_representation u_def) done have "dist (x + (\i\Basis. (u \ i) *\<^sub>R inv f i)) x \ (\i\Basis. norm ((u \ i) *\<^sub>R inv f i))" by (simp add: dist_norm sum_norm_le) also have "... = (\i\Basis. \u \ i\ * norm (inv f i))" by simp also have "... \ (\i\Basis. \u \ i\) * B" by (simp add: B sum_distrib_right sum_mono mult_left_mono) also have "... \ DIM('b) * dist y (f x) * B" apply (rule mult_right_mono [OF sum_bounded_above]) using \0 < B\ by (auto simp: Basis_le_norm dist_norm u_def) also have "... < e" by (metis mult.commute mult.left_commute that) finally show "x + (\i\Basis. (u \ i) *\<^sub>R inv f i) \ A" by (rule e) qed qed then show "\y. dist y (f x) < \ \ y \ f ` A" using \e > 0\ \B > 0\ by (auto simp: \_def field_split_simps) qed qed corollary open_bijective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "open(f ` A) \ open A" proof assume "open(f ` A)" then have "open(f -` (f ` A))" using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) then show "open A" by (simp add: assms bij_is_inj inj_vimage_image_eq) next assume "open A" then show "open(f ` A)" by (simp add: assms bij_is_surj open_surjective_linear_image) qed corollary interior_bijective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") proof safe fix x assume x: "x \ ?lhs" then obtain T where "open T" and "x \ T" and "T \ f ` S" by (metis interiorE) then show "x \ ?rhs" by (metis (no_types, opaque_lifting) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) next fix x assume x: "x \ interior S" then show "f x \ interior (f ` S)" by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) qed lemma interior_injective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "inj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) lemma interior_surjective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "surj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) lemma interior_negations: fixes S :: "'a::euclidean_space set" shows "interior(uminus ` S) = image uminus (interior S)" by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) lemma connected_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "linear f" and "connected s" shows "connected (f ` s)" using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast subsection\<^marker>\tag unimportant\ \"Isometry" (up to constant bounds) of Injective Linear Map\ proposition injective_imp_isometric: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes s: "closed s" "subspace s" and f: "bounded_linear f" "\x\s. f x = 0 \ x = 0" shows "\e>0. \x\s. norm (f x) \ e * norm x" proof (cases "s \ {0::'a}") case True have "norm x \ norm (f x)" if "x \ s" for x proof - from True that have "x = 0" by auto then show ?thesis by simp qed then show ?thesis by (auto intro!: exI[where x=1]) next case False interpret f: bounded_linear f by fact from False obtain a where a: "a \ 0" "a \ s" by auto from False have "s \ {}" by auto let ?S = "{f x| x. x \ s \ norm x = norm a}" let ?S' = "{x::'a. x\s \ norm x = norm a}" let ?S'' = "{x::'a. norm x = norm a}" have "?S'' = frontier (cball 0 (norm a))" by (simp add: sphere_def dist_norm) then have "compact ?S''" by (metis compact_cball compact_frontier) moreover have "?S' = s \ ?S''" by auto ultimately have "compact ?S'" using closed_Int_compact[of s ?S''] using s(1) by auto moreover have *:"f ` ?S' = ?S" by auto ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto then have "closed ?S" using compact_imp_closed by auto moreover from a have "?S \ {}" by auto ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto then obtain b where "b\s" and ba: "norm b = norm a" and b: "\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[symmetric] unfolding image_iff by auto let ?e = "norm (f b) / norm b" have "norm b > 0" using ba and a and norm_ge_zero by auto moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF \b\s\] using \norm b >0\ by simp ultimately have "0 < norm (f b) / norm b" by simp moreover have "norm (f b) / norm b * norm x \ norm (f x)" if "x\s" for x proof (cases "x = 0") case True then show "norm (f b) / norm b * norm x \ norm (f x)" by auto next case False with \a \ 0\ have *: "0 < norm a / norm x" unfolding zero_less_norm_iff[symmetric] by simp have "\x\s. c *\<^sub>R x \ s" for c using s[unfolded subspace_def] by simp with \x \ s\ \x \ 0\ have "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" by simp with \x \ 0\ \a \ 0\ show "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] unfolding f.scaleR and ba by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) qed ultimately show ?thesis by auto qed proposition closed_injective_image_subspace: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 \ x = 0" "closed s" shows "closed(f ` s)" proof - obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto show ?thesis using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) unfolding complete_eq_closed[symmetric] by auto qed lemma closure_bounded_linear_image_subset: assumes f: "bounded_linear f" shows "f ` closure S \ closure (f ` S)" using linear_continuous_on [OF f] closed_closure closure_subset by (rule image_closure_subset) lemma closure_linear_image_subset: fixes f :: "'m::euclidean_space \ 'n::real_normed_vector" assumes "linear f" shows "f ` (closure S) \ closure (f ` S)" using assms unfolding linear_conv_bounded_linear by (rule closure_bounded_linear_image_subset) lemma closed_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "closed S" and f: "linear f" "inj f" shows "closed (f ` S)" proof - obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF f] by blast then have confg: "continuous_on (range f) g" using linear_continuous_on linear_conv_bounded_linear by blast have [simp]: "g ` f ` S = S" using g by (simp add: image_comp) have cgf: "closed (g ` f ` S)" by (simp add: \g \ f = id\ S image_comp) have [simp]: "(range f \ g -` S) = f ` S" using g unfolding o_def id_def image_def by auto metis+ show ?thesis proof (rule closedin_closed_trans [of "range f"]) show "closedin (top_of_set (range f)) (f ` S)" using continuous_closedin_preimage [OF confg cgf] by simp show "closed (range f)" apply (rule closed_injective_image_subspace) using f apply (auto simp: linear_linear linear_injective_0) done qed qed lemma closed_injective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" shows "(closed(image f s) \ closed s)" by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) lemma closure_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; inj f\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym) apply (simp add: closure_linear_image_subset) by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) lemma closure_bounded_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; bounded S\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym, simp add: closure_linear_image_subset) apply (rule closure_minimal, simp add: closure_subset image_mono) by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) lemma closure_scaleR: fixes S :: "'a::real_normed_vector set" shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" proof show "((*\<^sub>R) c) ` (closure S) \ closure (((*\<^sub>R) c) ` S)" using bounded_linear_scaleR_right by (rule closure_bounded_linear_image_subset) show "closure (((*\<^sub>R) c) ` S) \ ((*\<^sub>R) c) ` (closure S)" by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) qed subsection\<^marker>\tag unimportant\ \Some properties of a canonical subspace\ lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" (is "closed ?A") proof - let ?D = "{i\Basis. P i}" have "closed (\i\?D. {x::'a. x\i = 0})" by (simp add: closed_INT closed_Collect_eq continuous_on_inner) also have "(\i\?D. {x::'a. x\i = 0}) = ?A" by auto finally show "closed ?A" . qed lemma closed_subspace: fixes s :: "'a::euclidean_space set" assumes "subspace s" shows "closed s" proof - have "dim s \ card (Basis :: 'a set)" using dim_subset_UNIV by auto with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" by auto let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" using dim_substandard[of d] t d assms by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) then obtain f where f: "linear f" "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" by blast interpret f: bounded_linear f using f by (simp add: linear_conv_bounded_linear) have "x \ ?t \ f x = 0 \ x = 0" for x using f.zero d f(3)[THEN inj_onD, of x 0] by auto moreover have "closed ?t" by (rule closed_substandard) moreover have "subspace ?t" by (rule subspace_substandard) ultimately show ?thesis using closed_injective_image_subspace[of ?t f] unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto qed lemma complete_subspace: "subspace s \ complete s" for s :: "'a::euclidean_space set" using complete_eq_closed closed_subspace by auto lemma closed_span [iff]: "closed (span s)" for s :: "'a::euclidean_space set" by (simp add: closed_subspace) lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") for s :: "'a::euclidean_space set" proof - have "?dc \ ?d" using closure_minimal[OF span_superset, of s] using closed_subspace[OF subspace_span, of s] using dim_subset[of "closure s" "span s"] by simp then show ?thesis using dim_subset[OF closure_subset, of s] by simp qed subsection \Set Distance\ lemma setdist_compact_closed: fixes A :: "'a::heine_borel set" assumes A: "compact A" and B: "closed B" and "A \ {}" "B \ {}" shows "\x \ A. \y \ B. dist x y = setdist A B" proof - obtain x where "x \ A" "setdist A B = infdist x B" by (metis A assms(3) setdist_attains_inf setdist_sym) moreover obtain y where"y \ B" "infdist x B = dist x y" using B \B \ {}\ infdist_attains_inf by blast ultimately show ?thesis using \x \ A\ \y \ B\ by auto qed lemma setdist_closed_compact: fixes S :: "'a::heine_borel set" assumes S: "closed S" and T: "compact T" and "S \ {}" "T \ {}" shows "\x \ S. \y \ T. dist x y = setdist S T" using setdist_compact_closed [OF T S \T \ {}\ \S \ {}\] by (metis dist_commute setdist_sym) lemma setdist_eq_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False then show ?thesis by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) qed corollary setdist_gt_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T > 0 \ (S \ {} \ T \ {} \ S \ T = {})" using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] by linarith lemma setdist_eq_0_closed_compact: assumes S: "closed S" and T: "compact T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" using setdist_eq_0_compact_closed [OF T S] by (metis Int_commute setdist_sym) lemma setdist_eq_0_bounded: fixes S :: "'a::heine_borel set" assumes "bounded S \ bounded T" shows "setdist S T = 0 \ S = {} \ T = {} \ closure S \ closure T \ {}" proof (cases "S = {} \ T = {}") case False then show ?thesis using setdist_eq_0_compact_closed [of "closure S" "closure T"] setdist_eq_0_closed_compact [of "closure S" "closure T"] assms by (force simp: bounded_closure compact_eq_bounded_closed) qed force lemma setdist_eq_0_sing_1: "setdist {x} S = 0 \ S = {} \ x \ closure S" by (metis in_closure_iff_infdist_zero infdist_def infdist_eq_setdist) lemma setdist_eq_0_sing_2: "setdist S {x} = 0 \ S = {} \ x \ closure S" by (metis setdist_eq_0_sing_1 setdist_sym) lemma setdist_neq_0_sing_1: "\setdist {x} S = a; a \ 0\ \ S \ {} \ x \ closure S" by (metis setdist_closure_2 setdist_empty2 setdist_eq_0I singletonI) lemma setdist_neq_0_sing_2: "\setdist S {x} = a; a \ 0\ \ S \ {} \ x \ closure S" by (simp add: setdist_neq_0_sing_1 setdist_sym) lemma setdist_sing_in_set: "x \ S \ setdist {x} S = 0" by (simp add: setdist_eq_0I) lemma setdist_eq_0_closed: "closed S \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (simp add: setdist_eq_0_sing_1) lemma setdist_eq_0_closedin: shows "\closedin (top_of_set U) S; x \ U\ \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) lemma setdist_gt_0_closedin: shows "\closedin (top_of_set U) S; x \ U; S \ {}; x \ S\ \ setdist {x} S > 0" using less_eq_real_def setdist_eq_0_closedin by fastforce no_notation eucl_less (infix "Differentiation\ theory Deriv imports Limits begin subsection \Frechet derivative\ definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" (infix "(has'_derivative)" 50) where "(f has_derivative f') F \ bounded_linear f' \ ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) \ 0) F" text \ Usually the filter \<^term>\F\ is \<^term>\at x within s\. \<^term>\(f has_derivative D) (at x within s)\ means: \<^term>\D\ is the derivative of function \<^term>\f\ at point \<^term>\x\ within the set \<^term>\s\. Where \<^term>\s\ is used to express left or right sided derivatives. In most cases \<^term>\s\ is either a variable or \<^term>\UNIV\. \ text \These are the only cases we'll care about, probably.\ lemma has_derivative_within: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x within s)" unfolding has_derivative_def tendsto_iff by (subst eventually_Lim_ident_at) (auto simp add: field_simps) lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" by simp definition has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" (infix "(has'_field'_derivative)" 50) where "(f has_field_derivative D) F \ (f has_derivative (*) D) F" lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" by simp definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" (infix "has'_vector'_derivative" 50) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" by simp named_theorems derivative_intros "structural introduction rules for derivatives" setup \ let val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs} fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\derivative_intros\ |> map_filter eq_rule) end \ text \ The following syntax is only used as a legacy syntax. \ abbreviation (input) FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> f' \ (f has_derivative f') (at x)" lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" by (simp add: has_derivative_def) lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" using bounded_linear.linear[OF has_derivative_bounded_linear] . lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" by (simp add: has_derivative_def) lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" by (metis eq_id_iff has_derivative_ident) lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" by (simp add: has_derivative_def) lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. lemma (in bounded_linear) has_derivative: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" unfolding has_derivative_def by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) lemmas has_derivative_scaleR_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_right] lemmas has_derivative_scaleR_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_left] lemmas has_derivative_mult_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_right] lemmas has_derivative_mult_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_left] lemmas has_derivative_of_real[derivative_intros, simp] = bounded_linear.has_derivative[OF bounded_linear_of_real] lemma has_derivative_add[simp, derivative_intros]: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" unfolding has_derivative_def proof safe let ?x = "Lim F (\x. x)" let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" have "((\x. ?D f f' x + ?D g g' x) \ (0 + 0)) F" using f g by (intro tendsto_add) (auto simp: has_derivative_def) then show "(?D (\x. f x + g x) (\x. f' x + g' x) \ 0) F" by (simp add: field_simps scaleR_add_right scaleR_diff_right) qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) lemma has_derivative_sum[simp, derivative_intros]: "(\i. i \ I \ (f i has_derivative f' i) F) \ ((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" by (induct I rule: infinite_finite_induct) simp_all lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" using has_derivative_scaleR_right[of f f' F "-1"] by simp lemma has_derivative_diff[simp, derivative_intros]: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) lemma has_derivative_at_within: "(f has_derivative f') (at x within s) \ (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s))" proof (cases "at x within s = bot") case True then show ?thesis by (metis (no_types, lifting) has_derivative_within tendsto_bot) next case False then show ?thesis by (simp add: Lim_ident_at has_derivative_def) qed lemma has_derivative_iff_norm: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] by (simp add: has_derivative_at_within divide_inverse ac_simps) lemma has_derivative_at: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) \0\ 0)" by (simp add: has_derivative_iff_norm LIM_offset_zero_iff) lemma field_has_derivative_at: fixes x :: "'a::real_normed_field" shows "(f has_derivative (*) D) (at x) \ (\h. (f (x + h) - f x) / h) \0\ D" (is "?lhs = ?rhs") proof - have "?lhs = (\h. norm (f (x + h) - f x - D * h) / norm h) \0 \ 0" by (simp add: bounded_linear_mult_right has_derivative_at) also have "... = (\y. norm ((f (x + y) - f x - D * y) / y)) \0\ 0" by (simp cong: LIM_cong flip: nonzero_norm_divide) also have "... = (\y. norm ((f (x + y) - f x) / y - D / y * y)) \0\ 0" by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) also have "... = ?rhs" by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) finally show ?thesis . qed lemma has_derivative_iff_Ex: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e. (\h. f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" unfolding has_derivative_at by force lemma has_derivative_at_within_iff_Ex: assumes "x \ S" "open S" shows "(f has_derivative f') (at x within S) \ bounded_linear f' \ (\e. (\h. x+h \ S \ f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" (is "?lhs = ?rhs") proof safe show "bounded_linear f'" if "(f has_derivative f') (at x within S)" using has_derivative_bounded_linear that by blast show "\e. (\h. x + h \ S \ f (x + h) = f x + f' h + e h) \ (\h. norm (e h) / norm h) \0\ 0" if "(f has_derivative f') (at x within S)" by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) show "(f has_derivative f') (at x within S)" if "bounded_linear f'" and eq [rule_format]: "\h. x + h \ S \ f (x + h) = f x + f' h + e h" and 0: "(\h. norm (e (h::'a)::'b) / norm h) \0\ 0" for e proof - have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \ S" for y using eq [of "y-x"] that by simp have 2: "((\y. norm (e (y-x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: "0" assms tendsto_offset_zero_iff) have "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: Lim_cong_within 1 2) then show ?thesis by (simp add: has_derivative_iff_norm \bounded_linear f'\) qed qed lemma has_derivativeI: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s) \ (f has_derivative f') (at x within s)" by (simp add: has_derivative_at_within) lemma has_derivativeI_sandwich: assumes e: "0 < e" and bounded: "bounded_linear f'" and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" and "(H \ 0) (at x within s)" shows "(f has_derivative f') (at x within s)" unfolding has_derivative_iff_norm proof safe show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" proof (rule tendsto_sandwich[where f="\x. 0"]) show "(H \ 0) (at x within s)" by fact show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" unfolding eventually_at using e sandwich by auto qed (auto simp: le_divide_eq) qed fact lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) lemma has_derivative_within_singleton_iff: "(f has_derivative g) (at x within {x}) \ bounded_linear g" by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) subsubsection \Limit transformation for derivatives\ lemma has_derivative_transform_within: assumes "(f has_derivative f') (at x within s)" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_open: assumes "(f has_derivative f') (at x within t)" and "open s" and "x \ s" and "\x. x\s \ f x = g x" shows "(g has_derivative f') (at x within t)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within_open) lemma has_derivative_transform: assumes "x \ s" "\x. x \ s \ g x = f x" assumes "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto lemma has_derivative_transform_eventually: assumes "(f has_derivative f') (at x within s)" "(\\<^sub>F x' in at x within s. f x' = g x')" assumes "f x = g x" "x \ s" shows "(g has_derivative f') (at x within s)" using assms proof - from assms(2,3) obtain d where "d > 0" "\x'. x' \ s \ dist x' x < d \ f x' = g x'" by (force simp: eventually_at) from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] show ?thesis . qed lemma has_field_derivative_transform_within: assumes "(f has_field_derivative f') (at a within S)" and "0 < d" and "a \ S" and "\x. \x \ S; dist x a < d\ \ f x = g x" shows "(g has_field_derivative f') (at a within S)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within) lemma has_field_derivative_transform_within_open: assumes "(f has_field_derivative f') (at a)" and "open S" "a \ S" and "\x. x \ S \ f x = g x" shows "(g has_field_derivative f') (at a)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within_open) subsection \Continuity\ lemma has_derivative_continuous: assumes f: "(f has_derivative f') (at x within s)" shows "continuous (at x within s) f" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) note F.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" using f unfolding has_derivative_iff_norm by blast then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" by (intro filterlim_cong) (simp_all add: eventually_at_filter) finally have "?L (\y. (f y - f x) - f' (y - x))" by (rule tendsto_norm_zero_cancel) then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) then have "?L (\y. f y - f x)" by simp from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis by (simp add: continuous_within) qed subsection \Composition\ lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f \ y) (at x within s) \ (f \ y) (inf (nhds x) (principal s))" unfolding tendsto_def eventually_inf_principal eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma has_derivative_in_compose: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at (f x) within (f`s))" shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast note G.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" define Nf where "Nf = ?N f f' x" define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\x. g' (f' x))" using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) next fix y :: 'a assume neq: "y \ x" have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" by (simp add: G.diff G.add field_simps) also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" proof (intro add_mono mult_left_mono) have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" by simp also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" by (rule norm_triangle_ineq) also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" using kF by (intro add_mono) simp finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" by (simp add: neq Nf_def field_simps) qed (use kG in \simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\) finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . next have [tendsto_intros]: "?L Nf" using f unfolding has_derivative_iff_norm Nf_def .. from f have "(f \ f x) (at x within s)" by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" unfolding filterlim_def by (simp add: eventually_filtermap eventually_at_filter le_principal) have "((?N g g' (f x)) \ 0) (at (f x) within f`s)" using g unfolding has_derivative_iff_norm .. then have g': "((?N g g' (f x)) \ 0) (inf (nhds (f x)) (principal (f`s)))" by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp have [tendsto_intros]: "?L Ng" unfolding Ng_def by (rule filterlim_compose[OF g' f']) show "((\y. Nf y * kG + Ng y * (Nf y + kF)) \ 0) (at x within s)" by (intro tendsto_eq_intros) auto qed simp qed lemma has_derivative_compose: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" by (blast intro: has_derivative_in_compose has_derivative_subset) lemma has_derivative_in_compose2: assumes "\x. x \ t \ (g has_derivative g' x) (at x within t)" assumes "f ` s \ t" "x \ s" assumes "(f has_derivative f') (at x within s)" shows "((\x. g (f x)) has_derivative (\y. g' (f x) (f' y))) (at x within s)" using assms by (auto intro: has_derivative_subset intro!: has_derivative_in_compose[of f f' x s g]) lemma (in bounded_bilinear) FDERIV: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" proof - from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" define Ng where "Ng = ?N g g'" define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" let ?F = "at x within s" show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\h. f x ** g' h + f' h ** g x)" by (intro bounded_linear_add bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) next from g have "(g \ g x) ?F" by (intro continuous_within[THEN iffD1] has_derivative_continuous) moreover from f g have "(Nf \ 0) ?F" "(Ng \ 0) ?F" by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) ultimately have "(?fun2 \ norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" by (intro tendsto_intros) (simp_all add: LIM_zero_iff) then show "(?fun2 \ 0) ?F" by simp next fix y :: 'd assume "y \ x" have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" by (simp add: diff_left diff_right add_left add_right field_simps) also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le]) also have "\ = ?fun2 y" by (simp add: add_divide_distrib Ng_def Nf_def) finally show "?fun1 y \ ?fun2 y" . qed simp qed lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] lemma has_derivative_prod[simp, derivative_intros]: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_field" shows "(\i. i \ I \ (f i has_derivative f' i) (at x within S)) \ ((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within S)" proof (induct I rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert i I) let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within S)" using insert by (intro has_derivative_mult) auto also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" using insert(1,2) by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) finally show ?case using insert by simp qed lemma has_derivative_power[simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" assumes f: "(f has_derivative f') (at x within S)" shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within S)" using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps) lemma has_derivative_inverse': fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within S)" (is "(_ has_derivative ?f) _") proof (rule has_derivativeI_sandwich) show "bounded_linear (\h. - (inverse x * h * inverse x))" by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) show "0 < norm x" using x by simp have "(inverse \ inverse x) (at x within S)" using tendsto_inverse tendsto_ident_at x by auto then show "((\y. norm (inverse y - inverse x) * norm (inverse x)) \ 0) (at x within S)" by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) next fix y :: 'a assume h: "y \ x" "dist y x < norm x" then have "y \ 0" by auto have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / norm (y - x)" by (simp add: \y \ 0\ inverse_diff_inverse x) also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" by (simp add: left_diff_distrib norm_minus_commute) also have "\ \ norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" by (simp add: norm_mult) also have "\ = norm (inverse y - inverse x) * norm (inverse x)" by simp finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \ norm (inverse y - inverse x) * norm (inverse x)" . qed lemma has_derivative_inverse[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within S)" using has_derivative_compose[OF f has_derivative_inverse', OF x] . lemma has_derivative_divide[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" assumes x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" using has_derivative_mult[OF f has_derivative_inverse[OF x g]] by (simp add: field_simps) lemma has_derivative_power_int': fixes x :: "'a::real_normed_field" assumes x: "x \ 0" shows "((\x. power_int x n) has_derivative (\y. y * (of_int n * power_int x (n - 1)))) (at x within S)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using x by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff fun_eq_iff simp flip: power_Suc) next case (neg n) thus ?thesis using x by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma has_derivative_power_int[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. power_int (f x) n) has_derivative (\h. f' h * (of_int n * power_int (f x) (n - 1)))) (at x within S)" using has_derivative_compose[OF f has_derivative_power_int', OF x] . text \Conventional form requires mult-AC laws. Types real and complex only.\ lemma has_derivative_divide'[derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" and x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" proof - have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = (f' h * g x - f x * g' h) / (g x * g x)" for h by (simp add: field_simps x) then show ?thesis using has_derivative_divide [OF f g] x by simp qed subsection \Uniqueness\ text \ This can not generally shown for \<^const>\has_derivative\, as we need to approach the point from all directions. There is a proof in \Analysis\ for \euclidean_space\. \ lemma has_derivative_at2: "(f has_derivative f') (at x) \ bounded_linear f' \ ((\y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x)" using has_derivative_within [of f f' x UNIV] by simp lemma has_derivative_zero_unique: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" proof - interpret F: bounded_linear F using assms by (rule has_derivative_bounded_linear) let ?r = "\h. norm (F h) / norm h" have *: "?r \0\ 0" using assms unfolding has_derivative_at by simp show "F = (\h. 0)" proof show "F h = 0" for h proof (rule ccontr) assume **: "\ ?thesis" then have h: "h \ 0" by (auto simp add: F.zero) with ** have "0 < ?r h" by simp from LIM_D [OF * this] obtain S where S: "0 < S" and r: "\x. x \ 0 \ norm x < S \ ?r x < ?r h" by auto from dense [OF S] obtain t where t: "0 < t \ t < S" .. let ?x = "scaleR (t / norm h) h" have "?x \ 0" and "norm ?x < S" using t h by simp_all then have "?r ?x < ?r h" by (rule r) then show False using t h by (simp add: F.scaleR) qed qed qed lemma has_derivative_unique: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" proof - have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" using has_derivative_diff [OF assms] by simp then have "(\h. F h - F' h) = (\h. 0)" by (rule has_derivative_zero_unique) then show "F = F'" unfolding fun_eq_iff right_minus_eq . qed lemma has_derivative_Uniq: "\\<^sub>\\<^sub>1F. (f has_derivative F) (at x)" by (simp add: Uniq_def has_derivative_unique) subsection \Differentiability predicate\ definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infix "differentiable" 50) where "f differentiable F \ (\D. (f has_derivative D) F)" lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" unfolding differentiable_def by (blast intro: has_derivative_subset) lemmas differentiable_within_subset = differentiable_subset lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_ident) lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_const) lemma differentiable_in_compose: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_in_compose) lemma differentiable_compose: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" by (blast intro: differentiable_in_compose differentiable_subset) lemma differentiable_add [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_add) lemma differentiable_sum[simp, derivative_intros]: assumes "finite s" "\a\s. (f a) differentiable net" shows "(\x. sum (\a. f a x) s) differentiable net" proof - from bchoice[OF assms(2)[unfolded differentiable_def]] show ?thesis by (auto intro!: has_derivative_sum simp: differentiable_def) qed lemma differentiable_minus [simp, derivative_intros]: "f differentiable F \ (\x. - f x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_minus) lemma differentiable_diff [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_diff) lemma differentiable_mult [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_mult) lemma differentiable_cmult_left_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. c * q t) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") proof assume L: ?lhs {assume "c \ 0" then have "q differentiable at t" using differentiable_mult [OF differentiable_const L, of concl: "1/c"] by auto } then show ?rhs by auto qed auto lemma differentiable_cmult_right_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. q t * c) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") by (simp add: mult.commute flip: differentiable_cmult_left_iff) lemma differentiable_inverse [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_inverse) lemma differentiable_divide [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" unfolding divide_inverse by simp lemma differentiable_power [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power) lemma differentiable_power_int [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. power_int (f x) n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power_int) lemma differentiable_scaleR [simp, derivative_intros]: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_scaleR) lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative (*) D) F" by (simp add: has_field_derivative_def) lemma DERIV_subset: "(f has_field_derivative f') (at x within s) \ t \ s \ (f has_field_derivative f') (at x within t)" by (simp add: has_field_derivative_def has_derivative_subset) lemma has_field_derivative_at_within: "(f has_field_derivative f') (at x) \ (f has_field_derivative f') (at x within s)" using DERIV_subset by blast abbreviation (input) DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D \ (f has_field_derivative D) (at x)" abbreviation has_real_derivative :: "(real \ real) \ real \ real filter \ bool" (infix "(has'_real'_derivative)" 50) where "(f has_real_derivative D) F \ (f has_field_derivative D) F" lemma real_differentiable_def: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" proof safe assume "f differentiable at x within s" then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = ((*) c)" by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) lemma real_differentiableE [elim?]: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" using assms by (auto simp: real_differentiable_def) lemma has_field_derivative_iff: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" proof - have "((\y. norm (f y - f x - D * (y - x)) / norm (y - x)) \ 0) (at x within S) = ((\y. (f y - f x) / (y - x) - D) \ 0) (at x within S)" by (smt (verit, best) Lim_cong_within divide_diff_eq_iff norm_divide right_minus_eq tendsto_norm_zero_iff) then show ?thesis by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) qed lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. text \due to Christian Pardillo Laursen, replacing a proper epsilon-delta horror\ lemma field_derivative_lim_unique: assumes f: "(f has_field_derivative df) (at z)" and s: "s \ 0" "\n. s n \ 0" and a: "(\n. (f (z + s n) - f z) / s n) \ a" shows "df = a" proof - have "((\k. (f (z + k) - f z) / k) \ df) (at 0)" using f by (simp add: DERIV_def) with s have "((\n. (f (z + s n) - f z) / s n) \ df)" by (simp flip: LIMSEQ_SEQ_conv) then show ?thesis using a by (rule LIMSEQ_unique) qed lemma mult_commute_abs: "(\x. x * c) = (*) c" for c :: "'a::ab_semigroup_mult" by (simp add: fun_eq_iff mult.commute) lemma DERIV_compose_FDERIV: fixes f::"real\real" assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) subsection \Vector derivative\ -lemma has_field_derivative_iff_has_vector_derivative: - "(f has_field_derivative y) F \ (f has_vector_derivative y) F" +text \It's for real derivatives only, and not obviously generalisable to field derivatives\ +lemma has_real_derivative_iff_has_vector_derivative: + "(f has_real_derivative y) F \ (f has_vector_derivative y) F" unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" - unfolding has_field_derivative_def by (rule has_derivative_subset) + by (fact DERIV_subset) lemma has_vector_derivative_const[simp, derivative_intros]: "((\x. c) has_vector_derivative 0) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_id[simp, derivative_intros]: "((\x. x) has_vector_derivative 1) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_minus[derivative_intros]: "(f has_vector_derivative f') net \ ((\x. - f x) has_vector_derivative (- f')) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_add[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x + g x) has_vector_derivative (f' + g')) net" by (auto simp: has_vector_derivative_def scaleR_right_distrib) lemma has_vector_derivative_sum[derivative_intros]: "(\i. i \ I \ (f i has_vector_derivative f' i) net) \ ((\x. \i\I. f i x) has_vector_derivative (\i\I. f' i)) net" by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) lemma has_vector_derivative_diff[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) lemma has_vector_derivative_add_const: "((\t. g t + z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" apply (intro iffI) apply (force dest: has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const]) apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) done lemma has_vector_derivative_diff_const: "((\t. g t - z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" using has_vector_derivative_add_const [where z = "-z"] by simp lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" using has_derivative[OF assms[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR) lemma (in bounded_bilinear) has_vector_derivative: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at x within s)" shows "((\x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) lemma has_vector_derivative_scaleR[derivative_intros]: "(f has_field_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" - unfolding has_field_derivative_iff_has_vector_derivative + unfolding has_real_derivative_iff_has_vector_derivative by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) lemma has_vector_derivative_mult[derivative_intros]: "(f has_vector_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" for f g :: "real \ 'a::real_normed_algebra" by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) lemma has_vector_derivative_of_real[derivative_intros]: "(f has_field_derivative D) F \ ((\x. of_real (f x)) has_vector_derivative (of_real D)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) - (simp add: has_field_derivative_iff_has_vector_derivative) + (simp add: has_real_derivative_iff_has_vector_derivative) lemma has_vector_derivative_real_field: "(f has_field_derivative f') (at (of_real a)) \ ((\x. f (of_real x)) has_vector_derivative f') (at a within s)" using has_derivative_compose[of of_real of_real a _ f "(*) f'"] by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) lemma has_vector_derivative_continuous: "(f has_vector_derivative D) (at x within s) \ continuous (at x within s) f" by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) lemma continuous_on_vector_derivative: "(\x. x \ S \ (f has_vector_derivative f' x) (at x within S)) \ continuous_on S f" by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) lemma has_vector_derivative_mult_right[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. a * f x) has_vector_derivative (a * x)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) lemma has_vector_derivative_mult_left[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" using has_vector_derivative_mult_left [of f x F "inverse a"] by (simp add: field_class.field_divide_inverse) subsection \Derivatives\ lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" by (simp add: DERIV_def) lemma has_field_derivativeD: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" by (simp add: has_field_derivative_iff) lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto lemma field_differentiable_add[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z + g z) has_field_derivative f' + g') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_add: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x + g x) has_field_derivative D + E) (at x within s)" by (rule field_differentiable_add) lemma field_differentiable_minus[derivative_intros]: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" by (rule field_differentiable_minus) lemma field_differentiable_diff[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) corollary DERIV_diff: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x - g x) has_field_derivative D - E) (at x within s)" by (rule field_differentiable_diff) lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" by (rule DERIV_continuous) lemma DERIV_atLeastAtMost_imp_continuous_on: assumes "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y" shows "continuous_on {a..b} f" by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma DERIV_continuous_on: "(\x. x \ s \ (f has_field_derivative (D x)) (at x within s)) \ continuous_on s f" unfolding continuous_on_eq_continuous_within by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) lemma DERIV_mult': "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_mult[derivative_intros]: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) text \Derivative of linear multiplication\ lemma DERIV_cmult: "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" by (drule DERIV_mult' [OF DERIV_const]) simp lemma DERIV_cmult_right: "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" using DERIV_cmult by (auto simp add: ac_simps) lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp lemma DERIV_cdivide: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" using DERIV_cmult_right[of f D x s "1 / c"] by simp lemma DERIV_unique: "DERIV f x :> D \ DERIV f x :> E \ D = E" unfolding DERIV_def by (rule LIM_unique) lemma DERIV_Uniq: "\\<^sub>\\<^sub>1D. DERIV f x :> D" by (simp add: DERIV_unique Uniq_def) lemma DERIV_sum[derivative_intros]: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ ((\x. sum (f x) S) has_field_derivative sum (f' x) S) F" by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_inverse'[derivative_intros]: assumes "(f has_field_derivative D) (at x within s)" and "f x \ 0" shows "((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" proof - have "(f has_derivative (\x. x * D)) = (f has_derivative (*) D)" by (rule arg_cong [of "\x. x * D"]) (simp add: fun_eq_iff) with assms have "(f has_derivative (\x. x * D)) (at x within s)" by (auto dest!: has_field_derivative_imp_has_derivative) then show ?thesis using \f x \ 0\ by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) qed text \Power of \-1\\ lemma DERIV_inverse: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" by (drule DERIV_inverse' [OF DERIV_ident]) simp text \Derivative of inverse\ lemma DERIV_inverse_fun: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) text \Derivative of quotient\ lemma DERIV_divide[derivative_intros]: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ g x \ 0 \ ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) lemma DERIV_quotient: "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_power[derivative_intros]: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" using DERIV_power [OF DERIV_ident] by simp lemma DERIV_power_int [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative d) (at x within s)" and [simp]: "f x \ 0" shows "((\x. power_int (f x) n) has_field_derivative (of_int n * power_int (f x) (n - 1) * d)) (at x within s)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff simp flip: power_Suc power_Suc2 power_add) next case (neg n) thus ?thesis by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" using has_derivative_compose[of f "(*) D" x s g "(*) E"] by (simp only: has_field_derivative_def mult_commute_abs ac_simps) corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" by (rule DERIV_chain') text \Standard version\ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lemma DERIV_chain_s: assumes "(\x. x \ s \ DERIV g x :> g'(x))" and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" and "DERIV f x :> f'" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) text \Alternative definition for differentiability\ lemma DERIV_LIM_iff: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows "((\h. (f (a + h) - f a) / h) \0\ D) = ((\x. (f x - f a) / (x - a)) \a\ D)" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. (f (a + (x + - a)) - f a) / (x + - a)) \0 - - a\ D" by (rule LIM_offset) then show ?rhs by simp next assume ?rhs then have "(\x. (f (x+a) - f a) / ((x+a) - a)) \a-a\ D" by (rule LIM_offset) then show ?lhs by (simp add: add.commute) qed lemma has_field_derivative_cong_ev: assumes "x = y" and *: "eventually (\x. x \ S \ f x = g x) (nhds x)" and "u = v" "S = t" "x \ S" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" unfolding has_field_derivative_iff proof (rule filterlim_cong) from assms have "f y = g y" by (auto simp: eventually_nhds) with * show "\\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" unfolding eventually_at_filter by eventually_elim (auto simp: assms \f y = g y\) qed (simp_all add: assms) lemma has_field_derivative_cong_eventually: assumes "eventually (\x. f x = g x) (at x within S)" "f x = g x" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" unfolding has_field_derivative_iff proof (rule tendsto_cong) show "\\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" using assms by (auto elim: eventually_mono) qed lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" by (rule has_field_derivative_cong_ev) simp_all lemma DERIV_mirror: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x)) x :> - y)" for f :: "real \ real" and x y :: real by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right tendsto_minus_cancel_left field_simps conj_commute) lemma DERIV_shift: "(f has_field_derivative y) (at (x + z)) = ((\x. f (x + z)) has_field_derivative y) (at x)" by (simp add: DERIV_def field_simps) lemma DERIV_at_within_shift_lemma: assumes "(f has_field_derivative y) (at (z+x) within (+) z ` S)" shows "(f \ (+)z has_field_derivative y) (at x within S)" proof - have "((+)z has_field_derivative 1) (at x within S)" by (rule derivative_eq_intros | simp)+ with assms DERIV_image_chain show ?thesis by (metis mult.right_neutral) qed lemma DERIV_at_within_shift: "(f has_field_derivative y) (at (z+x) within (+) z ` S) \ ((\x. f (z+x)) has_field_derivative y) (at x within S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using DERIV_at_within_shift_lemma unfolding o_def by blast next have [simp]: "(\x. x - z) ` (+) z ` S = S" by force assume R: ?rhs have "(f \ (+) z \ (+) (- z) has_field_derivative y) (at (z + x) within (+) z ` S)" by (rule DERIV_at_within_shift_lemma) (use R in \simp add: o_def\) then show ?lhs by (simp add: o_def) qed lemma floor_has_real_derivative: fixes f :: "real \ 'a::{floor_ceiling,order_topology}" assumes "isCont f x" and "f x \ \" shows "((\x. floor (f x)) has_real_derivative 0) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "((\_. floor (f x)) has_real_derivative 0) (at x)" by simp have "\\<^sub>F y in at x. \f y\ = \f x\" by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) then show "\\<^sub>F y in nhds x. real_of_int \f y\ = real_of_int \f x\" unfolding eventually_at_filter by eventually_elim auto qed lemmas has_derivative_floor[derivative_intros] = floor_has_real_derivative[THEN DERIV_compose_FDERIV] lemma continuous_floor: fixes x::real shows "x \ \ \ continuous (at x) (real_of_int \ floor)" using floor_has_real_derivative [where f=id] by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) lemma continuous_frac: fixes x::real assumes "x \ \" shows "continuous (at x) frac" proof - have "isCont (\x. real_of_int \x\) x" using continuous_floor [OF assms] by (simp add: o_def) then have *: "continuous (at x) (\x. x - real_of_int \x\)" by (intro continuous_intros) moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" by (simp add: frac_def) ultimately show ?thesis by (simp add: LIM_imp_LIM frac_def isCont_def) qed text \Caratheodory formulation of derivative at a point\ lemma CARAT_DERIV: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show "\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l" proof (intro exI conjI) let ?g = "(\z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z - x)" by simp show "isCont ?g x" using \?lhs\ by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume ?rhs then show ?lhs by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) qed subsection \Local extrema\ text \If \<^term>\0 < f' x\ then \<^term>\x\ is Locally Strictly Increasing At The Right.\ lemma has_real_derivative_pos_inc_right: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and l: "0 < l" shows "\d > 0. \h > 0. x + h \ S \ h < d \ f x < f (x + h)" using assms proof - from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x + h \ S" with all [of "x + h"] show "f x < f (x+h)" proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) assume "\ (f (x + h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x + h) - f x) / h" by arith then show "f x < f (x + h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_pos_inc_right: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "0 < l" shows "\d > 0. \h > 0. h < d \ f x < f (x + h)" using has_real_derivative_pos_inc_right[OF assms] by auto lemma has_real_derivative_neg_dec_left: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and "l < 0" shows "\d > 0. \h > 0. x - h \ S \ h < d \ f x < f (x - h)" proof - from \l < 0\ have l: "- l > 0" by simp from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < - l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x - h \ S" with all [of "x - h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith then show "f x < f (x - h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_neg_dec_left: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "l < 0" shows "\d > 0. \h > 0. h < d \ f x < f (x - h)" using has_real_derivative_neg_dec_left[OF assms] by auto lemma has_real_derivative_pos_inc_left: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ 0 < l \ \d>0. \h>0. x - h \ S \ h < d \ f (x - h) < f x" by (rule has_real_derivative_neg_dec_left [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_pos_inc_left: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d \ f (x - h) < f x" using has_real_derivative_pos_inc_left by blast lemma has_real_derivative_neg_dec_right: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ l < 0 \ \d > 0. \h > 0. x + h \ S \ h < d \ f x > f (x + h)" by (rule has_real_derivative_pos_inc_right [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_neg_dec_right: fixes f :: "real \ real" shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d \ f x > f (x + h)" using has_real_derivative_neg_dec_right by blast lemma DERIV_local_max: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and d: "0 < d" and le: "\y. \x - y\ < d \ f y \ f x" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal then show ?thesis . next case less from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x - h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x - e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x + e"]] show ?thesis by (auto simp add: abs_if) qed text \Similar theorem for a local minimum\ lemma DERIV_local_min: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x \ f y \ l = 0" by (drule DERIV_minus [THEN DERIV_local_max]) auto text\In particular, if a function is locally flat\ lemma DERIV_local_const: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x = f y \ l = 0" by (auto dest!: DERIV_local_max) subsection \Rolle's Theorem\ text \Lemma about introducing open ball in open interval\ lemma lemma_interval_lt: fixes a b x :: real assumes "a < x" "x < b" shows "\d. 0 < d \ (\y. \x - y\ < d \ a < y \ y < b)" using linorder_linear [of "x - a" "b - x"] proof assume "x - a \ b - x" with assms show ?thesis by (rule_tac x = "x - a" in exI) auto next assume "b - x \ x - a" with assms show ?thesis by (rule_tac x = "b - x" in exI) auto qed lemma lemma_interval: "a < x \ x < b \ \d. 0 < d \ (\y. \x - y\ < d \ a \ y \ y \ b)" for a b x :: real by (force dest: lemma_interval_lt) text \Rolle's Theorem. If \<^term>\f\ is defined and continuous on the closed interval \[a,b]\ and differentiable on the open interval \(a,b)\, and \<^term>\f a = f b\, then there exists \x0 \ (a,b)\ such that \<^term>\f' x0 = 0\\ theorem Rolle_deriv: fixes f :: "real \ real" assumes "a < b" and fab: "f a = f b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\z. a < z \ z < b \ f' z = (\v. 0)" proof - have le: "a \ b" using \a < b\ by simp have "(a + b) / 2 \ {a..b}" using assms(1) by auto then have *: "{a..b} \ {}" by auto obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" and "a \ x" "x \ b" using continuous_attains_sup[OF compact_Icc * contf] by (meson atLeastAtMost_iff) obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" and "a \ x'" "x' \ b" using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) consider "a < x" "x < b" | "x = a \ x = b" using \a \ x\ \x \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its maximum within the interval\ then obtain l where der: "DERIV f x :> l" using derf differentiable_def real_differentiable_def by blast obtain d where d: "0 < d" and bound: "\y. \x - y\ < d \ a \ y \ y \ b" using lemma_interval [OF 1] by blast then have bound': "\y. \x - y\ < d \ f y \ f x" using x_max by blast \ \the derivative at a local maximum is zero\ have "l = 0" by (rule DERIV_local_max [OF der d bound']) with 1 der derf [of x] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 then have fx: "f b = f x" by (auto simp add: fab) consider "a < x'" "x' < b" | "x' = a \ x' = b" using \a \ x'\ \x' \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its minimum within the interval\ then obtain l where der: "DERIV f x' :> l" using derf differentiable_def real_differentiable_def by blast from lemma_interval [OF 1] obtain d where d: "0y. \x'-y\ < d \ a \ y \ y \ b" by blast then have bound': "\y. \x' - y\ < d \ f x' \ f y" using x'_min by blast have "l = 0" by (rule DERIV_local_min [OF der d bound']) \ \the derivative at a local minimum is zero\ then show ?thesis using 1 der derf [of x'] by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 \ \\<^term>\f\ is constant throughout the interval\ then have fx': "f b = f x'" by (auto simp: fab) from dense [OF \a < b\] obtain r where r: "a < r" "r < b" by blast obtain d where d: "0 < d" and bound: "\y. \r - y\ < d \ a \ y \ y \ b" using lemma_interval [OF r] by blast have eq_fb: "f z = f b" if "a \ z" and "z \ b" for z proof (rule order_antisym) show "f z \ f b" by (simp add: fx x_max that) show "f b \ f z" by (simp add: fx' x'_min that) qed have bound': "\y. \r - y\ < d \ f r = f y" proof (intro strip) fix y :: real assume lt: "\r - y\ < d" then have "f y = f b" by (simp add: eq_fb bound) then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) qed obtain l where der: "DERIV f r :> l" using derf differentiable_def r(1) r(2) real_differentiable_def by blast have "l = 0" by (rule DERIV_local_const [OF der d bound']) \ \the derivative of a constant function is zero\ with r der derf [of r] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) qed qed qed corollary Rolle: fixes a b :: real assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f" and dif [rule_format]: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\z. a < z \ z < b \ DERIV f z :> 0" proof - obtain f' where f': "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then have "\z. a < z \ z < b \ f' z = (\v. 0)" by (metis Rolle_deriv [OF ab]) then show ?thesis using f' has_derivative_imp_has_field_derivative by fastforce qed subsection \Mean Value Theorem\ theorem mvt: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" obtains \ where "a < \" "\ < b" "f b - f a = (f' \) (b - a)" proof - have "\\. a < \ \ \ < b \ (\y. f' \ y - (f b - f a) / (b - a) * y) = (\v. 0)" proof (intro Rolle_deriv[OF \a < b\]) fix x assume x: "a < x" "x < b" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\y. f' x y - (f b - f a) / (b - a) * y)) (at x)" by (intro derivative_intros derf[OF x]) qed (use assms in \auto intro!: continuous_intros simp: field_simps\) then show ?thesis by (smt (verit, ccfv_SIG) pos_le_divide_eq pos_less_divide_eq that) qed theorem MVT: fixes a b :: real assumes lt: "a < b" and contf: "continuous_on {a..b} f" and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - obtain f' :: "real \ real \ real" where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis by (simp add: ac_simps) (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: assumes "a < b" and der: "\x. \a \ x; x \ b\ \ DERIV f x :> f' x" shows "\z::real. a < z \ z < b \ (f b - f a = (b - a) * f' z)" proof - have "\l z. a < z \ z < b \ (f has_real_derivative l) (at z) \ f b - f a = (b - a) * l" proof (rule MVT [OF \a < b\]) show "continuous_on {a..b} f" by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) show "\x. \a < x; x < b\ \ f differentiable (at x)" using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) qed with assms show ?thesis by (blast dest: DERIV_unique order_less_imp_le) qed lemma pos_deriv_imp_strict_mono: assumes "\x. (f has_real_derivative f' x) (at x)" assumes "\x. f' x > 0" shows "strict_mono f" proof (rule strict_monoI) fix x y :: real assume xy: "x < y" from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" by (intro MVT2) (auto dest: connectedD_interval) then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast note \f y - f x = (y - x) * f' z\ also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto finally show "f x < f y" by simp qed proposition deriv_nonneg_imp_mono: assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes ab: "a \ b" shows "g a \ g b" proof (cases "a < b") assume "a < b" from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp with MVT2[OF \a < b\] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp qed (insert ab, simp) subsubsection \A function is constant if its derivative is 0 over an interval.\ lemma DERIV_isconst_end: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and 0: "\x. \a < x; x < b\ \ DERIV f x :> 0" shows "f b = f a" using MVT [OF \a < b\] "0" DERIV_unique contf real_differentiable_def by (fastforce simp: algebra_simps) lemma DERIV_isconst2: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ DERIV f x :> 0" and "a \ x" "x \ b" shows "f x = f a" proof (cases "a < x") case True have *: "continuous_on {a..x} f" using \x \ b\ contf continuous_on_subset by fastforce show ?thesis by (rule DERIV_isconst_end [OF True *]) (use \x \ b\ derf in auto) qed (use \a \ x\ in auto) lemma DERIV_isconst3: fixes a b x y :: real assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" and derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" shows "f x = f y" proof (cases "x = y") case False let ?a = "min x y" let ?b = "max x y" have *: "DERIV f z :> 0" if "?a \ z" "z \ ?b" for z proof - have "a < z" and "z < b" using that \x \ {a <..< b}\ and \y \ {a <..< b}\ by auto then have "z \ {a<.. 0" by (rule derivable) qed have isCont: "continuous_on {?a..?b} f" by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) have DERIV: "\z. \?a < z; z < ?b\ \ DERIV f z :> 0" using * by auto have "?a < ?b" using \x \ y\ by auto from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] show ?thesis by auto qed auto lemma DERIV_isconst_all: fixes f :: "real \ real" shows "\x. DERIV f x :> 0 \ f x = f y" apply (rule linorder_cases [of x y]) apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ done lemma DERIV_const_ratio_const: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "f b - f a = (b - a) * k" proof (cases a b rule: linorder_cases) case less show ?thesis using MVT [OF less] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) next case greater have "f a - f b = (a - b) * k" using MVT [OF greater] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) then show ?thesis by (simp add: algebra_simps) qed auto lemma DERIV_const_ratio_const2: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "(f b - f a) / (b - a) = k" using DERIV_const_ratio_const [OF assms] \a \ b\ by auto lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" for a b :: real by simp lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" for a b :: real by simp text \Gallileo's "trick": average velocity = av. of end velocities.\ lemma DERIV_const_average: fixes v :: "real \ real" and a b :: real assumes neq: "a \ b" and der: "\x. DERIV v x :> k" shows "v ((a + b) / 2) = (v a + v b) / 2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by (force simp add: add.commute) qed subsubsection\A function with positive derivative is increasing\ text \A simple proof using the MVT, by Jeremy Avigad. And variants.\ lemma DERIV_pos_imp_increasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y > 0)" and con: "continuous_on {a..b} f" shows "f a < f b" proof (rule ccontr) assume f: "\ ?thesis" have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" by (rule MVT) (use assms real_differentiable_def in \force+\) then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto with assms f have "\ l > 0" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) with assms z show False by (metis DERIV_unique) qed lemma DERIV_pos_imp_increasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y > 0" shows "f a < f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \a < b\] der) lemma DERIV_nonneg_imp_nondecreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof (rule ccontr, cases "a = b") assume "\ ?thesis" and "a = b" then show False by auto next assume *: "\ ?thesis" assume "a \ b" with \a \ b\ have "a < b" by linarith moreover have "continuous_on {a..b} f" by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) ultimately have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" using assms MVT [OF \a < b\, of f] real_differentiable_def less_eq_real_def by blast then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" by auto with * have "a < b" "f b < f a" by auto with ** have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) with assms lz show False by (metis DERIV_unique order_less_imp_le) qed lemma DERIV_neg_imp_decreasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y < 0" and con: "continuous_on {a..b} f" shows "f a > f b" proof - have "(\x. -f x) a < (\x. -f x) b" proof (rule DERIV_pos_imp_increasing_open [of a b]) show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 < y" using assms by simp (metis field_differentiable_minus neg_0_less_iff_less) show "continuous_on {a..b} (\x. - f x)" using con continuous_on_minus by blast qed (use assms in auto) then show ?thesis by simp qed lemma DERIV_neg_imp_decreasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y < 0" shows "f a > f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \a < b\] der) lemma DERIV_nonpos_imp_nonincreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof - have "(\x. -f x) a \ (\x. -f x) b" using DERIV_nonneg_imp_nondecreasing [of a b "\x. -f x"] assms DERIV_minus by fastforce then show ?thesis by simp qed lemma DERIV_pos_imp_increasing_at_bot: fixes f :: "real \ real" assumes "\x. x \ b \ (\y. DERIV f x :> y \ y > 0)" and lim: "(f \ flim) at_bot" shows "flim < f b" proof - have "\N. \n\N. f n \ f (b - 1)" by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) then have "flim \ f (b - 1)" by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) also have "\ < f b" by (force intro: DERIV_pos_imp_increasing [where f=f] assms) finally show ?thesis . qed lemma DERIV_neg_imp_decreasing_at_top: fixes f :: "real \ real" assumes der: "\x. x \ b \ \y. DERIV f x :> y \ y < 0" and lim: "(f \ flim) at_top" shows "flim < f b" apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) apply (metis filterlim_at_top_mirror lim) done text \Derivative of inverse function\ lemma DERIV_inverse_function: fixes f g :: "real \ real" assumes der: "DERIV f (g x) :> D" and neq: "D \ 0" and x: "a < x" "x < b" and inj: "\y. \a < y; y < b\ \ f (g y) = y" and cont: "isCont g x" shows "DERIV g x :> inverse D" unfolding has_field_derivative_iff proof (rule LIM_equal2) show "0 < min (x - a) (b - x)" using x by arith next fix y assume "norm (y - x) < min (x - a) (b - x)" then have "a < y" and "y < b" by (simp_all add: abs_less_iff) then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" by (simp add: inj) next have "(\z. (f z - f (g x)) / (z - g x)) \g x\ D" by (rule der [unfolded has_field_derivative_iff]) then have 1: "(\z. (f z - x) / (z - g x)) \g x\ D" using inj x by simp have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" proof (rule exI, safe) show "0 < min (x - a) (b - x)" using x by simp next fix y assume "norm (y - x) < min (x - a) (b - x)" then have y: "a < y" "y < b" by (simp_all add: abs_less_iff) assume "g y = g x" then have "f (g y) = f (g x)" by simp then have "y = x" using inj y x by simp also assume "y \ x" finally show False by simp qed have "(\y. (f (g y) - x) / (g y - g x)) \x\ D" using cont 1 2 by (rule isCont_LIM_compose2) then show "(\y. inverse ((f (g y) - x) / (g y - g x))) \x\ inverse D" using neq by (rule tendsto_inverse) qed subsection \Generalized Mean Value Theorem\ theorem GMVT: fixes a b :: real assumes alb: "a < b" and fc: "\x. a \ x \ x \ b \ isCont f x" and fd: "\x. a < x \ x < b \ f differentiable (at x)" and gc: "\x. a \ x \ x \ b \ isCont g x" and gd: "\x. a < x \ x < b \ g differentiable (at x)" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" proof - let ?h = "\x. (f b - f a) * g x - (g b - g a) * f x" have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" proof (rule MVT) from assms show "a < b" by simp show "continuous_on {a..b} ?h" by (simp add: continuous_at_imp_continuous_on fc gc) show "\x. \a < x; x < b\ \ ?h differentiable (at x)" using fd gd by simp qed then obtain l where l: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where c: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. from c have cint: "a < c \ c < b" by auto then obtain g'c where g'c: "DERIV g c :> g'c" using gd real_differentiable_def by blast from c have "a < c \ c < b" by auto then obtain f'c where f'c: "DERIV f c :> f'c" using fd real_differentiable_def by blast from c have "DERIV ?h c :> l" by auto moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" using g'c f'c by (auto intro!: derivative_eq_intros) ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" proof - from c have "?h b - ?h a = (b - a) * l" by auto also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp finally show ?thesis by simp qed moreover have "?h b - ?h a = 0" proof - have "?h b - ?h a = ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" by (simp add: algebra_simps) then show ?thesis by auto qed ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) with g'c f'c cint show ?thesis by auto qed lemma GMVT': fixes f g :: "real \ real" assumes "a < b" and isCont_f: "\z. a \ z \ z \ b \ isCont f z" and isCont_g: "\z. a \ z \ z \ b \ isCont g z" and DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" and DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" proof - have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" using assms by (intro GMVT) (force simp: real_differentiable_def)+ then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" using DERIV_f DERIV_g by (force dest: DERIV_unique) then show ?thesis by auto qed subsection \L'Hopitals rule\ lemma isCont_If_ge: fixes a :: "'a :: linorder_topology" assumes "continuous (at_left a) g" and f: "(f \ g a) (at_right a)" shows "isCont (\x. if x \ a then g x else f x) a" (is "isCont ?gf a") proof - have g: "(g \ g a) (at_left a)" using assms continuous_within by blast show ?thesis unfolding isCont_def continuous_within proof (intro filterlim_split_at; simp) show "(?gf \ g a) (at_left a)" by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) show "(?gf \ g a) (at_right a)" by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) qed qed lemma lhopital_right_0: fixes f0 g0 :: "real \ real" assumes f_0: "(f0 \ 0) (at_right 0)" and g_0: "(g0 \ 0) (at_right 0)" and ev: "eventually (\x. g0 x \ 0) (at_right 0)" "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" and lim: "filterlim (\ x. (f' x / g' x)) F (at_right 0)" shows "filterlim (\ x. f0 x / g0 x) F (at_right 0)" proof - define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" using ev by eventually_elim auto then obtain a where [arith]: "0 < a" and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" unfolding eventually_at by (auto simp: dist_real_def) have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" using g0_neq_0 by (simp add: g_def) have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have "isCont f 0" unfolding f_def by (intro isCont_If_ge f_0 continuous_const) have "isCont g 0" unfolding g_def by (intro isCont_If_ge g_0 continuous_const) have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" proof (rule bchoice, rule ballI) fix x assume "x \ {0 <..< a}" then have x[arith]: "0 < x" "x < a" by auto with g'_neq_0 g_neq_0 \g 0 = 0\ have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" by auto have "\x. 0 \ x \ x < a \ isCont f x" using \isCont f 0\ f by (auto intro: DERIV_isCont simp: le_less) moreover have "\x. 0 \ x \ x < a \ isCont g x" using \isCont g 0\ g by (auto intro: DERIV_isCont simp: le_less) ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" using f g \x < a\ by (intro GMVT') auto then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" by blast moreover from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" by (simp add: field_simps) ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" using \f 0 = 0\ \g 0 = 0\ by (auto intro!: exI[of _ c]) qed then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) moreover from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" by eventually_elim auto then have "((\x. norm (\ x)) \ 0) (at_right 0)" by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) auto then have "(\ \ 0) (at_right 0)" by (rule tendsto_norm_zero_cancel) with \ have "filterlim \ (at_right 0) (at_right 0)" by (auto elim!: eventually_mono simp: filterlim_at) from this lim have "filterlim (\t. f' (\ t) / g' (\ t)) F (at_right 0)" by (rule_tac filterlim_compose[of _ _ _ \]) ultimately have "filterlim (\t. f t / g t) F (at_right 0)" (is ?P) by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) (auto elim: eventually_mono) also have "?P \ ?thesis" by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) finally show ?thesis . qed lemma lhopital_right: "(f \ 0) (at_right x) \ (g \ 0) (at_right x) \ eventually (\x. g x \ 0) (at_right x) \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ filterlim (\ x. (f' x / g' x)) F (at_right x) \ filterlim (\ x. f x / g x) F (at_right x)" for x :: real unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0) lemma lhopital_left: "(f \ 0) (at_left x) \ (g \ 0) (at_left x) \ eventually (\x. g x \ 0) (at_left x) \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ filterlim (\ x. (f' x / g' x)) F (at_left x) \ filterlim (\ x. f x / g x) F (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital: "(f \ 0) (at x) \ (g \ 0) (at x) \ eventually (\x. g x \ 0) (at x) \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ filterlim (\ x. (f' x / g' x)) F (at x) \ filterlim (\ x. f x / g x) F (at x)" for x :: real unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) lemma lhopital_right_0_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_right 0. g x :> at_top" and ev: "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f x :> f' x) (at_right 0)" "eventually (\x. DERIV g x :> g' x) (at_right 0)" and lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f x / g x) \ x) (at_right 0)" unfolding tendsto_iff proof safe fix e :: real assume "0 < e" with lim[unfolded tendsto_iff, rule_format, of "e / 4"] have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] obtain a where [arith]: "0 < a" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" unfolding eventually_at_le by (auto simp: dist_real_def) from Df have "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) moreover have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) moreover have inv_g: "((\x. inverse (g x)) \ 0) (at_right 0)" using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] by (rule filterlim_compose) then have "((\x. norm (1 - g a * inverse (g x))) \ norm (1 - g a * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\x. norm (1 - g a / g x)) \ 1) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of 1] have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" by (auto elim!: eventually_mono simp: dist_real_def) moreover from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) \ norm ((f a - x * g a) * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\t. norm (f a - x * g a) / norm (g t)) \ 0) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of "e / 2"] \0 < e\ have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" by (auto simp: dist_real_def) ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" proof eventually_elim fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ then obtain y where [arith]: "t < y" "y < a" and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" by blast from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" using \g a < g t\ g'_neq_0[of y] by (auto simp add: field_simps) have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" by (simp add: field_simps) have "norm (f t / g t - x) \ norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" unfolding * by (rule norm_triangle_ineq) also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" by (simp add: abs_mult D_eq dist_real_def) also have "\ < (e / 4) * 2 + e / 2" using ineq Df[of y] \0 < e\ by (intro add_le_less_mono mult_mono) auto finally show "dist (f t / g t) x < e" by (simp add: dist_real_def) qed qed lemma lhopital_right_at_top: "LIM x at_right x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ ((\ x. (f' x / g' x)) \ y) (at_right x) \ ((\ x. f x / g x) \ y) (at_right x)" unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0_at_top) lemma lhopital_left_at_top: "LIM x at_left x. g x :> at_top \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ ((\ x. (f' x / g' x)) \ y) (at_left x) \ ((\ x. f x / g x) \ y) (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital_at_top: "LIM x at x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ ((\ x. (f' x / g' x)) \ y) (at x) \ ((\ x. f x / g x) \ y) (at x)" unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) lemma lhospital_at_top_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_top. g x :> at_top" and g': "eventually (\x. g' x \ 0) at_top" and Df: "eventually (\x. DERIV f x :> f' x) at_top" and Dg: "eventually (\x. DERIV g x :> g' x) at_top" and lim: "((\ x. (f' x / g' x)) \ x) at_top" shows "((\ x. f x / g x) \ x) at_top" unfolding filterlim_at_top_to_right proof (rule lhopital_right_0_at_top) let ?F = "\x. f (inverse x)" let ?G = "\x. g (inverse x)" let ?R = "at_right (0::real)" let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" show "LIM x ?R. ?G x :> at_top" using g_0 unfolding filterlim_at_top_to_right . show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" unfolding eventually_at_right_to_top using Dg eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" unfolding eventually_at_right_to_top using Df eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. ?D g' x \ 0) ?R" unfolding eventually_at_right_to_top using g' eventually_ge_at_top[where c=1] by eventually_elim auto show "((\x. ?D f' x / ?D g' x) \ x) ?R" unfolding filterlim_at_right_to_top apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) using eventually_ge_at_top[where c=1] by eventually_elim simp qed lemma lhopital_right_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_right a)" shows "filterlim (\ x. f x / g x) at_top (at_right a)" proof - from lim have pos: "eventually (\x. f' x / g' x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast have "((\x. g x / f x) \ 0) (at_right a)" proof (rule lhopital_right_at_top) from pos show "eventually (\x. f' x \ 0) (at_right a)" by eventually_elim auto from tendsto_inverse_0_at_top[OF lim] show "((\x. g' x / f' x) \ 0) (at_right a)" by simp qed fact+ moreover from f_0 g_0 have "eventually (\x. f x > 0) (at_right a)" "eventually (\x. g x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast+ hence "eventually (\x. g x / f x > 0) (at_right a)" by eventually_elim simp ultimately have "filterlim (\x. inverse (g x / f x)) at_top (at_right a)" by (rule filterlim_inverse_at_top) thus ?thesis by simp qed lemma lhopital_right_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_right a)" shows "filterlim (\ x. f x / g x) at_bot (at_right a)" proof - from ev(2) have ev': "eventually (\x. DERIV (\x. -g x) x :> -g' x) (at_right a)" by eventually_elim (auto intro: derivative_intros) have "filterlim (\x. f x / (-g x)) at_top (at_right a)" by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\x. -g' x"]) (insert assms ev', auto simp: filterlim_uminus_at_bot) hence "filterlim (\x. -(f x / g x)) at_top (at_right a)" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma lhopital_left_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_left a)" shows "filterlim (\ x. f x / g x) at_top (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_top[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_left_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_left a)" shows "filterlim (\ x. f x / g x) at_bot (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_bot[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at a)" shows "filterlim (\ x. f x / g x) at_top (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] lhopital_left_at_top_at_top[of f a g f' g']) lemma lhopital_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at a)" shows "filterlim (\ x. f x / g x) at_bot (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] lhopital_left_at_top_at_bot[of f a g f' g']) end diff --git a/src/HOL/Library/Landau_Symbols.thy b/src/HOL/Library/Landau_Symbols.thy --- a/src/HOL/Library/Landau_Symbols.thy +++ b/src/HOL/Library/Landau_Symbols.thy @@ -1,2188 +1,2198 @@ (* File: Landau_Symbols_Definition.thy Author: Manuel Eberl Landau symbols for reasoning about the asymptotic growth of functions. *) section \Definition of Landau symbols\ theory Landau_Symbols imports Complex_Main begin lemma eventually_subst': "eventually (\x. f x = g x) F \ eventually (\x. P x (f x)) F = eventually (\x. P x (g x)) F" by (rule eventually_subst, erule eventually_rev_mp) simp subsection \Definition of Landau symbols\ text \ Our Landau symbols are sign-oblivious, i.e. any function always has the same growth as its absolute. This has the advantage of making some cancelling rules for sums nicer, but introduces some problems in other places. Nevertheless, we found this definition more convenient to work with. \ definition bigo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1O[_]'(_'))\) where "bigo F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" definition smallo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1o[_]'(_'))\) where "smallo F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" definition bigomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) where "bigomega F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" definition smallomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) where "smallomega F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" definition bigtheta :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) where "bigtheta F g = bigo F g \ bigomega F g" abbreviation bigo_at_top (\(2O'(_'))\) where "O(g) \ bigo at_top g" abbreviation smallo_at_top (\(2o'(_'))\) where "o(g) \ smallo at_top g" abbreviation bigomega_at_top (\(2\'(_'))\) where "\(g) \ bigomega at_top g" abbreviation smallomega_at_top (\(2\'(_'))\) where "\(g) \ smallomega at_top g" abbreviation bigtheta_at_top (\(2\'(_'))\) where "\(g) \ bigtheta at_top g" text \The following is a set of properties that all Landau symbols satisfy.\ named_theorems landau_divide_simps locale landau_symbol = fixes L :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" and L' :: "'c filter \ ('c \ ('b :: real_normed_field)) \ ('c \ 'b) set" and Lr :: "'a filter \ ('a \ real) \ ('a \ real) set" assumes bot': "L bot f = UNIV" assumes filter_mono': "F1 \ F2 \ L F2 f \ L F1 f" assumes in_filtermap_iff: "f' \ L (filtermap h' F') g' \ (\x. f' (h' x)) \ L' F' (\x. g' (h' x))" assumes filtercomap: "f' \ L F'' g' \ (\x. f' (h' x)) \ L' (filtercomap h' F'') (\x. g' (h' x))" assumes sup: "f \ L F1 g \ f \ L F2 g \ f \ L (sup F1 F2) g" assumes in_cong: "eventually (\x. f x = g x) F \ f \ L F (h) \ g \ L F (h)" assumes cong: "eventually (\x. f x = g x) F \ L F (f) = L F (g)" assumes cong_bigtheta: "f \ \[F](g) \ L F (f) = L F (g)" assumes in_cong_bigtheta: "f \ \[F](g) \ f \ L F (h) \ g \ L F (h)" assumes cmult [simp]: "c \ 0 \ L F (\x. c * f x) = L F (f)" assumes cmult_in_iff [simp]: "c \ 0 \ (\x. c * f x) \ L F (g) \ f \ L F (g)" assumes mult_left [simp]: "f \ L F (g) \ (\x. h x * f x) \ L F (\x. h x * g x)" assumes inverse: "eventually (\x. f x \ 0) F \ eventually (\x. g x \ 0) F \ f \ L F (g) \ (\x. inverse (g x)) \ L F (\x. inverse (f x))" assumes subsetI: "f \ L F (g) \ L F (f) \ L F (g)" assumes plus_subset1: "f \ o[F](g) \ L F (g) \ L F (\x. f x + g x)" assumes trans: "f \ L F (g) \ g \ L F (h) \ f \ L F (h)" assumes compose: "f \ L F (g) \ filterlim h' F G \ (\x. f (h' x)) \ L' G (\x. g (h' x))" assumes norm_iff [simp]: "(\x. norm (f x)) \ Lr F (\x. norm (g x)) \ f \ L F (g)" assumes abs [simp]: "Lr Fr (\x. \fr x\) = Lr Fr fr" assumes abs_in_iff [simp]: "(\x. \fr x\) \ Lr Fr gr \ fr \ Lr Fr gr" begin lemma bot [simp]: "f \ L bot g" by (simp add: bot') lemma filter_mono: "F1 \ F2 \ f \ L F2 g \ f \ L F1 g" using filter_mono'[of F1 F2] by blast lemma cong_ex: "eventually (\x. f1 x = f2 x) F \ eventually (\x. g1 x = g2 x) F \ f1 \ L F (g1) \ f2 \ L F (g2)" by (subst cong, assumption, subst in_cong, assumption, rule refl) lemma cong_ex_bigtheta: "f1 \ \[F](f2) \ g1 \ \[F](g2) \ f1 \ L F (g1) \ f2 \ L F (g2)" by (subst cong_bigtheta, assumption, subst in_cong_bigtheta, assumption, rule refl) lemma bigtheta_trans1: "f \ L F (g) \ g \ \[F](h) \ f \ L F (h)" by (subst cong_bigtheta[symmetric]) lemma bigtheta_trans2: "f \ \[F](g) \ g \ L F (h) \ f \ L F (h)" by (subst in_cong_bigtheta) lemma cmult' [simp]: "c \ 0 \ L F (\x. f x * c) = L F (f)" by (subst mult.commute) (rule cmult) lemma cmult_in_iff' [simp]: "c \ 0 \ (\x. f x * c) \ L F (g) \ f \ L F (g)" by (subst mult.commute) (rule cmult_in_iff) lemma cdiv [simp]: "c \ 0 \ L F (\x. f x / c) = L F (f)" using cmult'[of "inverse c" F f] by (simp add: field_simps) lemma cdiv_in_iff' [simp]: "c \ 0 \ (\x. f x / c) \ L F (g) \ f \ L F (g)" using cmult_in_iff'[of "inverse c" f] by (simp add: field_simps) lemma uminus [simp]: "L F (\x. -g x) = L F (g)" using cmult[of "-1"] by simp lemma uminus_in_iff [simp]: "(\x. -f x) \ L F (g) \ f \ L F (g)" using cmult_in_iff[of "-1"] by simp lemma const: "c \ 0 \ L F (\_. c) = L F (\_. 1)" by (subst (2) cmult[symmetric]) simp_all (* Simplifier loops without the NO_MATCH *) lemma const' [simp]: "NO_MATCH 1 c \ c \ 0 \ L F (\_. c) = L F (\_. 1)" by (rule const) lemma const_in_iff: "c \ 0 \ (\_. c) \ L F (f) \ (\_. 1) \ L F (f)" using cmult_in_iff'[of c "\_. 1"] by simp lemma const_in_iff' [simp]: "NO_MATCH 1 c \ c \ 0 \ (\_. c) \ L F (f) \ (\_. 1) \ L F (f)" by (rule const_in_iff) lemma plus_subset2: "g \ o[F](f) \ L F (f) \ L F (\x. f x + g x)" by (subst add.commute) (rule plus_subset1) lemma mult_right [simp]: "f \ L F (g) \ (\x. f x * h x) \ L F (\x. g x * h x)" using mult_left by (simp add: mult.commute) lemma mult: "f1 \ L F (g1) \ f2 \ L F (g2) \ (\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" by (rule trans, erule mult_left, erule mult_right) lemma inverse_cancel: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" shows "(\x. inverse (f x)) \ L F (\x. inverse (g x)) \ g \ L F (f)" proof assume "(\x. inverse (f x)) \ L F (\x. inverse (g x))" from inverse[OF _ _ this] assms show "g \ L F (f)" by simp qed (intro inverse assms) lemma divide_right: assumes "eventually (\x. h x \ 0) F" assumes "f \ L F (g)" shows "(\x. f x / h x) \ L F (\x. g x / h x)" by (subst (1 2) divide_inverse) (intro mult_right inverse assms) lemma divide_right_iff: assumes "eventually (\x. h x \ 0) F" shows "(\x. f x / h x) \ L F (\x. g x / h x) \ f \ L F (g)" proof assume "(\x. f x / h x) \ L F (\x. g x / h x)" from mult_right[OF this, of h] assms show "f \ L F (g)" by (subst (asm) cong_ex[of _ f F _ g]) (auto elim!: eventually_mono) qed (simp add: divide_right assms) lemma divide_left: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" assumes "g \ L F(f)" shows "(\x. h x / f x) \ L F (\x. h x / g x)" by (subst (1 2) divide_inverse) (intro mult_left inverse assms) lemma divide_left_iff: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" assumes "eventually (\x. h x \ 0) F" shows "(\x. h x / f x) \ L F (\x. h x / g x) \ g \ L F (f)" proof assume A: "(\x. h x / f x) \ L F (\x. h x / g x)" from assms have B: "eventually (\x. h x / f x / h x = inverse (f x)) F" by eventually_elim (simp add: divide_inverse) from assms have C: "eventually (\x. h x / g x / h x = inverse (g x)) F" by eventually_elim (simp add: divide_inverse) from divide_right[OF assms(3) A] assms show "g \ L F (f)" by (subst (asm) cong_ex[OF B C]) (simp add: inverse_cancel) qed (simp add: divide_left assms) lemma divide: assumes "eventually (\x. g1 x \ 0) F" assumes "eventually (\x. g2 x \ 0) F" assumes "f1 \ L F (f2)" "g2 \ L F (g1)" shows "(\x. f1 x / g1 x) \ L F (\x. f2 x / g2 x)" by (subst (1 2) divide_inverse) (intro mult inverse assms) lemma divide_eq1: assumes "eventually (\x. h x \ 0) F" shows "f \ L F (\x. g x / h x) \ (\x. f x * h x) \ L F (g)" proof- have "f \ L F (\x. g x / h x) \ (\x. f x * h x / h x) \ L F (\x. g x / h x)" using assms by (intro in_cong) (auto elim: eventually_mono) thus ?thesis by (simp only: divide_right_iff assms) qed lemma divide_eq2: assumes "eventually (\x. h x \ 0) F" shows "(\x. f x / h x) \ L F (\x. g x) \ f \ L F (\x. g x * h x)" proof- have "L F (\x. g x) = L F (\x. g x * h x / h x)" using assms by (intro cong) (auto elim: eventually_mono) thus ?thesis by (simp only: divide_right_iff assms) qed lemma inverse_eq1: assumes "eventually (\x. g x \ 0) F" shows "f \ L F (\x. inverse (g x)) \ (\x. f x * g x) \ L F (\_. 1)" using divide_eq1[of g F f "\_. 1"] by (simp add: divide_inverse assms) lemma inverse_eq2: assumes "eventually (\x. f x \ 0) F" shows "(\x. inverse (f x)) \ L F (g) \ (\x. 1) \ L F (\x. f x * g x)" using divide_eq2[of f F "\_. 1" g] by (simp add: divide_inverse assms mult_ac) lemma inverse_flip: assumes "eventually (\x. g x \ 0) F" assumes "eventually (\x. h x \ 0) F" assumes "(\x. inverse (g x)) \ L F (h)" shows "(\x. inverse (h x)) \ L F (g)" using assms by (simp add: divide_eq1 divide_eq2 inverse_eq_divide mult.commute) lemma lift_trans: assumes "f \ L F (g)" assumes "(\x. t x (g x)) \ L F (h)" assumes "\f g. f \ L F (g) \ (\x. t x (f x)) \ L F (\x. t x (g x))" shows "(\x. t x (f x)) \ L F (h)" by (rule trans[OF assms(3)[OF assms(1)] assms(2)]) lemma lift_trans': assumes "f \ L F (\x. t x (g x))" assumes "g \ L F (h)" assumes "\g h. g \ L F (h) \ (\x. t x (g x)) \ L F (\x. t x (h x))" shows "f \ L F (\x. t x (h x))" by (rule trans[OF assms(1) assms(3)[OF assms(2)]]) lemma lift_trans_bigtheta: assumes "f \ L F (g)" assumes "(\x. t x (g x)) \ \[F](h)" assumes "\f g. f \ L F (g) \ (\x. t x (f x)) \ L F (\x. t x (g x))" shows "(\x. t x (f x)) \ L F (h)" using cong_bigtheta[OF assms(2)] assms(3)[OF assms(1)] by simp lemma lift_trans_bigtheta': assumes "f \ L F (\x. t x (g x))" assumes "g \ \[F](h)" assumes "\g h. g \ \[F](h) \ (\x. t x (g x)) \ \[F](\x. t x (h x))" shows "f \ L F (\x. t x (h x))" using cong_bigtheta[OF assms(3)[OF assms(2)]] assms(1) by simp lemma (in landau_symbol) mult_in_1: assumes "f \ L F (\_. 1)" "g \ L F (\_. 1)" shows "(\x. f x * g x) \ L F (\_. 1)" using mult[OF assms] by simp lemma (in landau_symbol) of_real_cancel: "(\x. of_real (f x)) \ L F (\x. of_real (g x)) \ f \ Lr F g" by (subst (asm) norm_iff [symmetric], subst (asm) (1 2) norm_of_real) simp_all lemma (in landau_symbol) of_real_iff: "(\x. of_real (f x)) \ L F (\x. of_real (g x)) \ f \ Lr F g" by (subst norm_iff [symmetric], subst (1 2) norm_of_real) simp_all lemmas [landau_divide_simps] = inverse_cancel divide_left_iff divide_eq1 divide_eq2 inverse_eq1 inverse_eq2 end text \ The symbols $O$ and $o$ and $\Omega$ and $\omega$ are dual, so for many rules, replacing $O$ with $\Omega$, $o$ with $\omega$, and $\leq$ with $\geq$ in a theorem yields another valid theorem. The following locale captures this fact. \ locale landau_pair = fixes L l :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" fixes L' l' :: "'c filter \ ('c \ ('b :: real_normed_field)) \ ('c \ 'b) set" fixes Lr lr :: "'a filter \ ('a \ real) \ ('a \ real) set" and R :: "real \ real \ bool" assumes L_def: "L F g = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g x))) F}" and l_def: "l F g = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g x))) F}" and L'_def: "L' F' g' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g' x))) F'}" and l'_def: "l' F' g' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g' x))) F'}" and Lr_def: "Lr F'' g'' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g'' x))) F''}" and lr_def: "lr F'' g'' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g'' x))) F''}" and R: "R = (\) \ R = (\)" interpretation landau_o: landau_pair bigo smallo bigo smallo bigo smallo "(\)" by unfold_locales (auto simp: bigo_def smallo_def intro!: ext) interpretation landau_omega: landau_pair bigomega smallomega bigomega smallomega bigomega smallomega "(\)" by unfold_locales (auto simp: bigomega_def smallomega_def intro!: ext) context landau_pair begin lemmas R_E = disjE [OF R, case_names le ge] lemma bigI: "c > 0 \ eventually (\x. R (norm (f x)) (c * norm (g x))) F \ f \ L F (g)" unfolding L_def by blast lemma bigE: assumes "f \ L F (g)" obtains c where "c > 0" "eventually (\x. R (norm (f x)) (c * (norm (g x)))) F" using assms unfolding L_def by blast lemma smallI: "(\c. c > 0 \ eventually (\x. R (norm (f x)) (c * (norm (g x)))) F) \ f \ l F (g)" unfolding l_def by blast lemma smallD: "f \ l F (g) \ c > 0 \ eventually (\x. R (norm (f x)) (c * (norm (g x)))) F" unfolding l_def by blast lemma bigE_nonneg_real: assumes "f \ Lr F (g)" "eventually (\x. f x \ 0) F" obtains c where "c > 0" "eventually (\x. R (f x) (c * \g x\)) F" proof- from assms(1) obtain c where c: "c > 0" "eventually (\x. R (norm (f x)) (c * norm (g x))) F" by (auto simp: Lr_def) from c(2) assms(2) have "eventually (\x. R (f x) (c * \g x\)) F" by eventually_elim simp from c(1) and this show ?thesis by (rule that) qed lemma smallD_nonneg_real: assumes "f \ lr F (g)" "eventually (\x. f x \ 0) F" "c > 0" shows "eventually (\x. R (f x) (c * \g x\)) F" using assms by (auto simp: lr_def dest!: spec[of _ c] elim: eventually_elim2) lemma small_imp_big: "f \ l F (g) \ f \ L F (g)" by (rule bigI[OF _ smallD, of 1]) simp_all lemma small_subset_big: "l F (g) \ L F (g)" using small_imp_big by blast lemma R_refl [simp]: "R x x" using R by auto lemma R_linear: "\R x y \ R y x" using R by auto lemma R_trans [trans]: "R a b \ R b c \ R a c" using R by auto lemma R_mult_left_mono: "R a b \ c \ 0 \ R (c*a) (c*b)" using R by (auto simp: mult_left_mono) lemma R_mult_right_mono: "R a b \ c \ 0 \ R (a*c) (b*c)" using R by (auto simp: mult_right_mono) lemma big_trans: assumes "f \ L F (g)" "g \ L F (h)" shows "f \ L F (h)" proof- from assms obtain c d where *: "0 < c" "0 < d" and **: "\\<^sub>F x in F. R (norm (f x)) (c * norm (g x))" "\\<^sub>F x in F. R (norm (g x)) (d * norm (h x))" by (elim bigE) from ** have "eventually (\x. R (norm (f x)) (c * d * (norm (h x)))) F" proof eventually_elim fix x assume "R (norm (f x)) (c * (norm (g x)))" also assume "R (norm (g x)) (d * (norm (h x)))" with \0 < c\ have "R (c * (norm (g x))) (c * (d * (norm (h x))))" by (intro R_mult_left_mono) simp_all finally show "R (norm (f x)) (c * d * (norm (h x)))" by (simp add: algebra_simps) qed with * show ?thesis by (intro bigI[of "c*d"]) simp_all qed lemma big_small_trans: assumes "f \ L F (g)" "g \ l F (h)" shows "f \ l F (h)" proof (rule smallI) fix c :: real assume c: "c > 0" from assms(1) obtain d where d: "d > 0" and *: "\\<^sub>F x in F. R (norm (f x)) (d * norm (g x))" by (elim bigE) from assms(2) c d have "eventually (\x. R (norm (g x)) (c * inverse d * norm (h x))) F" by (intro smallD) simp_all with * show "eventually (\x. R (norm (f x)) (c * (norm (h x)))) F" proof eventually_elim case (elim x) show ?case by (use elim(1) in \rule R_trans\) (use elim(2) R d in \auto simp: field_simps\) qed qed lemma small_big_trans: assumes "f \ l F (g)" "g \ L F (h)" shows "f \ l F (h)" proof (rule smallI) fix c :: real assume c: "c > 0" from assms(2) obtain d where d: "d > 0" and *: "\\<^sub>F x in F. R (norm (g x)) (d * norm (h x))" by (elim bigE) from assms(1) c d have "eventually (\x. R (norm (f x)) (c * inverse d * norm (g x))) F" by (intro smallD) simp_all with * show "eventually (\x. R (norm (f x)) (c * norm (h x))) F" by eventually_elim (rotate_tac 2, erule R_trans, insert R c d, auto simp: field_simps) qed lemma small_trans: "f \ l F (g) \ g \ l F (h) \ f \ l F (h)" by (rule big_small_trans[OF small_imp_big]) lemma small_big_trans': "f \ l F (g) \ g \ L F (h) \ f \ L F (h)" by (rule small_imp_big[OF small_big_trans]) lemma big_small_trans': "f \ L F (g) \ g \ l F (h) \ f \ L F (h)" by (rule small_imp_big[OF big_small_trans]) lemma big_subsetI [intro]: "f \ L F (g) \ L F (f) \ L F (g)" by (intro subsetI) (drule (1) big_trans) lemma small_subsetI [intro]: "f \ L F (g) \ l F (f) \ l F (g)" by (intro subsetI) (drule (1) small_big_trans) lemma big_refl [simp]: "f \ L F (f)" by (rule bigI[of 1]) simp_all lemma small_refl_iff: "f \ l F (f) \ eventually (\x. f x = 0) F" proof (rule iffI[OF _ smallI]) assume f: "f \ l F f" have "(1/2::real) > 0" "(2::real) > 0" by simp_all from smallD[OF f this(1)] smallD[OF f this(2)] show "eventually (\x. f x = 0) F" by eventually_elim (insert R, auto) next fix c :: real assume "c > 0" "eventually (\x. f x = 0) F" from this(2) show "eventually (\x. R (norm (f x)) (c * norm (f x))) F" by eventually_elim simp_all qed lemma big_small_asymmetric: "f \ L F (g) \ g \ l F (f) \ eventually (\x. f x = 0) F" by (drule (1) big_small_trans) (simp add: small_refl_iff) lemma small_big_asymmetric: "f \ l F (g) \ g \ L F (f) \ eventually (\x. f x = 0) F" by (drule (1) small_big_trans) (simp add: small_refl_iff) lemma small_asymmetric: "f \ l F (g) \ g \ l F (f) \ eventually (\x. f x = 0) F" by (drule (1) small_trans) (simp add: small_refl_iff) lemma plus_aux: assumes "f \ o[F](g)" shows "g \ L F (\x. f x + g x)" proof (rule R_E) assume R: "R = (\)" have A: "1/2 > (0::real)" by simp have B: "1/2 * (norm (g x)) \ norm (f x + g x)" if "norm (f x) \ 1/2 * norm (g x)" for x proof - from that have "1/2 * (norm (g x)) \ (norm (g x)) - (norm (f x))" by simp also have "norm (g x) - norm (f x) \ norm (f x + g x)" by (subst add.commute) (rule norm_diff_ineq) finally show ?thesis by simp qed show "g \ L F (\x. f x + g x)" apply (rule bigI[of "2"]) apply simp apply (use landau_o.smallD[OF assms A] in eventually_elim) apply (use B in \simp add: R algebra_simps\) done next assume R: "R = (\x y. x \ y)" show "g \ L F (\x. f x + g x)" proof (rule bigI[of "1/2"]) show "eventually (\x. R (norm (g x)) (1/2 * norm (f x + g x))) F" using landau_o.smallD[OF assms zero_less_one] proof eventually_elim case (elim x) have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also note elim finally show ?case by (simp add: R) qed qed simp_all qed end +lemma summable_comparison_test_bigo: + fixes f :: "nat \ real" + assumes "summable (\n. norm (g n))" "f \ O(g)" + shows "summable f" +proof - + from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" + by (auto elim: landau_o.bigE) + thus ?thesis + by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) +qed lemma bigomega_iff_bigo: "g \ \[F](f) \ f \ O[F](g)" proof assume "f \ O[F](g)" then obtain c where "0 < c" "\\<^sub>F x in F. norm (f x) \ c * norm (g x)" by (rule landau_o.bigE) then show "g \ \[F](f)" by (intro landau_omega.bigI[of "inverse c"]) (simp_all add: field_simps) next assume "g \ \[F](f)" then obtain c where "0 < c" "\\<^sub>F x in F. c * norm (f x) \ norm (g x)" by (rule landau_omega.bigE) then show "f \ O[F](g)" by (intro landau_o.bigI[of "inverse c"]) (simp_all add: field_simps) qed lemma smallomega_iff_smallo: "g \ \[F](f) \ f \ o[F](g)" proof assume "f \ o[F](g)" from landau_o.smallD[OF this, of "inverse c" for c] show "g \ \[F](f)" by (intro landau_omega.smallI) (simp_all add: field_simps) next assume "g \ \[F](f)" from landau_omega.smallD[OF this, of "inverse c" for c] show "f \ o[F](g)" by (intro landau_o.smallI) (simp_all add: field_simps) qed context landau_pair begin lemma big_mono: "eventually (\x. R (norm (f x)) (norm (g x))) F \ f \ L F (g)" by (rule bigI[OF zero_less_one]) simp lemma big_mult: assumes "f1 \ L F (g1)" "f2 \ L F (g2)" shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" proof- from assms obtain c1 c2 where *: "c1 > 0" "c2 > 0" and **: "\\<^sub>F x in F. R (norm (f1 x)) (c1 * norm (g1 x))" "\\<^sub>F x in F. R (norm (f2 x)) (c2 * norm (g2 x))" by (elim bigE) from * have "c1 * c2 > 0" by simp moreover have "eventually (\x. R (norm (f1 x * f2 x)) (c1 * c2 * norm (g1 x * g2 x))) F" using ** proof eventually_elim case (elim x) show ?case proof (cases rule: R_E) case le have "norm (f1 x) * norm (f2 x) \ (c1 * norm (g1 x)) * (c2 * norm (g2 x))" using elim le * by (intro mult_mono mult_nonneg_nonneg) auto with le show ?thesis by (simp add: le norm_mult mult_ac) next case ge have "(c1 * norm (g1 x)) * (c2 * norm (g2 x)) \ norm (f1 x) * norm (f2 x)" using elim ge * by (intro mult_mono mult_nonneg_nonneg) auto with ge show ?thesis by (simp_all add: norm_mult mult_ac) qed qed ultimately show ?thesis by (rule bigI) qed lemma small_big_mult: assumes "f1 \ l F (g1)" "f2 \ L F (g2)" shows "(\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" proof (rule smallI) fix c1 :: real assume c1: "c1 > 0" from assms(2) obtain c2 where c2: "c2 > 0" and *: "\\<^sub>F x in F. R (norm (f2 x)) (c2 * norm (g2 x))" by (elim bigE) from assms(1) c1 c2 have "eventually (\x. R (norm (f1 x)) (c1 * inverse c2 * norm (g1 x))) F" by (auto intro!: smallD) with * show "eventually (\x. R (norm (f1 x * f2 x)) (c1 * norm (g1 x * g2 x))) F" proof eventually_elim case (elim x) show ?case proof (cases rule: R_E) case le have "norm (f1 x) * norm (f2 x) \ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))" using elim le c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto with le c2 show ?thesis by (simp add: le norm_mult field_simps) next case ge have "norm (f1 x) * norm (f2 x) \ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))" using elim ge c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto with ge c2 show ?thesis by (simp add: ge norm_mult field_simps) qed qed qed lemma big_small_mult: "f1 \ L F (g1) \ f2 \ l F (g2) \ (\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" by (subst (1 2) mult.commute) (rule small_big_mult) lemma small_mult: "f1 \ l F (g1) \ f2 \ l F (g2) \ (\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" by (rule small_big_mult, assumption, rule small_imp_big) lemmas mult = big_mult small_big_mult big_small_mult small_mult sublocale big: landau_symbol L L' Lr proof have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff) have A: "(\x. c * f x) \ L F f" if "c \ 0" for c :: 'b and F and f :: "'a \ 'b" using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult) show "L F (\x. c * f x) = L F f" if "c \ 0" for c :: 'b and F and f :: "'a \ 'b" using \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] by (intro equalityI big_subsetI) (simp_all add: field_simps) show "((\x. c * f x) \ L F g) = (f \ L F g)" if "c \ 0" for c :: 'b and F and f g :: "'a \ 'b" proof - from \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" by (simp_all add: field_simps) then show ?thesis by (intro iffI) (erule (1) big_trans)+ qed show "(\x. inverse (g x)) \ L F (\x. inverse (f x))" if *: "f \ L F (g)" and **: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" for f g :: "'a \ 'b" and F proof - from * obtain c where c: "c > 0" and ***: "\\<^sub>F x in F. R (norm (f x)) (c * norm (g x))" by (elim bigE) from ** *** have "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide c) with c show ?thesis by (rule bigI) qed show "L F g \ L F (\x. f x + g x)" if "f \ o[F](g)" for f g :: "'a \ 'b" and F using plus_aux that by (blast intro!: big_subsetI) show "L F (f) = L F (g)" if "eventually (\x. f x = g x) F" for f g :: "'a \ 'b" and F unfolding L_def by (subst eventually_subst'[OF that]) (rule refl) show "f \ L F (h) \ g \ L F (h)" if "eventually (\x. f x = g x) F" for f g h :: "'a \ 'b" and F unfolding L_def mem_Collect_eq by (subst (1) eventually_subst'[OF that]) (rule refl) show "L F f \ L F g" if "f \ L F g" for f g :: "'a \ 'b" and F using that by (rule big_subsetI) show "L F (f) = L F (g)" if "f \ \[F](g)" for f g :: "'a \ 'b" and F using L that unfolding bigtheta_def by (intro equalityI big_subsetI) (auto simp: bigomega_iff_bigo) show "f \ L F (h) \ g \ L F (h)" if "f \ \[F](g)" for f g h :: "'a \ 'b" and F by (rule disjE[OF L]) (use that in \auto simp: bigtheta_def bigomega_iff_bigo intro: landau_o.big_trans\) show "(\x. h x * f x) \ L F (\x. h x * g x)" if "f \ L F g" for f g h :: "'a \ 'b" and F using that by (intro big_mult) simp show "f \ L F (h)" if "f \ L F g" "g \ L F h" for f g h :: "'a \ 'b" and F using that by (rule big_trans) show "(\x. f (h x)) \ L' G (\x. g (h x))" if "f \ L F g" and "filterlim h F G" for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G using that by (auto simp: L_def L'_def filterlim_iff) show "f \ L (sup F G) g" if "f \ L F g" "f \ L G g" for f g :: "'a \ 'b" and F G :: "'a filter" proof - from that [THEN bigE] obtain c1 c2 where *: "c1 > 0" "c2 > 0" and **: "\\<^sub>F x in F. R (norm (f x)) (c1 * norm (g x))" "\\<^sub>F x in G. R (norm (f x)) (c2 * norm (g x))" . define c where "c = (if R c1 c2 then c2 else c1)" from * have c: "R c1 c" "R c2 c" "c > 0" by (auto simp: c_def dest: R_linear) with ** have "eventually (\x. R (norm (f x)) (c * norm (g x))) F" "eventually (\x. R (norm (f x)) (c * norm (g x))) G" by (force elim: eventually_mono intro: R_trans[OF _ R_mult_right_mono])+ with c show "f \ L (sup F G) g" by (auto simp: L_def eventually_sup) qed show "((\x. f (h x)) \ L' (filtercomap h F) (\x. g (h x)))" if "(f \ L F g)" for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" using that unfolding L_def L'_def by auto qed (auto simp: L_def Lr_def eventually_filtermap L'_def intro: filter_leD exI[of _ "1::real"]) sublocale small: landau_symbol l l' lr proof have A: "(\x. c * f x) \ L F f" if "c \ 0" for c :: 'b and f :: "'a \ 'b" and F using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult) show "l F (\x. c * f x) = l F f" if "c \ 0" for c :: 'b and f :: "'a \ 'b" and F using that and A[of c f] and A[of "inverse c" "\x. c * f x"] by (intro equalityI small_subsetI) (simp_all add: field_simps) show "((\x. c * f x) \ l F g) = (f \ l F g)" if "c \ 0" for c :: 'b and f g :: "'a \ 'b" and F proof - from that and A[of c f] and A[of "inverse c" "\x. c * f x"] have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" by (simp_all add: field_simps) then show ?thesis by (intro iffI) (erule (1) big_small_trans)+ qed show "l F g \ l F (\x. f x + g x)" if "f \ o[F](g)" for f g :: "'a \ 'b" and F using plus_aux that by (blast intro!: small_subsetI) show "(\x. inverse (g x)) \ l F (\x. inverse (f x))" if A: "f \ l F (g)" and B: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" for f g :: "'a \ 'b" and F proof (rule smallI) fix c :: real assume c: "c > 0" from B smallD[OF A c] show "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide) qed show "l F (f) = l F (g)" if "eventually (\x. f x = g x) F" for f g :: "'a \ 'b" and F unfolding l_def by (subst eventually_subst'[OF that]) (rule refl) show "f \ l F (h) \ g \ l F (h)" if "eventually (\x. f x = g x) F" for f g h :: "'a \ 'b" and F unfolding l_def mem_Collect_eq by (subst (1) eventually_subst'[OF that]) (rule refl) show "l F f \ l F g" if "f \ l F g" for f g :: "'a \ 'b" and F using that by (intro small_subsetI small_imp_big) show "l F (f) = l F (g)" if "f \ \[F](g)" for f g :: "'a \ 'b" and F proof - have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff) with that show ?thesis unfolding bigtheta_def by (intro equalityI small_subsetI) (auto simp: bigomega_iff_bigo) qed show "f \ l F (h) \ g \ l F (h)" if "f \ \[F](g)" for f g h :: "'a \ 'b" and F proof - have l: "l = smallo \ l = smallomega" by (rule R_E) (auto simp: smallo_def l_def smallomega_def fun_eq_iff) show ?thesis by (rule disjE[OF l]) (use that in \auto simp: bigtheta_def bigomega_iff_bigo smallomega_iff_smallo intro: landau_o.big_small_trans landau_o.small_big_trans\) qed show "(\x. h x * f x) \ l F (\x. h x * g x)" if "f \ l F g" for f g h :: "'a \ 'b" and F using that by (intro big_small_mult) simp show "f \ l F (h)" if "f \ l F g" "g \ l F h" for f g h :: "'a \ 'b" and F using that by (rule small_trans) show "(\x. f (h x)) \ l' G (\x. g (h x))" if "f \ l F g" and "filterlim h F G" for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G using that by (auto simp: l_def l'_def filterlim_iff) show "((\x. f (h x)) \ l' (filtercomap h F) (\x. g (h x)))" if "f \ l F g" for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" using that unfolding l_def l'_def by auto qed (auto simp: l_def lr_def eventually_filtermap l'_def eventually_sup intro: filter_leD) text \These rules allow chaining of Landau symbol propositions in Isar with "also".\ lemma big_mult_1: "f \ L F (g) \ (\_. 1) \ L F (h) \ f \ L F (\x. g x * h x)" and big_mult_1': "(\_. 1) \ L F (g) \ f \ L F (h) \ f \ L F (\x. g x * h x)" and small_mult_1: "f \ l F (g) \ (\_. 1) \ L F (h) \ f \ l F (\x. g x * h x)" and small_mult_1': "(\_. 1) \ L F (g) \ f \ l F (h) \ f \ l F (\x. g x * h x)" and small_mult_1'': "f \ L F (g) \ (\_. 1) \ l F (h) \ f \ l F (\x. g x * h x)" and small_mult_1''': "(\_. 1) \ l F (g) \ f \ L F (h) \ f \ l F (\x. g x * h x)" by (drule (1) big.mult big_small_mult small_big_mult, simp)+ lemma big_1_mult: "f \ L F (g) \ h \ L F (\_. 1) \ (\x. f x * h x) \ L F (g)" and big_1_mult': "h \ L F (\_. 1) \ f \ L F (g) \ (\x. f x * h x) \ L F (g)" and small_1_mult: "f \ l F (g) \ h \ L F (\_. 1) \ (\x. f x * h x) \ l F (g)" and small_1_mult': "h \ L F (\_. 1) \ f \ l F (g) \ (\x. f x * h x) \ l F (g)" and small_1_mult'': "f \ L F (g) \ h \ l F (\_. 1) \ (\x. f x * h x) \ l F (g)" and small_1_mult''': "h \ l F (\_. 1) \ f \ L F (g) \ (\x. f x * h x) \ l F (g)" by (drule (1) big.mult big_small_mult small_big_mult, simp)+ lemmas mult_1_trans = big_mult_1 big_mult_1' small_mult_1 small_mult_1' small_mult_1'' small_mult_1''' big_1_mult big_1_mult' small_1_mult small_1_mult' small_1_mult'' small_1_mult''' lemma big_equal_iff_bigtheta: "L F (f) = L F (g) \ f \ \[F](g)" proof have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: fun_eq_iff L_def bigo_def bigomega_def) fix f g :: "'a \ 'b" assume "L F (f) = L F (g)" with big_refl[of f F] big_refl[of g F] have "f \ L F (g)" "g \ L F (f)" by simp_all thus "f \ \[F](g)" using L unfolding bigtheta_def by (auto simp: bigomega_iff_bigo) qed (rule big.cong_bigtheta) lemma big_prod: assumes "\x. x \ A \ f x \ L F (g x)" shows "(\y. \x\A. f x y) \ L F (\y. \x\A. g x y)" using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult) lemma big_prod_in_1: assumes "\x. x \ A \ f x \ L F (\_. 1)" shows "(\y. \x\A. f x y) \ L F (\_. 1)" using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult_in_1) end context landau_symbol begin lemma plus_absorb1: assumes "f \ o[F](g)" shows "L F (\x. f x + g x) = L F (g)" proof (intro equalityI) from plus_subset1 and assms show "L F g \ L F (\x. f x + g x)" . from landau_o.small.plus_subset1[OF assms] and assms have "(\x. -f x) \ o[F](\x. f x + g x)" by (auto simp: landau_o.small.uminus_in_iff) from plus_subset1[OF this] show "L F (\x. f x + g x) \ L F (g)" by simp qed lemma plus_absorb2: "g \ o[F](f) \ L F (\x. f x + g x) = L F (f)" using plus_absorb1[of g F f] by (simp add: add.commute) lemma diff_absorb1: "f \ o[F](g) \ L F (\x. f x - g x) = L F (g)" by (simp only: diff_conv_add_uminus plus_absorb1 landau_o.small.uminus uminus) lemma diff_absorb2: "g \ o[F](f) \ L F (\x. f x - g x) = L F (f)" by (simp only: diff_conv_add_uminus plus_absorb2 landau_o.small.uminus_in_iff) lemmas absorb = plus_absorb1 plus_absorb2 diff_absorb1 diff_absorb2 end lemma bigthetaI [intro]: "f \ O[F](g) \ f \ \[F](g) \ f \ \[F](g)" unfolding bigtheta_def bigomega_def by blast lemma bigthetaD1 [dest]: "f \ \[F](g) \ f \ O[F](g)" and bigthetaD2 [dest]: "f \ \[F](g) \ f \ \[F](g)" unfolding bigtheta_def bigo_def bigomega_def by blast+ lemma bigtheta_refl [simp]: "f \ \[F](f)" unfolding bigtheta_def by simp lemma bigtheta_sym: "f \ \[F](g) \ g \ \[F](f)" unfolding bigtheta_def by (auto simp: bigomega_iff_bigo) lemmas landau_flip = bigomega_iff_bigo[symmetric] smallomega_iff_smallo[symmetric] bigomega_iff_bigo smallomega_iff_smallo bigtheta_sym interpretation landau_theta: landau_symbol bigtheta bigtheta bigtheta proof fix f g :: "'a \ 'b" and F assume "f \ o[F](g)" hence "O[F](g) \ O[F](\x. f x + g x)" "\[F](g) \ \[F](\x. f x + g x)" by (rule landau_o.big.plus_subset1 landau_omega.big.plus_subset1)+ thus "\[F](g) \ \[F](\x. f x + g x)" unfolding bigtheta_def by blast next fix f g :: "'a \ 'b" and F assume "f \ \[F](g)" thus A: "\[F](f) = \[F](g)" apply (subst (1 2) bigtheta_def) apply (subst landau_o.big.cong_bigtheta landau_omega.big.cong_bigtheta, assumption)+ apply (rule refl) done thus "\[F](f) \ \[F](g)" by simp fix h :: "'a \ 'b" show "f \ \[F](h) \ g \ \[F](h)" by (subst (1 2) bigtheta_sym) (simp add: A) next fix f g h :: "'a \ 'b" and F assume "f \ \[F](g)" "g \ \[F](h)" thus "f \ \[F](h)" unfolding bigtheta_def by (blast intro: landau_o.big.trans landau_omega.big.trans) next fix f :: "'a \ 'b" and F1 F2 :: "'a filter" assume "F1 \ F2" thus "\[F2](f) \ \[F1](f)" by (auto simp: bigtheta_def intro: landau_o.big.filter_mono landau_omega.big.filter_mono) qed (auto simp: bigtheta_def landau_o.big.norm_iff landau_o.big.cmult landau_omega.big.cmult landau_o.big.cmult_in_iff landau_omega.big.cmult_in_iff landau_o.big.in_cong landau_omega.big.in_cong landau_o.big.mult landau_omega.big.mult landau_o.big.inverse landau_omega.big.inverse landau_o.big.compose landau_omega.big.compose landau_o.big.bot' landau_omega.big.bot' landau_o.big.in_filtermap_iff landau_omega.big.in_filtermap_iff landau_o.big.sup landau_omega.big.sup landau_o.big.filtercomap landau_omega.big.filtercomap dest: landau_o.big.cong landau_omega.big.cong) lemmas landau_symbols = landau_o.big.landau_symbol_axioms landau_o.small.landau_symbol_axioms landau_omega.big.landau_symbol_axioms landau_omega.small.landau_symbol_axioms landau_theta.landau_symbol_axioms lemma bigoI [intro]: assumes "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" shows "f \ O[F](g)" proof (rule landau_o.bigI) show "max 1 c > 0" by simp have "c * (norm (g x)) \ max 1 c * (norm (g x))" for x by (simp add: mult_right_mono) with assms show "eventually (\x. (norm (f x)) \ max 1 c * (norm (g x))) F" by (auto elim!: eventually_mono dest: order.trans) qed lemma smallomegaD [dest]: assumes "f \ \[F](g)" shows "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" proof (cases "c > 0") case False show ?thesis by (intro always_eventually allI, rule order.trans[of _ 0]) (insert False, auto intro!: mult_nonpos_nonneg) qed (blast dest: landau_omega.smallD[OF assms, of c]) lemma bigthetaI': assumes "c1 > 0" "c2 > 0" assumes "eventually (\x. c1 * (norm (g x)) \ (norm (f x)) \ (norm (f x)) \ c2 * (norm (g x))) F" shows "f \ \[F](g)" apply (rule bigthetaI) apply (rule landau_o.bigI[OF assms(2)]) using assms(3) apply (eventually_elim, simp) apply (rule landau_omega.bigI[OF assms(1)]) using assms(3) apply (eventually_elim, simp) done lemma bigthetaI_cong: "eventually (\x. f x = g x) F \ f \ \[F](g)" by (intro bigthetaI'[of 1 1]) (auto elim!: eventually_mono) lemma (in landau_symbol) ev_eq_trans1: "f \ L F (\x. g x (h x)) \ eventually (\x. h x = h' x) F \ f \ L F (\x. g x (h' x))" by (rule bigtheta_trans1[OF _ bigthetaI_cong]) (auto elim!: eventually_mono) lemma (in landau_symbol) ev_eq_trans2: "eventually (\x. f x = f' x) F \ (\x. g x (f' x)) \ L F (h) \ (\x. g x (f x)) \ L F (h)" by (rule bigtheta_trans2[OF bigthetaI_cong]) (auto elim!: eventually_mono) declare landau_o.smallI landau_omega.bigI landau_omega.smallI [intro] declare landau_o.bigE landau_omega.bigE [elim] declare landau_o.smallD lemma (in landau_symbol) bigtheta_trans1': "f \ L F (g) \ h \ \[F](g) \ f \ L F (h)" by (subst cong_bigtheta[symmetric]) (simp add: bigtheta_sym) lemma (in landau_symbol) bigtheta_trans2': "g \ \[F](f) \ g \ L F (h) \ f \ L F (h)" by (rule bigtheta_trans2, subst bigtheta_sym) lemma bigo_bigomega_trans: "f \ O[F](g) \ h \ \[F](g) \ f \ O[F](h)" and bigo_smallomega_trans: "f \ O[F](g) \ h \ \[F](g) \ f \ o[F](h)" and smallo_bigomega_trans: "f \ o[F](g) \ h \ \[F](g) \ f \ o[F](h)" and smallo_smallomega_trans: "f \ o[F](g) \ h \ \[F](g) \ f \ o[F](h)" and bigomega_bigo_trans: "f \ \[F](g) \ h \ O[F](g) \ f \ \[F](h)" and bigomega_smallo_trans: "f \ \[F](g) \ h \ o[F](g) \ f \ \[F](h)" and smallomega_bigo_trans: "f \ \[F](g) \ h \ O[F](g) \ f \ \[F](h)" and smallomega_smallo_trans: "f \ \[F](g) \ h \ o[F](g) \ f \ \[F](h)" by (unfold bigomega_iff_bigo smallomega_iff_smallo) (erule (1) landau_o.big_trans landau_o.big_small_trans landau_o.small_big_trans landau_o.big_trans landau_o.small_trans)+ lemmas landau_trans_lift [trans] = landau_symbols[THEN landau_symbol.lift_trans] landau_symbols[THEN landau_symbol.lift_trans'] landau_symbols[THEN landau_symbol.lift_trans_bigtheta] landau_symbols[THEN landau_symbol.lift_trans_bigtheta'] lemmas landau_mult_1_trans [trans] = landau_o.mult_1_trans landau_omega.mult_1_trans lemmas landau_trans [trans] = landau_symbols[THEN landau_symbol.bigtheta_trans1] landau_symbols[THEN landau_symbol.bigtheta_trans2] landau_symbols[THEN landau_symbol.bigtheta_trans1'] landau_symbols[THEN landau_symbol.bigtheta_trans2'] landau_symbols[THEN landau_symbol.ev_eq_trans1] landau_symbols[THEN landau_symbol.ev_eq_trans2] landau_o.big_trans landau_o.small_trans landau_o.small_big_trans landau_o.big_small_trans landau_omega.big_trans landau_omega.small_trans landau_omega.small_big_trans landau_omega.big_small_trans bigo_bigomega_trans bigo_smallomega_trans smallo_bigomega_trans smallo_smallomega_trans bigomega_bigo_trans bigomega_smallo_trans smallomega_bigo_trans smallomega_smallo_trans lemma bigtheta_inverse [simp]: shows "(\x. inverse (f x)) \ \[F](\x. inverse (g x)) \ f \ \[F](g)" proof - have "(\x. inverse (f x)) \ O[F](\x. inverse (g x))" if A: "f \ \[F](g)" for f g :: "'a \ 'b" and F proof - from A obtain c1 c2 :: real where *: "c1 > 0" "c2 > 0" and **: "\\<^sub>F x in F. norm (f x) \ c1 * norm (g x)" "\\<^sub>F x in F. c2 * norm (g x) \ norm (f x)" unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) from \c2 > 0\ have c2: "inverse c2 > 0" by simp from ** have "eventually (\x. norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))) F" proof eventually_elim fix x assume A: "norm (f x) \ c1 * norm (g x)" "c2 * norm (g x) \ norm (f x)" from A * have "f x = 0 \ g x = 0" by (auto simp: field_simps mult_le_0_iff) with A * show "norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))" by (force simp: field_simps norm_inverse norm_divide) qed with c2 show ?thesis by (rule landau_o.bigI) qed then show ?thesis unfolding bigtheta_def by (force simp: bigomega_iff_bigo bigtheta_sym) qed lemma bigtheta_divide: assumes "f1 \ \(f2)" "g1 \ \(g2)" shows "(\x. f1 x / g1 x) \ \(\x. f2 x / g2 x)" by (subst (1 2) divide_inverse, intro landau_theta.mult) (simp_all add: bigtheta_inverse assms) lemma eventually_nonzero_bigtheta: assumes "f \ \[F](g)" shows "eventually (\x. f x \ 0) F \ eventually (\x. g x \ 0) F" proof - have "eventually (\x. g x \ 0) F" if A: "f \ \[F](g)" and B: "eventually (\x. f x \ 0) F" for f g :: "'a \ 'b" proof - from A obtain c1 c2 where "\\<^sub>F x in F. norm (f x) \ c1 * norm (g x)" "\\<^sub>F x in F. c2 * norm (g x) \ norm (f x)" unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) with B show ?thesis by eventually_elim auto qed with assms show ?thesis by (force simp: bigtheta_sym) qed subsection \Landau symbols and limits\ lemma bigoI_tendsto_norm: fixes f g assumes "((\x. norm (f x / g x)) \ c) F" assumes "eventually (\x. g x \ 0) F" shows "f \ O[F](g)" proof (rule bigoI) from assms have "eventually (\x. dist (norm (f x / g x)) c < 1) F" using tendstoD by force thus "eventually (\x. (norm (f x)) \ (norm c + 1) * (norm (g x))) F" unfolding dist_real_def using assms(2) proof eventually_elim case (elim x) have "(norm (f x)) - norm c * (norm (g x)) \ norm ((norm (f x)) - c * (norm (g x)))" unfolding norm_mult [symmetric] using norm_triangle_ineq2[of "norm (f x)" "c * norm (g x)"] by (simp add: norm_mult abs_mult) also from elim have "\ = norm (norm (g x)) * norm (norm (f x / g x) - c)" unfolding norm_mult [symmetric] by (simp add: algebra_simps norm_divide) also from elim have "norm (norm (f x / g x) - c) \ 1" by simp hence "norm (norm (g x)) * norm (norm (f x / g x) - c) \ norm (norm (g x)) * 1" by (rule mult_left_mono) simp_all finally show ?case by (simp add: algebra_simps) qed qed lemma bigoI_tendsto: assumes "((\x. f x / g x) \ c) F" assumes "eventually (\x. g x \ 0) F" shows "f \ O[F](g)" using assms by (rule bigoI_tendsto_norm[OF tendsto_norm]) lemma bigomegaI_tendsto_norm: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. norm (f x / g x)) \ c) F" shows "f \ \[F](g)" proof (cases "F = bot") case False show ?thesis proof (rule landau_omega.bigI) from lim have "c \ 0" by (rule tendsto_lowerbound) (insert False, simp_all) with c_not_0 have "c > 0" by simp with c_not_0 show "c/2 > 0" by simp from lim have ev: "\\. \ > 0 \ eventually (\x. norm (norm (f x / g x) - c) < \) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) from ev[OF \c/2 > 0\] show "eventually (\x. (norm (f x)) \ c/2 * (norm (g x))) F" proof (eventually_elim) fix x assume B: "norm (norm (f x / g x) - c) < c / 2" from B have g: "g x \ 0" by auto from B have "-c/2 < -norm (norm (f x / g x) - c)" by simp also have "... \ norm (f x / g x) - c" by simp finally show "(norm (f x)) \ c/2 * (norm (g x))" using g by (simp add: field_simps norm_mult norm_divide) qed qed qed simp lemma bigomegaI_tendsto: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. f x / g x) \ c) F" shows "f \ \[F](g)" by (rule bigomegaI_tendsto_norm[OF _ tendsto_norm, of c]) (insert assms, simp_all) lemma smallomegaI_filterlim_at_top_norm: assumes lim: "filterlim (\x. norm (f x / g x)) at_top F" shows "f \ \[F](g)" proof (rule landau_omega.smallI) fix c :: real assume c_pos: "c > 0" from lim have ev: "eventually (\x. norm (f x / g x) \ c) F" by (subst (asm) filterlim_at_top) simp thus "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" proof eventually_elim fix x assume A: "norm (f x / g x) \ c" from A c_pos have "g x \ 0" by auto with A show "(norm (f x)) \ c * (norm (g x))" by (simp add: field_simps norm_divide) qed qed lemma smallomegaI_filterlim_at_infinity: assumes lim: "filterlim (\x. f x / g x) at_infinity F" shows "f \ \[F](g)" proof (rule smallomegaI_filterlim_at_top_norm) from lim show "filterlim (\x. norm (f x / g x)) at_top F" by (rule filterlim_at_infinity_imp_norm_at_top) qed lemma smallomegaD_filterlim_at_top_norm: assumes "f \ \[F](g)" assumes "eventually (\x. g x \ 0) F" shows "LIM x F. norm (f x / g x) :> at_top" proof (subst filterlim_at_top_gt, clarify) fix c :: real assume c: "c > 0" from landau_omega.smallD[OF assms(1) this] assms(2) show "eventually (\x. norm (f x / g x) \ c) F" by eventually_elim (simp add: field_simps norm_divide) qed lemma smallomegaD_filterlim_at_infinity: assumes "f \ \[F](g)" assumes "eventually (\x. g x \ 0) F" shows "LIM x F. f x / g x :> at_infinity" using assms by (intro filterlim_norm_at_top_imp_at_infinity smallomegaD_filterlim_at_top_norm) lemma smallomega_1_conv_filterlim: "f \ \[F](\_. 1) \ filterlim f at_infinity F" by (auto intro: smallomegaI_filterlim_at_infinity dest: smallomegaD_filterlim_at_infinity) lemma smalloI_tendsto: assumes lim: "((\x. f x / g x) \ 0) F" assumes "eventually (\x. g x \ 0) F" shows "f \ o[F](g)" proof (rule landau_o.smallI) fix c :: real assume c_pos: "c > 0" from c_pos and lim have ev: "eventually (\x. norm (f x / g x) < c) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) with assms(2) show "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" by eventually_elim (simp add: field_simps norm_divide) qed lemma smalloD_tendsto: assumes "f \ o[F](g)" shows "((\x. f x / g x) \ 0) F" unfolding tendsto_iff proof clarify fix e :: real assume e: "e > 0" hence "e/2 > 0" by simp from landau_o.smallD[OF assms this] show "eventually (\x. dist (f x / g x) 0 < e) F" proof eventually_elim fix x assume "(norm (f x)) \ e/2 * (norm (g x))" with e have "dist (f x / g x) 0 \ e/2" by (cases "g x = 0") (simp_all add: dist_real_def norm_divide field_simps) also from e have "... < e" by simp finally show "dist (f x / g x) 0 < e" by simp qed qed lemma bigthetaI_tendsto_norm: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. norm (f x / g x)) \ c) F" shows "f \ \[F](g)" proof (rule bigthetaI) from c_not_0 have "\c\ > 0" by simp with lim have "eventually (\x. norm (norm (f x / g x) - c) < \c\) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) hence g: "eventually (\x. g x \ 0) F" by eventually_elim (auto simp add: field_simps) from lim g show "f \ O[F](g)" by (rule bigoI_tendsto_norm) from c_not_0 and lim show "f \ \[F](g)" by (rule bigomegaI_tendsto_norm) qed lemma bigthetaI_tendsto: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. f x / g x) \ c) F" shows "f \ \[F](g)" using assms by (intro bigthetaI_tendsto_norm[OF _ tendsto_norm, of "c"]) simp_all lemma tendsto_add_smallo: assumes "(f1 \ a) F" assumes "f2 \ o[F](f1)" shows "((\x. f1 x + f2 x) \ a) F" proof (subst filterlim_cong[OF refl refl]) from landau_o.smallD[OF assms(2) zero_less_one] have "eventually (\x. norm (f2 x) \ norm (f1 x)) F" by simp thus "eventually (\x. f1 x + f2 x = f1 x * (1 + f2 x / f1 x)) F" by eventually_elim (auto simp: field_simps) next from assms(1) show "((\x. f1 x * (1 + f2 x / f1 x)) \ a) F" by (force intro: tendsto_eq_intros smalloD_tendsto[OF assms(2)]) qed lemma tendsto_diff_smallo: shows "(f1 \ a) F \ f2 \ o[F](f1) \ ((\x. f1 x - f2 x) \ a) F" using tendsto_add_smallo[of f1 a F "\x. -f2 x"] by simp lemma tendsto_add_smallo_iff: assumes "f2 \ o[F](f1)" shows "(f1 \ a) F \ ((\x. f1 x + f2 x) \ a) F" proof assume "((\x. f1 x + f2 x) \ a) F" hence "((\x. f1 x + f2 x - f2 x) \ a) F" by (rule tendsto_diff_smallo) (simp add: landau_o.small.plus_absorb2 assms) thus "(f1 \ a) F" by simp qed (rule tendsto_add_smallo[OF _ assms]) lemma tendsto_diff_smallo_iff: shows "f2 \ o[F](f1) \ (f1 \ a) F \ ((\x. f1 x - f2 x) \ a) F" using tendsto_add_smallo_iff[of "\x. -f2 x" F f1 a] by simp lemma tendsto_divide_smallo: assumes "((\x. f1 x / g1 x) \ a) F" assumes "f2 \ o[F](f1)" "g2 \ o[F](g1)" assumes "eventually (\x. g1 x \ 0) F" shows "((\x. (f1 x + f2 x) / (g1 x + g2 x)) \ a) F" (is "(?f \ _) _") proof (subst tendsto_cong) let ?f' = "\x. (f1 x / g1 x) * (1 + f2 x / f1 x) / (1 + g2 x / g1 x)" have "(?f' \ a * (1 + 0) / (1 + 0)) F" by (rule tendsto_mult tendsto_divide tendsto_add assms tendsto_const smalloD_tendsto[OF assms(2)] smalloD_tendsto[OF assms(3)])+ simp_all thus "(?f' \ a) F" by simp have "(1/2::real) > 0" by simp from landau_o.smallD[OF assms(2) this] landau_o.smallD[OF assms(3) this] have "eventually (\x. norm (f2 x) \ norm (f1 x)/2) F" "eventually (\x. norm (g2 x) \ norm (g1 x)/2) F" by simp_all with assms(4) show "eventually (\x. ?f x = ?f' x) F" proof eventually_elim fix x assume A: "norm (f2 x) \ norm (f1 x)/2" and B: "norm (g2 x) \ norm (g1 x)/2" and C: "g1 x \ 0" show "?f x = ?f' x" proof (cases "f1 x = 0") assume D: "f1 x \ 0" from D have "f1 x + f2 x = f1 x * (1 + f2 x/f1 x)" by (simp add: field_simps) moreover from C have "g1 x + g2 x = g1 x * (1 + g2 x/g1 x)" by (simp add: field_simps) ultimately have "?f x = (f1 x * (1 + f2 x/f1 x)) / (g1 x * (1 + g2 x/g1 x))" by (simp only:) also have "... = ?f' x" by simp finally show ?thesis . qed (insert A, simp) qed qed lemma bigo_powr: fixes f :: "'a \ real" assumes "f \ O[F](g)" "p \ 0" shows "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" proof- from assms(1) obtain c where c: "c > 0" and *: "\\<^sub>F x in F. norm (f x) \ c * norm (g x)" by (elim landau_o.bigE landau_omega.bigE IntE) from assms(2) * have "eventually (\x. (norm (f x)) powr p \ (c * norm (g x)) powr p) F" by (auto elim!: eventually_mono intro!: powr_mono2) with c show "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" by (intro bigoI[of _ "c powr p"]) (simp_all add: powr_mult) qed lemma smallo_powr: fixes f :: "'a \ real" assumes "f \ o[F](g)" "p > 0" shows "(\x. \f x\ powr p) \ o[F](\x. \g x\ powr p)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "c powr (1/p) > 0" by simp from landau_o.smallD[OF assms(1) this] show "eventually (\x. norm (\f x\ powr p) \ c * norm (\g x\ powr p)) F" proof eventually_elim fix x assume "(norm (f x)) \ c powr (1 / p) * (norm (g x))" with assms(2) have "(norm (f x)) powr p \ (c powr (1 / p) * (norm (g x))) powr p" by (intro powr_mono2) simp_all also from assms(2) c have "... = c * (norm (g x)) powr p" by (simp add: field_simps powr_mult powr_powr) finally show "norm (\f x\ powr p) \ c * norm (\g x\ powr p)" by simp qed qed lemma smallo_powr_nonneg: fixes f :: "'a \ real" assumes "f \ o[F](g)" "p > 0" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr p) \ o[F](\x. g x powr p)" proof - from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also have "(\x. \f x\ powr p) \ o[F](\x. \g x\ powr p)" by (intro smallo_powr) fact+ also from assms(4) have "(\x. \g x\ powr p) \ \[F](\x. g x powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show ?thesis . qed lemma bigtheta_powr: fixes f :: "'a \ real" shows "f \ \[F](g) \ (\x. \f x\ powr p) \ \[F](\x. \g x\ powr p)" apply (cases "p < 0") apply (subst bigtheta_inverse[symmetric], subst (1 2) powr_minus[symmetric]) unfolding bigtheta_def apply (auto simp: bigomega_iff_bigo intro!: bigo_powr) done lemma bigo_powr_nonneg: fixes f :: "'a \ real" assumes "f \ O[F](g)" "p \ 0" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr p) \ O[F](\x. g x powr p)" proof - from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also have "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" by (intro bigo_powr) fact+ also from assms(4) have "(\x. \g x\ powr p) \ \[F](\x. g x powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show ?thesis . qed lemma zero_in_smallo [simp]: "(\_. 0) \ o[F](f)" by (intro landau_o.smallI) simp_all lemma zero_in_bigo [simp]: "(\_. 0) \ O[F](f)" by (intro landau_o.bigI[of 1]) simp_all lemma in_bigomega_zero [simp]: "f \ \[F](\x. 0)" by (rule landau_omega.bigI[of 1]) simp_all lemma in_smallomega_zero [simp]: "f \ \[F](\x. 0)" by (simp add: smallomega_iff_smallo) lemma in_smallo_zero_iff [simp]: "f \ o[F](\_. 0) \ eventually (\x. f x = 0) F" proof assume "f \ o[F](\_. 0)" from landau_o.smallD[OF this, of 1] show "eventually (\x. f x = 0) F" by simp next assume "eventually (\x. f x = 0) F" hence "\c>0. eventually (\x. (norm (f x)) \ c * \0\) F" by simp thus "f \ o[F](\_. 0)" unfolding smallo_def by simp qed lemma in_bigo_zero_iff [simp]: "f \ O[F](\_. 0) \ eventually (\x. f x = 0) F" proof assume "f \ O[F](\_. 0)" thus "eventually (\x. f x = 0) F" by (elim landau_o.bigE) simp next assume "eventually (\x. f x = 0) F" hence "eventually (\x. (norm (f x)) \ 1 * \0\) F" by simp thus "f \ O[F](\_. 0)" by (intro landau_o.bigI[of 1]) simp_all qed lemma zero_in_smallomega_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" by (simp add: smallomega_iff_smallo) lemma zero_in_bigomega_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" by (simp add: bigomega_iff_bigo) lemma zero_in_bigtheta_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" unfolding bigtheta_def by simp lemma in_bigtheta_zero_iff [simp]: "f \ \[F](\x. 0) \ eventually (\x. f x = 0) F" unfolding bigtheta_def by simp lemma cmult_in_bigo_iff [simp]: "(\x. c * f x) \ O[F](g) \ c = 0 \ f \ O[F](g)" and cmult_in_bigo_iff' [simp]: "(\x. f x * c) \ O[F](g) \ c = 0 \ f \ O[F](g)" and cmult_in_smallo_iff [simp]: "(\x. c * f x) \ o[F](g) \ c = 0 \ f \ o[F](g)" and cmult_in_smallo_iff' [simp]: "(\x. f x * c) \ o[F](g) \ c = 0 \ f \ o[F](g)" by (cases "c = 0", simp, simp)+ lemma bigo_const [simp]: "(\_. c) \ O[F](\_. 1)" by (rule bigoI[of _ "norm c"]) simp lemma bigo_const_iff [simp]: "(\_. c1) \ O[F](\_. c2) \ F = bot \ c1 = 0 \ c2 \ 0" by (cases "c1 = 0"; cases "c2 = 0") (auto simp: bigo_def eventually_False intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"]) lemma bigomega_const_iff [simp]: "(\_. c1) \ \[F](\_. c2) \ F = bot \ c1 \ 0 \ c2 = 0" by (cases "c1 = 0"; cases "c2 = 0") (auto simp: bigomega_def eventually_False mult_le_0_iff intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"]) lemma smallo_real_nat_transfer: "(f :: real \ real) \ o(g) \ (\x::nat. f (real x)) \ o(\x. g (real x))" by (rule landau_o.small.compose[OF _ filterlim_real_sequentially]) lemma bigo_real_nat_transfer: "(f :: real \ real) \ O(g) \ (\x::nat. f (real x)) \ O(\x. g (real x))" by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) lemma smallomega_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" by (rule landau_omega.small.compose[OF _ filterlim_real_sequentially]) lemma bigomega_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" by (rule landau_omega.big.compose[OF _ filterlim_real_sequentially]) lemma bigtheta_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" unfolding bigtheta_def using bigo_real_nat_transfer bigomega_real_nat_transfer by blast lemmas landau_real_nat_transfer [intro] = bigo_real_nat_transfer smallo_real_nat_transfer bigomega_real_nat_transfer smallomega_real_nat_transfer bigtheta_real_nat_transfer lemma landau_symbol_if_at_top_eq [simp]: assumes "landau_symbol L L' Lr" shows "L at_top (\x::'a::linordered_semidom. if x = a then f x else g x) = L at_top (g)" apply (rule landau_symbol.cong[OF assms]) using less_add_one[of a] apply (auto intro: eventually_mono eventually_ge_at_top[of "a + 1"]) done lemmas landau_symbols_if_at_top_eq [simp] = landau_symbols[THEN landau_symbol_if_at_top_eq] lemma sum_in_smallo: assumes "f \ o[F](h)" "g \ o[F](h)" shows "(\x. f x + g x) \ o[F](h)" "(\x. f x - g x) \ o[F](h)" proof - have "(\x. f x + g x) \ o[F](h)" if "f \ o[F](h)" "g \ o[F](h)" for f g proof (rule landau_o.smallI) fix c :: real assume "c > 0" hence "c/2 > 0" by simp from that[THEN landau_o.smallD[OF _ this]] show "eventually (\x. norm (f x + g x) \ c * (norm (h x))) F" by eventually_elim (auto intro: order.trans[OF norm_triangle_ineq]) qed from this[of f g] this[of f "\x. -g x"] assms show "(\x. f x + g x) \ o[F](h)" "(\x. f x - g x) \ o[F](h)" by simp_all qed lemma big_sum_in_smallo: assumes "\x. x \ A \ f x \ o[F](g)" shows "(\x. sum (\y. f y x) A) \ o[F](g)" using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_smallo) lemma sum_in_bigo: assumes "f \ O[F](h)" "g \ O[F](h)" shows "(\x. f x + g x) \ O[F](h)" "(\x. f x - g x) \ O[F](h)" proof - have "(\x. f x + g x) \ O[F](h)" if "f \ O[F](h)" "g \ O[F](h)" for f g proof - from that obtain c1 c2 where *: "c1 > 0" "c2 > 0" and **: "\\<^sub>F x in F. norm (f x) \ c1 * norm (h x)" "\\<^sub>F x in F. norm (g x) \ c2 * norm (h x)" by (elim landau_o.bigE) from ** have "eventually (\x. norm (f x + g x) \ (c1 + c2) * (norm (h x))) F" by eventually_elim (auto simp: algebra_simps intro: order.trans[OF norm_triangle_ineq]) then show ?thesis by (rule bigoI) qed from assms this[of f g] this[of f "\x. - g x"] show "(\x. f x + g x) \ O[F](h)" "(\x. f x - g x) \ O[F](h)" by simp_all qed lemma big_sum_in_bigo: assumes "\x. x \ A \ f x \ O[F](g)" shows "(\x. sum (\y. f y x) A) \ O[F](g)" using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_bigo) context landau_symbol begin lemma mult_cancel_left: assumes "f1 \ \[F](g1)" and "eventually (\x. g1 x \ 0) F" notes [trans] = bigtheta_trans1 bigtheta_trans2 shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x) \ f2 \ L F (g2)" proof assume A: "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" from assms have nz: "eventually (\x. f1 x \ 0) F" by (simp add: eventually_nonzero_bigtheta) hence "f2 \ \[F](\x. f1 x * f2 x / f1 x)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also from A assms nz have "(\x. f1 x * f2 x / f1 x) \ L F (\x. g1 x * g2 x / f1 x)" by (intro divide_right) simp_all also from assms nz have "(\x. g1 x * g2 x / f1 x) \ \[F](\x. g1 x * g2 x / g1 x)" by (intro landau_theta.mult landau_theta.divide) (simp_all add: bigtheta_sym) also from assms have "(\x. g1 x * g2 x / g1 x) \ \[F](g2)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show "f2 \ L F (g2)" . next assume "f2 \ L F (g2)" hence "(\x. f1 x * f2 x) \ L F (\x. f1 x * g2 x)" by (rule mult_left) also have "(\x. f1 x * g2 x) \ \[F](\x. g1 x * g2 x)" by (intro landau_theta.mult_right assms) finally show "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" . qed lemma mult_cancel_right: assumes "f2 \ \[F](g2)" and "eventually (\x. g2 x \ 0) F" shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x) \ f1 \ L F (g1)" by (subst (1 2) mult.commute) (rule mult_cancel_left[OF assms]) lemma divide_cancel_right: assumes "f2 \ \[F](g2)" and "eventually (\x. g2 x \ 0) F" shows "(\x. f1 x / f2 x) \ L F (\x. g1 x / g2 x) \ f1 \ L F (g1)" by (subst (1 2) divide_inverse, intro mult_cancel_right bigtheta_inverse) (simp_all add: assms) lemma divide_cancel_left: assumes "f1 \ \[F](g1)" and "eventually (\x. g1 x \ 0) F" shows "(\x. f1 x / f2 x) \ L F (\x. g1 x / g2 x) \ (\x. inverse (f2 x)) \ L F (\x. inverse (g2 x))" by (simp only: divide_inverse mult_cancel_left[OF assms]) end lemma powr_smallo_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ o[F](\x. g x powr q) \ p < q" proof- from assms have "eventually (\x. g x \ 1) F" by (force simp: filterlim_at_top) hence A: "eventually (\x. g x \ 0) F" by eventually_elim simp have B: "(\x. g x powr q) \ O[F](\x. g x powr p) \ (\x. g x powr p) \ o[F](\x. g x powr q)" proof assume "(\x. g x powr q) \ O[F](\x. g x powr p)" "(\x. g x powr p) \ o[F](\x. g x powr q)" from landau_o.big_small_asymmetric[OF this] have "eventually (\x. g x = 0) F" by simp with A have "eventually (\_::'a. False) F" by eventually_elim simp thus False by (simp add: eventually_False assms) qed show ?thesis proof (cases p q rule: linorder_cases) assume "p < q" hence "(\x. g x powr p) \ o[F](\x. g x powr q)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with \p < q\ show ?thesis by auto next assume "p = q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" by (auto intro!: bigthetaD1) with B \p = q\ show ?thesis by auto next assume "p > q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr landau_o.small_imp_big simp flip: powr_diff) with B \p > q\ show ?thesis by auto qed qed lemma powr_bigo_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ O[F](\x. g x powr q) \ p \ q" proof- from assms have "eventually (\x. g x \ 1) F" by (force simp: filterlim_at_top) hence A: "eventually (\x. g x \ 0) F" by eventually_elim simp have B: "(\x. g x powr q) \ o[F](\x. g x powr p) \ (\x. g x powr p) \ O[F](\x. g x powr q)" proof assume "(\x. g x powr q) \ o[F](\x. g x powr p)" "(\x. g x powr p) \ O[F](\x. g x powr q)" from landau_o.small_big_asymmetric[OF this] have "eventually (\x. g x = 0) F" by simp with A have "eventually (\_::'a. False) F" by eventually_elim simp thus False by (simp add: eventually_False assms) qed show ?thesis proof (cases p q rule: linorder_cases) assume "p < q" hence "(\x. g x powr p) \ o[F](\x. g x powr q)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with \p < q\ show ?thesis by (auto intro: landau_o.small_imp_big) next assume "p = q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" by (auto intro!: bigthetaD1) with B \p = q\ show ?thesis by auto next assume "p > q" hence "(\x. g x powr q) \ o[F](\x. g x powr p)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with B \p > q\ show ?thesis by (auto intro: landau_o.small_imp_big) qed qed lemma powr_bigtheta_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ \[F](\x. g x powr q) \ p = q" using assms unfolding bigtheta_def by (auto simp: bigomega_iff_bigo powr_bigo_iff) subsection \Flatness of real functions\ text \ Given two real-valued functions $f$ and $g$, we say that $f$ is flatter than $g$ if any power of $f(x)$ is asymptotically dominated by any positive power of $g(x)$. This is a useful notion since, given two products of powers of functions sorted by flatness, we can compare them asymptotically by simply comparing the exponent lists lexicographically. A simple sufficient criterion for flatness it that $\ln f(x) \in o(\ln g(x))$, which we show now. \ lemma ln_smallo_imp_flat: fixes f g :: "real \ real" assumes lim_f: "filterlim f at_top at_top" assumes lim_g: "filterlim g at_top at_top" assumes ln_o_ln: "(\x. ln (f x)) \ o(\x. ln (g x))" assumes q: "q > 0" shows "(\x. f x powr p) \ o(\x. g x powr q)" proof (rule smalloI_tendsto) from lim_f have "eventually (\x. f x > 0) at_top" by (simp add: filterlim_at_top_dense) hence f_nz: "eventually (\x. f x \ 0) at_top" by eventually_elim simp from lim_g have g_gt_1: "eventually (\x. g x > 1) at_top" by (simp add: filterlim_at_top_dense) hence g_nz: "eventually (\x. g x \ 0) at_top" by eventually_elim simp thus "eventually (\x. g x powr q \ 0) at_top" by eventually_elim simp have eq: "eventually (\x. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x) = p * ln (f x) - q * ln (g x)) at_top" using g_gt_1 by eventually_elim (insert q, simp_all add: field_simps) have "filterlim (\x. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x)) at_bot at_top" by (insert q) (rule filterlim_tendsto_neg_mult_at_bot tendsto_mult tendsto_const tendsto_diff smalloD_tendsto[OF ln_o_ln] lim_g filterlim_compose[OF ln_at_top] | simp)+ hence "filterlim (\x. p * ln (f x) - q * ln (g x)) at_bot at_top" by (subst (asm) filterlim_cong[OF refl refl eq]) hence *: "((\x. exp (p * ln (f x) - q * ln (g x))) \ 0) at_top" by (rule filterlim_compose[OF exp_at_bot]) have eq: "eventually (\x. exp (p * ln (f x) - q * ln (g x)) = f x powr p / g x powr q) at_top" using f_nz g_nz by eventually_elim (simp add: powr_def exp_diff) show "((\x. f x powr p / g x powr q) \ 0) at_top" using * by (subst (asm) filterlim_cong[OF refl refl eq]) qed lemma ln_smallo_imp_flat': fixes f g :: "real \ real" assumes lim_f: "filterlim f at_top at_top" assumes lim_g: "filterlim g at_top at_top" assumes ln_o_ln: "(\x. ln (f x)) \ o(\x. ln (g x))" assumes q: "q < 0" shows "(\x. g x powr q) \ o(\x. f x powr p)" proof - from lim_f lim_g have "eventually (\x. f x > 0) at_top" "eventually (\x. g x > 0) at_top" by (simp_all add: filterlim_at_top_dense) hence "eventually (\x. f x \ 0) at_top" "eventually (\x. g x \ 0) at_top" by (auto elim: eventually_mono) moreover from assms have "(\x. f x powr -p) \ o(\x. g x powr -q)" by (intro ln_smallo_imp_flat assms) simp_all ultimately show ?thesis unfolding powr_minus by (simp add: landau_o.small.inverse_cancel) qed subsection \Asymptotic Equivalence\ named_theorems asymp_equiv_intros named_theorems asymp_equiv_simps definition asymp_equiv :: "('a \ ('b :: real_normed_field)) \ 'a filter \ ('a \ 'b) \ bool" (\_ \[_] _\ [51, 10, 51] 50) where "f \[F] g \ ((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" abbreviation (input) asymp_equiv_at_top where "asymp_equiv_at_top f g \ f \[at_top] g" bundle asymp_equiv_notation begin notation asymp_equiv_at_top (infix \\\ 50) end lemma asymp_equivI: "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F \ f \[F] g" by (simp add: asymp_equiv_def) lemma asymp_equivD: "f \[F] g \ ((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (simp add: asymp_equiv_def) lemma asymp_equiv_filtermap_iff: "f \[filtermap h F] g \ (\x. f (h x)) \[F] (\x. g (h x))" by (simp add: asymp_equiv_def filterlim_filtermap) lemma asymp_equiv_refl [simp, asymp_equiv_intros]: "f \[F] f" proof (intro asymp_equivI) have "eventually (\x. 1 = (if f x = 0 \ f x = 0 then 1 else f x / f x)) F" by (intro always_eventually) simp moreover have "((\_. 1) \ 1) F" by simp ultimately show "((\x. if f x = 0 \ f x = 0 then 1 else f x / f x) \ 1) F" by (simp add: tendsto_eventually) qed lemma asymp_equiv_symI: assumes "f \[F] g" shows "g \[F] f" using tendsto_inverse[OF asymp_equivD[OF assms]] by (auto intro!: asymp_equivI simp: if_distrib conj_commute cong: if_cong) lemma asymp_equiv_sym: "f \[F] g \ g \[F] f" by (blast intro: asymp_equiv_symI) lemma asymp_equivI': assumes "((\x. f x / g x) \ 1) F" shows "f \[F] g" proof (cases "F = bot") case False have "eventually (\x. f x \ 0) F" proof (rule ccontr) assume "\eventually (\x. f x \ 0) F" hence "frequently (\x. f x = 0) F" by (simp add: frequently_def) hence "frequently (\x. f x / g x = 0) F" by (auto elim!: frequently_elim1) from limit_frequently_eq[OF False this assms] show False by simp_all qed hence "eventually (\x. f x / g x = (if f x = 0 \ g x = 0 then 1 else f x / g x)) F" by eventually_elim simp with assms show "f \[F] g" unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed (simp_all add: asymp_equiv_def) lemma asymp_equiv_cong: assumes "eventually (\x. f1 x = f2 x) F" "eventually (\x. g1 x = g2 x) F" shows "f1 \[F] g1 \ f2 \[F] g2" unfolding asymp_equiv_def proof (rule tendsto_cong, goal_cases) case 1 from assms show ?case by eventually_elim simp qed lemma asymp_equiv_eventually_zeros: fixes f g :: "'a \ 'b :: real_normed_field" assumes "f \[F] g" shows "eventually (\x. f x = 0 \ g x = 0) F" proof - let ?h = "\x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "eventually (\x. x \ 0) (nhds (1::'b))" by (rule t1_space_nhds) auto hence "eventually (\x. x \ 0) (filtermap ?h F)" using assms unfolding asymp_equiv_def filterlim_def by (rule filter_leD [rotated]) hence "eventually (\x. ?h x \ 0) F" by (simp add: eventually_filtermap) thus ?thesis by eventually_elim (auto split: if_splits) qed lemma asymp_equiv_transfer: assumes "f1 \[F] g1" "eventually (\x. f1 x = f2 x) F" "eventually (\x. g1 x = g2 x) F" shows "f2 \[F] g2" using assms(1) asymp_equiv_cong[OF assms(2,3)] by simp lemma asymp_equiv_transfer_trans [trans]: assumes "(\x. f x (h1 x)) \[F] (\x. g x (h1 x))" assumes "eventually (\x. h1 x = h2 x) F" shows "(\x. f x (h2 x)) \[F] (\x. g x (h2 x))" by (rule asymp_equiv_transfer[OF assms(1)]) (insert assms(2), auto elim!: eventually_mono) lemma asymp_equiv_trans [trans]: fixes f g h assumes "f \[F] g" "g \[F] h" shows "f \[F] h" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" from tendsto_mult[OF assms[THEN asymp_equivD]] have "((\x. ?T f g x * ?T g h x) \ 1) F" by simp moreover from assms[THEN asymp_equiv_eventually_zeros] have "eventually (\x. ?T f g x * ?T g h x = ?T f h x) F" by eventually_elim simp ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed lemma asymp_equiv_trans_lift1 [trans]: assumes "a \[F] f b" "b \[F] c" "\c d. c \[F] d \ f c \[F] f d" shows "a \[F] f c" using assms by (blast intro: asymp_equiv_trans) lemma asymp_equiv_trans_lift2 [trans]: assumes "f a \[F] b" "a \[F] c" "\c d. c \[F] d \ f c \[F] f d" shows "f c \[F] b" using asymp_equiv_symI[OF assms(3)[OF assms(2)]] assms(1) by (blast intro: asymp_equiv_trans) lemma asymp_equivD_const: assumes "f \[F] (\_. c)" shows "(f \ c) F" proof (cases "c = 0") case False with tendsto_mult_right[OF asymp_equivD[OF assms], of c] show ?thesis by simp next case True with asymp_equiv_eventually_zeros[OF assms] show ?thesis by (simp add: tendsto_eventually) qed lemma asymp_equiv_refl_ev: assumes "eventually (\x. f x = g x) F" shows "f \[F] g" by (intro asymp_equivI tendsto_eventually) (insert assms, auto elim!: eventually_mono) lemma asymp_equiv_sandwich: fixes f g h :: "'a \ 'b :: {real_normed_field, order_topology, linordered_field}" assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. f x \ g x) F" assumes "eventually (\x. g x \ h x) F" assumes "f \[F] h" shows "g \[F] f" "g \[F] h" proof - show "g \[F] f" proof (rule asymp_equivI, rule tendsto_sandwich) from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)] show "eventually (\n. (if h n = 0 \ f n = 0 then 1 else h n / f n) \ (if g n = 0 \ f n = 0 then 1 else g n / f n)) F" by eventually_elim (auto intro!: divide_right_mono) from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)] show "eventually (\n. 1 \ (if g n = 0 \ f n = 0 then 1 else g n / f n)) F" by eventually_elim (auto intro!: divide_right_mono) qed (insert asymp_equiv_symI[OF assms(4)], simp_all add: asymp_equiv_def) also note \f \[F] h\ finally show "g \[F] h" . qed lemma asymp_equiv_imp_eventually_same_sign: fixes f g :: "real \ real" assumes "f \[F] g" shows "eventually (\x. sgn (f x) = sgn (g x)) F" proof - from assms have "((\x. sgn (if f x = 0 \ g x = 0 then 1 else f x / g x)) \ sgn 1) F" unfolding asymp_equiv_def by (rule tendsto_sgn) simp_all from order_tendstoD(1)[OF this, of "1/2"] have "eventually (\x. sgn (if f x = 0 \ g x = 0 then 1 else f x / g x) > 1/2) F" by simp thus "eventually (\x. sgn (f x) = sgn (g x)) F" proof eventually_elim case (elim x) thus ?case by (cases "f x" "0 :: real" rule: linorder_cases; cases "g x" "0 :: real" rule: linorder_cases) simp_all qed qed lemma fixes f g :: "_ \ real" assumes "f \[F] g" shows asymp_equiv_eventually_same_sign: "eventually (\x. sgn (f x) = sgn (g x)) F" (is ?th1) and asymp_equiv_eventually_neg_iff: "eventually (\x. f x < 0 \ g x < 0) F" (is ?th2) and asymp_equiv_eventually_pos_iff: "eventually (\x. f x > 0 \ g x > 0) F" (is ?th3) proof - from assms have "filterlim (\x. if f x = 0 \ g x = 0 then 1 else f x / g x) (nhds 1) F" by (rule asymp_equivD) from order_tendstoD(1)[OF this zero_less_one] show ?th1 ?th2 ?th3 by (eventually_elim; force simp: sgn_if field_split_simps split: if_splits)+ qed lemma asymp_equiv_tendsto_transfer: assumes "f \[F] g" and "(f \ c) F" shows "(g \ c) F" proof - let ?h = "\x. (if g x = 0 \ f x = 0 then 1 else g x / f x) * f x" from assms(1) have "g \[F] f" by (rule asymp_equiv_symI) hence "filterlim (\x. if g x = 0 \ f x = 0 then 1 else g x / f x) (nhds 1) F" by (rule asymp_equivD) from tendsto_mult[OF this assms(2)] have "(?h \ c) F" by simp moreover have "eventually (\x. ?h x = g x) F" using asymp_equiv_eventually_zeros[OF assms(1)] by eventually_elim simp ultimately show ?thesis by (rule Lim_transform_eventually) qed lemma tendsto_asymp_equiv_cong: assumes "f \[F] g" shows "(f \ c) F \ (g \ c) F" proof - have "(f \ c * 1) F" if fg: "f \[F] g" and "(g \ c) F" for f g :: "'a \ 'b" proof - from that have *: "((\x. g x * (if f x = 0 \ g x = 0 then 1 else f x / g x)) \ c * 1) F" by (intro tendsto_intros asymp_equivD) have "eventually (\x. g x * (if f x = 0 \ g x = 0 then 1 else f x / g x) = f x) F" using asymp_equiv_eventually_zeros[OF fg] by eventually_elim simp with * show ?thesis by (rule Lim_transform_eventually) qed from this[of f g] this[of g f] assms show ?thesis by (auto simp: asymp_equiv_sym) qed lemma smallo_imp_eventually_sgn: fixes f g :: "real \ real" assumes "g \ o(f)" shows "eventually (\x. sgn (f x + g x) = sgn (f x)) at_top" proof - have "0 < (1/2 :: real)" by simp from landau_o.smallD[OF assms, OF this] have "eventually (\x. \g x\ \ 1/2 * \f x\) at_top" by simp thus ?thesis proof eventually_elim case (elim x) thus ?case by (cases "f x" "0::real" rule: linorder_cases; cases "f x + g x" "0::real" rule: linorder_cases) simp_all qed qed context begin private lemma asymp_equiv_add_rightI: assumes "f \[F] g" "h \ o[F](g)" shows "(\x. f x + h x) \[F] g" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" from landau_o.smallD[OF assms(2) zero_less_one] have ev: "eventually (\x. g x = 0 \ h x = 0) F" by eventually_elim auto have "(\x. f x + h x) \[F] g \ ((\x. ?T f g x + h x / g x) \ 1) F" unfolding asymp_equiv_def using ev by (intro tendsto_cong) (auto elim!: eventually_mono simp: field_split_simps) also have "\ \ ((\x. ?T f g x + h x / g x) \ 1 + 0) F" by simp also have \ by (intro tendsto_intros asymp_equivD assms smalloD_tendsto) finally show "(\x. f x + h x) \[F] g" . qed lemma asymp_equiv_add_right [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "(\x. f x + h x) \[F] g \ f \[F] g" proof assume "(\x. f x + h x) \[F] g" from asymp_equiv_add_rightI[OF this, of "\x. -h x"] assms show "f \[F] g" by simp qed (simp_all add: asymp_equiv_add_rightI assms) end lemma asymp_equiv_add_left [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "(\x. h x + f x) \[F] g \ f \[F] g" using asymp_equiv_add_right[OF assms] by (simp add: add.commute) lemma asymp_equiv_add_right' [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "g \[F] (\x. f x + h x) \ g \[F] f" using asymp_equiv_add_right[OF assms] by (simp add: asymp_equiv_sym) lemma asymp_equiv_add_left' [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "g \[F] (\x. h x + f x) \ g \[F] f" using asymp_equiv_add_left[OF assms] by (simp add: asymp_equiv_sym) lemma smallo_imp_asymp_equiv: assumes "(\x. f x - g x) \ o[F](g)" shows "f \[F] g" proof - from assms have "(\x. f x - g x + g x) \[F] g" by (subst asymp_equiv_add_left) simp_all thus ?thesis by simp qed lemma asymp_equiv_uminus [asymp_equiv_intros]: "f \[F] g \ (\x. -f x) \[F] (\x. -g x)" by (simp add: asymp_equiv_def cong: if_cong) lemma asymp_equiv_uminus_iff [asymp_equiv_simps]: "(\x. -f x) \[F] g \ f \[F] (\x. -g x)" by (simp add: asymp_equiv_def cong: if_cong) lemma asymp_equiv_mult [asymp_equiv_intros]: fixes f1 f2 g1 g2 :: "'a \ 'b :: real_normed_field" assumes "f1 \[F] g1" "f2 \[F] g2" shows "(\x. f1 x * f2 x) \[F] (\x. g1 x * g2 x)" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" let ?S = "\x. (if f1 x = 0 \ g1 x = 0 then 1 - ?T f2 g2 x else if f2 x = 0 \ g2 x = 0 then 1 - ?T f1 g1 x else 0)" let ?S' = "\x. ?T (\x. f1 x * f2 x) (\x. g1 x * g2 x) x - ?T f1 g1 x * ?T f2 g2 x" have A: "((\x. 1 - ?T f g x) \ 0) F" if "f \[F] g" for f g :: "'a \ 'b" by (rule tendsto_eq_intros refl asymp_equivD[OF that])+ simp_all from assms have *: "((\x. ?T f1 g1 x * ?T f2 g2 x) \ 1 * 1) F" by (intro tendsto_mult asymp_equivD) { have "(?S \ 0) F" by (intro filterlim_If assms[THEN A, THEN tendsto_mono[rotated]]) (auto intro: le_infI1 le_infI2) moreover have "eventually (\x. ?S x = ?S' x) F" using assms[THEN asymp_equiv_eventually_zeros] by eventually_elim auto ultimately have "(?S' \ 0) F" by (rule Lim_transform_eventually) } with * have "(?T (\x. f1 x * f2 x) (\x. g1 x * g2 x) \ 1 * 1) F" by (rule Lim_transform) then show ?thesis by (simp add: asymp_equiv_def) qed lemma asymp_equiv_power [asymp_equiv_intros]: "f \[F] g \ (\x. f x ^ n) \[F] (\x. g x ^ n)" by (induction n) (simp_all add: asymp_equiv_mult) lemma asymp_equiv_inverse [asymp_equiv_intros]: assumes "f \[F] g" shows "(\x. inverse (f x)) \[F] (\x. inverse (g x))" proof - from tendsto_inverse[OF asymp_equivD[OF assms]] have "((\x. if f x = 0 \ g x = 0 then 1 else g x / f x) \ 1) F" by (simp add: if_distrib cong: if_cong) also have "(\x. if f x = 0 \ g x = 0 then 1 else g x / f x) = (\x. if f x = 0 \ g x = 0 then 1 else inverse (f x) / inverse (g x))" by (intro ext) (simp add: field_simps) finally show ?thesis by (simp add: asymp_equiv_def) qed lemma asymp_equiv_inverse_iff [asymp_equiv_simps]: "(\x. inverse (f x)) \[F] (\x. inverse (g x)) \ f \[F] g" proof assume "(\x. inverse (f x)) \[F] (\x. inverse (g x))" hence "(\x. inverse (inverse (f x))) \[F] (\x. inverse (inverse (g x)))" (is ?P) by (rule asymp_equiv_inverse) also have "?P \ f \[F] g" by (intro asymp_equiv_cong) simp_all finally show "f \[F] g" . qed (simp_all add: asymp_equiv_inverse) lemma asymp_equiv_divide [asymp_equiv_intros]: assumes "f1 \[F] g1" "f2 \[F] g2" shows "(\x. f1 x / f2 x) \[F] (\x. g1 x / g2 x)" using asymp_equiv_mult[OF assms(1) asymp_equiv_inverse[OF assms(2)]] by (simp add: field_simps) lemma asymp_equiv_compose [asymp_equiv_intros]: assumes "f \[G] g" "filterlim h G F" shows "f \ h \[F] g \ h" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "f \ h \[F] g \ h \ ((?T f g \ h) \ 1) F" by (simp add: asymp_equiv_def o_def) also have "\ \ (?T f g \ 1) (filtermap h F)" by (rule tendsto_compose_filtermap) also have "\" by (rule tendsto_mono[of _ G]) (insert assms, simp_all add: asymp_equiv_def filterlim_def) finally show ?thesis . qed lemma asymp_equiv_compose': assumes "f \[G] g" "filterlim h G F" shows "(\x. f (h x)) \[F] (\x. g (h x))" using asymp_equiv_compose[OF assms] by (simp add: o_def) lemma asymp_equiv_powr_real [asymp_equiv_intros]: fixes f g :: "'a \ real" assumes "f \[F] g" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr y) \[F] (\x. g x powr y)" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "((\x. ?T f g x powr y) \ 1 powr y) F" by (intro tendsto_intros asymp_equivD[OF assms(1)]) simp_all hence "((\x. ?T f g x powr y) \ 1) F" by simp moreover have "eventually (\x. ?T f g x powr y = ?T (\x. f x powr y) (\x. g x powr y) x) F" using asymp_equiv_eventually_zeros[OF assms(1)] assms(2,3) by eventually_elim (auto simp: powr_divide) ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed lemma asymp_equiv_norm [asymp_equiv_intros]: fixes f g :: "'a \ 'b :: real_normed_field" assumes "f \[F] g" shows "(\x. norm (f x)) \[F] (\x. norm (g x))" using tendsto_norm[OF asymp_equivD[OF assms]] by (simp add: if_distrib asymp_equiv_def norm_divide cong: if_cong) lemma asymp_equiv_abs_real [asymp_equiv_intros]: fixes f g :: "'a \ real" assumes "f \[F] g" shows "(\x. \f x\) \[F] (\x. \g x\)" using tendsto_rabs[OF asymp_equivD[OF assms]] by (simp add: if_distrib asymp_equiv_def cong: if_cong) lemma asymp_equiv_imp_eventually_le: assumes "f \[F] g" "c > 1" shows "eventually (\x. norm (f x) \ c * norm (g x)) F" proof - from order_tendstoD(2)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)] asymp_equiv_eventually_zeros[OF assms(1)] show ?thesis by eventually_elim (auto split: if_splits simp: field_simps) qed lemma asymp_equiv_imp_eventually_ge: assumes "f \[F] g" "c < 1" shows "eventually (\x. norm (f x) \ c * norm (g x)) F" proof - from order_tendstoD(1)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)] asymp_equiv_eventually_zeros[OF assms(1)] show ?thesis by eventually_elim (auto split: if_splits simp: field_simps) qed lemma asymp_equiv_imp_bigo: assumes "f \[F] g" shows "f \ O[F](g)" proof (rule bigoI) have "(3/2::real) > 1" by simp from asymp_equiv_imp_eventually_le[OF assms this] show "eventually (\x. norm (f x) \ 3/2 * norm (g x)) F" by eventually_elim simp qed lemma asymp_equiv_imp_bigomega: "f \[F] g \ f \ \[F](g)" using asymp_equiv_imp_bigo[of g F f] by (simp add: asymp_equiv_sym bigomega_iff_bigo) lemma asymp_equiv_imp_bigtheta: "f \[F] g \ f \ \[F](g)" by (intro bigthetaI asymp_equiv_imp_bigo asymp_equiv_imp_bigomega) lemma asymp_equiv_at_infinity_transfer: assumes "f \[F] g" "filterlim f at_infinity F" shows "filterlim g at_infinity F" proof - from assms(1) have "g \ \[F](f)" by (rule asymp_equiv_imp_bigtheta[OF asymp_equiv_symI]) also from assms have "f \ \[F](\_. 1)" by (simp add: smallomega_1_conv_filterlim) finally show ?thesis by (simp add: smallomega_1_conv_filterlim) qed lemma asymp_equiv_at_top_transfer: fixes f g :: "_ \ real" assumes "f \[F] g" "filterlim f at_top F" shows "filterlim g at_top F" proof (rule filterlim_at_infinity_imp_filterlim_at_top) show "filterlim g at_infinity F" by (rule asymp_equiv_at_infinity_transfer[OF assms(1) filterlim_mono[OF assms(2)]]) (auto simp: at_top_le_at_infinity) from assms(2) have "eventually (\x. f x > 0) F" using filterlim_at_top_dense by blast with asymp_equiv_eventually_pos_iff[OF assms(1)] show "eventually (\x. g x > 0) F" by eventually_elim blast qed lemma asymp_equiv_at_bot_transfer: fixes f g :: "_ \ real" assumes "f \[F] g" "filterlim f at_bot F" shows "filterlim g at_bot F" unfolding filterlim_uminus_at_bot by (rule asymp_equiv_at_top_transfer[of "\x. -f x" F "\x. -g x"]) (insert assms, auto simp: filterlim_uminus_at_bot asymp_equiv_uminus) lemma asymp_equivI'_const: assumes "((\x. f x / g x) \ c) F" "c \ 0" shows "f \[F] (\x. c * g x)" using tendsto_mult[OF assms(1) tendsto_const[of "inverse c"]] assms(2) by (intro asymp_equivI') (simp add: field_simps) lemma asymp_equivI'_inverse_const: assumes "((\x. f x / g x) \ inverse c) F" "c \ 0" shows "(\x. c * f x) \[F] g" using tendsto_mult[OF assms(1) tendsto_const[of "c"]] assms(2) by (intro asymp_equivI') (simp add: field_simps) lemma filterlim_at_bot_imp_at_infinity: "filterlim f at_bot F \ filterlim f at_infinity F" for f :: "_ \ real" using at_bot_le_at_infinity filterlim_mono by blast lemma asymp_equiv_imp_diff_smallo: assumes "f \[F] g" shows "(\x. f x - g x) \ o[F](g)" proof (rule landau_o.smallI) fix c :: real assume "c > 0" hence c: "min c 1 > 0" by simp let ?h = "\x. if f x = 0 \ g x = 0 then 1 else f x / g x" from assms have "((\x. ?h x - 1) \ 1 - 1) F" by (intro tendsto_diff asymp_equivD tendsto_const) from tendstoD[OF this c] show "eventually (\x. norm (f x - g x) \ c * norm (g x)) F" proof eventually_elim case (elim x) from elim have "norm (f x - g x) \ norm (f x / g x - 1) * norm (g x)" by (subst norm_mult [symmetric]) (auto split: if_splits simp add: algebra_simps) also have "norm (f x / g x - 1) * norm (g x) \ c * norm (g x)" using elim by (auto split: if_splits simp: mult_right_mono) finally show ?case . qed qed lemma asymp_equiv_altdef: "f \[F] g \ (\x. f x - g x) \ o[F](g)" by (rule iffI[OF asymp_equiv_imp_diff_smallo smallo_imp_asymp_equiv]) lemma asymp_equiv_0_left_iff [simp]: "(\_. 0) \[F] f \ eventually (\x. f x = 0) F" and asymp_equiv_0_right_iff [simp]: "f \[F] (\_. 0) \ eventually (\x. f x = 0) F" by (simp_all add: asymp_equiv_altdef landau_o.small_refl_iff) lemma asymp_equiv_sandwich_real: fixes f g l u :: "'a \ real" assumes "l \[F] g" "u \[F] g" "eventually (\x. f x \ {l x..u x}) F" shows "f \[F] g" unfolding asymp_equiv_altdef proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" have "eventually (\x. norm (f x - g x) \ max (norm (l x - g x)) (norm (u x - g x))) F" using assms(3) by eventually_elim auto moreover have "eventually (\x. norm (l x - g x) \ c * norm (g x)) F" "eventually (\x. norm (u x - g x) \ c * norm (g x)) F" using assms(1,2) by (auto simp: asymp_equiv_altdef dest: landau_o.smallD[OF _ c]) hence "eventually (\x. max (norm (l x - g x)) (norm (u x - g x)) \ c * norm (g x)) F" by eventually_elim simp ultimately show "eventually (\x. norm (f x - g x) \ c * norm (g x)) F" by eventually_elim (rule order.trans) qed lemma asymp_equiv_sandwich_real': fixes f g l u :: "'a \ real" assumes "f \[F] l" "f \[F] u" "eventually (\x. g x \ {l x..u x}) F" shows "f \[F] g" using asymp_equiv_sandwich_real[of l F f u g] assms by (simp add: asymp_equiv_sym) lemma asymp_equiv_sandwich_real'': fixes f g l u :: "'a \ real" assumes "l1 \[F] u1" "u1 \[F] l2" "l2 \[F] u2" "eventually (\x. f x \ {l1 x..u1 x}) F" "eventually (\x. g x \ {l2 x..u2 x}) F" shows "f \[F] g" by (rule asymp_equiv_sandwich_real[OF asymp_equiv_sandwich_real'[OF _ _ assms(5)] asymp_equiv_sandwich_real'[OF _ _ assms(5)] assms(4)]; blast intro: asymp_equiv_trans assms(1,2,3))+ end diff --git a/src/HOL/Probability/Characteristic_Functions.thy b/src/HOL/Probability/Characteristic_Functions.thy --- a/src/HOL/Probability/Characteristic_Functions.thy +++ b/src/HOL/Probability/Characteristic_Functions.thy @@ -1,554 +1,554 @@ (* Title: HOL/Probability/Characteristic_Functions.thy Authors: Jeremy Avigad (CMU), Luke Serafin (CMU), Johannes Hölzl (TUM) *) section \Characteristic Functions\ theory Characteristic_Functions imports Weak_Convergence Independent_Family Distributions begin lemma mult_min_right: "a \ 0 \ (a :: real) * min b c = min (a * b) (a * c)" by (metis min.absorb_iff2 min_def mult_left_mono) lemma sequentially_even_odd: assumes E: "eventually (\n. P (2 * n)) sequentially" and O: "eventually (\n. P (2 * n + 1)) sequentially" shows "eventually P sequentially" proof - from E obtain n_e where [intro]: "\n. n \ n_e \ P (2 * n)" by (auto simp: eventually_sequentially) moreover from O obtain n_o where [intro]: "\n. n \ n_o \ P (Suc (2 * n))" by (auto simp: eventually_sequentially) show ?thesis unfolding eventually_sequentially proof (intro exI allI impI) fix n assume "max (2 * n_e) (2 * n_o + 1) \ n" then show "P n" by (cases "even n") (auto elim!: evenE oddE ) qed qed lemma limseq_even_odd: assumes "(\n. f (2 * n)) \ (l :: 'a :: topological_space)" and "(\n. f (2 * n + 1)) \ l" shows "f \ l" using assms by (auto simp: filterlim_iff intro: sequentially_even_odd) subsection \Application of the FTC: integrating $e^ix$\ abbreviation iexp :: "real \ complex" where "iexp \ (\x. exp (\ * complex_of_real x))" lemma isCont_iexp [simp]: "isCont iexp x" by (intro continuous_intros) lemma has_vector_derivative_iexp[derivative_intros]: "(iexp has_vector_derivative \ * iexp x) (at x within s)" by (auto intro!: derivative_eq_intros simp: Re_exp Im_exp has_vector_derivative_complex_iff) lemma interval_integral_iexp: fixes a b :: real shows "(CLBINT x=a..b. iexp x) = \ * iexp a - \ * iexp b" by (subst interval_integral_FTC_finite [where F = "\x. -\ * iexp x"]) (auto intro!: derivative_eq_intros continuous_intros) subsection \The Characteristic Function of a Real Measure.\ definition char :: "real measure \ real \ complex" where "char M t = CLINT x|M. iexp (t * x)" lemma (in real_distribution) char_zero: "char M 0 = 1" unfolding char_def by (simp del: space_eq_univ add: prob_space) lemma (in prob_space) integrable_iexp: assumes f: "f \ borel_measurable M" "\x. Im (f x) = 0" shows "integrable M (\x. exp (\ * (f x)))" proof (intro integrable_const_bound [of _ 1]) from f have "\x. of_real (Re (f x)) = f x" by (simp add: complex_eq_iff) then show "AE x in M. cmod (exp (\ * f x)) \ 1" using norm_exp_i_times[of "Re (f x)" for x] by simp qed (insert f, simp) lemma (in real_distribution) cmod_char_le_1: "norm (char M t) \ 1" proof - have "norm (char M t) \ (\x. norm (iexp (t * x)) \M)" unfolding char_def by (intro integral_norm_bound) also have "\ \ 1" by (simp del: of_real_mult) finally show ?thesis . qed lemma (in real_distribution) isCont_char: "isCont (char M) t" unfolding continuous_at_sequentially proof safe fix X assume X: "X \ t" show "(char M \ X) \ char M t" unfolding comp_def char_def by (rule integral_dominated_convergence[where w="\_. 1"]) (auto intro!: tendsto_intros X) qed lemma (in real_distribution) char_measurable [measurable]: "char M \ borel_measurable borel" by (auto intro!: borel_measurable_continuous_onI continuous_at_imp_continuous_on isCont_char) subsection \Independence\ (* the automation can probably be improved *) lemma (in prob_space) char_distr_add: fixes X1 X2 :: "'a \ real" and t :: real assumes "indep_var borel X1 borel X2" shows "char (distr M borel (\\. X1 \ + X2 \)) t = char (distr M borel X1) t * char (distr M borel X2) t" proof - from assms have [measurable]: "random_variable borel X1" by (elim indep_var_rv1) from assms have [measurable]: "random_variable borel X2" by (elim indep_var_rv2) have "char (distr M borel (\\. X1 \ + X2 \)) t = (CLINT x|M. iexp (t * (X1 x + X2 x)))" by (simp add: char_def integral_distr) also have "\ = (CLINT x|M. iexp (t * (X1 x)) * iexp (t * (X2 x))) " by (simp add: field_simps exp_add) also have "\ = (CLINT x|M. iexp (t * (X1 x))) * (CLINT x|M. iexp (t * (X2 x)))" by (intro indep_var_lebesgue_integral indep_var_compose[unfolded comp_def, OF assms]) (auto intro!: integrable_iexp) also have "\ = char (distr M borel X1) t * char (distr M borel X2) t" by (simp add: char_def integral_distr) finally show ?thesis . qed lemma (in prob_space) char_distr_sum: "indep_vars (\i. borel) X A \ char (distr M borel (\\. \i\A. X i \)) t = (\i\A. char (distr M borel (X i)) t)" proof (induct A rule: infinite_finite_induct) case (insert x F) with indep_vars_subset[of "\_. borel" X "insert x F" F] show ?case by (auto simp add: char_distr_add indep_vars_sum) qed (simp_all add: char_def integral_distr prob_space del: distr_const) subsection \Approximations to $e^{ix}$\ text \Proofs from Billingsley, page 343.\ lemma CLBINT_I0c_power_mirror_iexp: fixes x :: real and n :: nat defines "f s m \ complex_of_real ((x - s) ^ m)" shows "(CLBINT s=0..x. f s n * iexp s) = x^Suc n / Suc n + (\ / Suc n) * (CLBINT s=0..x. f s (Suc n) * iexp s)" proof - have 1: "((\s. complex_of_real(-((x - s) ^ (Suc n) / (Suc n))) * iexp s) has_vector_derivative complex_of_real((x - s)^n) * iexp s + (\ * iexp s) * complex_of_real(-((x - s) ^ (Suc n) / (Suc n)))) (at s within A)" for s A by (intro derivative_eq_intros) auto let ?F = "\s. complex_of_real(-((x - s) ^ (Suc n) / (Suc n))) * iexp s" have "x^(Suc n) / (Suc n) = (CLBINT s=0..x. (f s n * iexp s + (\ * iexp s) * -(f s (Suc n) / (Suc n))))" (is "?LHS = ?RHS") proof - have "?RHS = (CLBINT s=0..x. (f s n * iexp s + (\ * iexp s) * complex_of_real(-((x - s) ^ (Suc n) / (Suc n)))))" by (cases "0 \ x") (auto intro!: simp: f_def[abs_def]) also have "... = ?F x - ?F 0" unfolding zero_ereal_def using 1 by (intro interval_integral_FTC_finite) (auto simp: f_def add_nonneg_eq_0_iff complex_eq_iff intro!: continuous_at_imp_continuous_on continuous_intros) finally show ?thesis by auto qed show ?thesis unfolding \?LHS = ?RHS\ f_def interval_lebesgue_integral_mult_right [symmetric] by (subst interval_lebesgue_integral_add(2) [symmetric]) (auto intro!: interval_integrable_isCont continuous_intros simp: zero_ereal_def complex_eq_iff) qed lemma iexp_eq1: fixes x :: real defines "f s m \ complex_of_real ((x - s) ^ m)" shows "iexp x = (\k \ n. (\ * x)^k / (fact k)) + ((\ ^ (Suc n)) / (fact n)) * (CLBINT s=0..x. (f s n) * (iexp s))" (is "?P n") proof (induction n) show "?P 0" by (auto simp add: field_simps interval_integral_iexp f_def zero_ereal_def) next fix n assume ih: "?P n" have **: "\a b :: real. a = -b \ a + b = 0" by linarith have *: "of_nat n * of_nat (fact n) \ - (of_nat (fact n)::complex)" unfolding of_nat_mult[symmetric] by (simp add: complex_eq_iff ** of_nat_add[symmetric] del: of_nat_mult of_nat_fact of_nat_add) show "?P (Suc n)" unfolding sum.atMost_Suc ih f_def CLBINT_I0c_power_mirror_iexp[of _ n] by (simp add: divide_simps add_eq_0_iff *) (simp add: field_simps) qed lemma iexp_eq2: fixes x :: real defines "f s m \ complex_of_real ((x - s) ^ m)" shows "iexp x = (\k\Suc n. (\*x)^k/fact k) + \^Suc n/fact n * (CLBINT s=0..x. f s n*(iexp s - 1))" proof - have isCont_f: "isCont (\s. f s n) x" for n x by (auto simp: f_def) let ?F = "\s. complex_of_real (-((x - s) ^ (Suc n) / real (Suc n)))" have calc1: "(CLBINT s=0..x. f s n * (iexp s - 1)) = (CLBINT s=0..x. f s n * iexp s) - (CLBINT s=0..x. f s n)" unfolding zero_ereal_def by (subst interval_lebesgue_integral_diff(2) [symmetric]) (simp_all add: interval_integrable_isCont isCont_f field_simps) have calc2: "(CLBINT s=0..x. f s n) = x^Suc n / Suc n" unfolding zero_ereal_def proof (subst interval_integral_FTC_finite [where a = 0 and b = x and f = "\s. f s n" and F = ?F]) show "(?F has_vector_derivative f y n) (at y within {min 0 x..max 0 x})" for y unfolding f_def by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros simp del: power_Suc simp add: add_nonneg_eq_0_iff) qed (auto intro: continuous_at_imp_continuous_on isCont_f) have calc3: "\ ^ (Suc (Suc n)) / (fact (Suc n)) = (\ ^ (Suc n) / (fact n)) * (\ / (Suc n))" by (simp add: field_simps) show ?thesis unfolding iexp_eq1 [where n = "Suc n" and x=x] calc1 calc2 calc3 unfolding f_def by (subst CLBINT_I0c_power_mirror_iexp [where n = n]) auto qed lemma abs_LBINT_I0c_abs_power_diff: "\LBINT s=0..x. \(x - s)^n\\ = \x ^ (Suc n) / (Suc n)\" proof - have "\LBINT s=0..x. \(x - s)^n\\ = \LBINT s=0..x. (x - s)^n\" proof cases assume "0 \ x \ even n" then have "(LBINT s=0..x. \(x - s)^n\) = LBINT s=0..x. (x - s)^n" by (auto simp add: zero_ereal_def power_even_abs power_abs min_absorb1 max_absorb2 intro!: interval_integral_cong) then show ?thesis by simp next assume "\ (0 \ x \ even n)" then have "(LBINT s=0..x. \(x - s)^n\) = LBINT s=0..x. -((x - s)^n)" by (auto simp add: zero_ereal_def power_abs min_absorb1 max_absorb2 ereal_min[symmetric] ereal_max[symmetric] power_minus_odd[symmetric] simp del: ereal_min ereal_max intro!: interval_integral_cong) also have "\ = - (LBINT s=0..x. (x - s)^n)" by (subst interval_lebesgue_integral_uminus, rule refl) finally show ?thesis by simp qed also have "LBINT s=0..x. (x - s)^n = x^Suc n / Suc n" proof - let ?F = "\t. - ((x - t)^(Suc n) / Suc n)" have "LBINT s=0..x. (x - s)^n = ?F x - ?F 0" unfolding zero_ereal_def by (intro interval_integral_FTC_finite continuous_at_imp_continuous_on - has_field_derivative_iff_has_vector_derivative[THEN iffD1]) + has_real_derivative_iff_has_vector_derivative[THEN iffD1]) (auto simp del: power_Suc intro!: derivative_eq_intros simp add: add_nonneg_eq_0_iff) also have "\ = x ^ (Suc n) / (Suc n)" by simp finally show ?thesis . qed finally show ?thesis . qed lemma iexp_approx1: "cmod (iexp x - (\k \ n. (\ * x)^k / fact k)) \ \x\^(Suc n) / fact (Suc n)" proof - have "iexp x - (\k \ n. (\ * x)^k / fact k) = ((\ ^ (Suc n)) / (fact n)) * (CLBINT s=0..x. (x - s)^n * (iexp s))" (is "?t1 = ?t2") by (subst iexp_eq1 [of _ n], simp add: field_simps) then have "cmod (?t1) = cmod (?t2)" by simp also have "\ = (1 / of_nat (fact n)) * cmod (CLBINT s=0..x. (x - s)^n * (iexp s))" by (simp add: norm_mult norm_divide norm_power) also have "\ \ (1 / of_nat (fact n)) * \LBINT s=0..x. cmod ((x - s)^n * (iexp s))\" by (intro mult_left_mono interval_integral_norm2) (auto simp: zero_ereal_def intro: interval_integrable_isCont) also have "\ \ (1 / of_nat (fact n)) * \LBINT s=0..x. \(x - s)^n\\" by (simp add: norm_mult del: of_real_diff of_real_power) also have "\ \ (1 / of_nat (fact n)) * \x ^ (Suc n) / (Suc n)\" by (simp add: abs_LBINT_I0c_abs_power_diff) also have "1 / real_of_nat (fact n::nat) * \x ^ Suc n / real (Suc n)\ = \x\ ^ Suc n / fact (Suc n)" by (simp add: abs_mult power_abs) finally show ?thesis . qed lemma iexp_approx2: "cmod (iexp x - (\k \ n. (\ * x)^k / fact k)) \ 2 * \x\^n / fact n" proof (induction n) \ \really cases\ case (Suc n) have *: "\a b. interval_lebesgue_integrable lborel a b f \ interval_lebesgue_integrable lborel a b g \ \LBINT s=a..b. f s\ \ \LBINT s=a..b. g s\" if f: "\s. 0 \ f s" and g: "\s. f s \ g s" for f g :: "_ \ real" using order_trans[OF f g] f g unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def set_integrable_def by (auto simp: integral_nonneg_AE[OF AE_I2] intro!: integral_mono mult_mono) have "iexp x - (\k \ Suc n. (\ * x)^k / fact k) = ((\ ^ (Suc n)) / (fact n)) * (CLBINT s=0..x. (x - s)^n * (iexp s - 1))" (is "?t1 = ?t2") unfolding iexp_eq2 [of x n] by (simp add: field_simps) then have "cmod (?t1) = cmod (?t2)" by simp also have "\ = (1 / (fact n)) * cmod (CLBINT s=0..x. (x - s)^n * (iexp s - 1))" by (simp add: norm_mult norm_divide norm_power) also have "\ \ (1 / (fact n)) * \LBINT s=0..x. cmod ((x - s)^n * (iexp s - 1))\" by (intro mult_left_mono interval_integral_norm2) (auto intro!: interval_integrable_isCont simp: zero_ereal_def) also have "\ = (1 / (fact n)) * \LBINT s=0..x. abs ((x - s)^n) * cmod((iexp s - 1))\" by (simp add: norm_mult del: of_real_diff of_real_power) also have "\ \ (1 / (fact n)) * \LBINT s=0..x. abs ((x - s)^n) * 2\" by (intro mult_left_mono * order_trans [OF norm_triangle_ineq4]) (auto simp: mult_ac zero_ereal_def intro!: interval_integrable_isCont) also have "\ = (2 / (fact n)) * \x ^ (Suc n) / (Suc n)\" by (simp add: abs_LBINT_I0c_abs_power_diff abs_mult) also have "2 / fact n * \x ^ Suc n / real (Suc n)\ = 2 * \x\ ^ Suc n / (fact (Suc n))" by (simp add: abs_mult power_abs) finally show ?case . qed (insert norm_triangle_ineq4[of "iexp x" 1], simp) lemma (in real_distribution) char_approx1: assumes integrable_moments: "\k. k \ n \ integrable M (\x. x^k)" shows "cmod (char M t - (\k \ n. ((\ * t)^k / fact k) * expectation (\x. x^k))) \ (2*\t\^n / fact n) * expectation (\x. \x\^n)" (is "cmod (char M t - ?t1) \ _") proof - have integ_iexp: "integrable M (\x. iexp (t * x))" by (intro integrable_const_bound) auto define c where [abs_def]: "c k x = (\ * t)^k / fact k * complex_of_real (x^k)" for k x have integ_c: "\k. k \ n \ integrable M (\x. c k x)" unfolding c_def by (intro integrable_mult_right integrable_of_real integrable_moments) have "k \ n \ expectation (c k) = (\*t) ^ k * (expectation (\x. x ^ k)) / fact k" for k unfolding c_def integral_mult_right_zero integral_complex_of_real by simp then have "norm (char M t - ?t1) = norm (char M t - (CLINT x | M. (\k \ n. c k x)))" by (simp add: integ_c) also have "\ = norm ((CLINT x | M. iexp (t * x) - (\k \ n. c k x)))" unfolding char_def by (subst Bochner_Integration.integral_diff[OF integ_iexp]) (auto intro!: integ_c) also have "\ \ expectation (\x. cmod (iexp (t * x) - (\k \ n. c k x)))" by (intro integral_norm_bound) also have "\ \ expectation (\x. 2 * \t\ ^ n / fact n * \x\ ^ n)" proof (rule integral_mono) show "integrable M (\x. cmod (iexp (t * x) - (\k\n. c k x)))" by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp show "integrable M (\x. 2 * \t\ ^ n / fact n * \x\ ^ n)" unfolding power_abs[symmetric] by (intro integrable_mult_right integrable_abs integrable_moments) simp show "cmod (iexp (t * x) - (\k\n. c k x)) \ 2 * \t\ ^ n / fact n * \x\ ^ n" for x using iexp_approx2[of "t * x" n] by (auto simp add: abs_mult field_simps c_def) qed finally show ?thesis unfolding integral_mult_right_zero . qed lemma (in real_distribution) char_approx2: assumes integrable_moments: "\k. k \ n \ integrable M (\x. x ^ k)" shows "cmod (char M t - (\k \ n. ((\ * t)^k / fact k) * expectation (\x. x^k))) \ (\t\^n / fact (Suc n)) * expectation (\x. min (2 * \x\^n * Suc n) (\t\ * \x\^Suc n))" (is "cmod (char M t - ?t1) \ _") proof - have integ_iexp: "integrable M (\x. iexp (t * x))" by (intro integrable_const_bound) auto define c where [abs_def]: "c k x = (\ * t)^k / fact k * complex_of_real (x^k)" for k x have integ_c: "\k. k \ n \ integrable M (\x. c k x)" unfolding c_def by (intro integrable_mult_right integrable_of_real integrable_moments) have *: "min (2 * \t * x\^n / fact n) (\t * x\^Suc n / fact (Suc n)) = \t\^n / fact (Suc n) * min (2 * \x\^n * real (Suc n)) (\t\ * \x\^(Suc n))" for x apply (subst mult_min_right) apply simp apply (rule arg_cong2[where f=min]) apply (simp_all add: field_simps abs_mult del: fact_Suc) apply (simp_all add: field_simps) done have "k \ n \ expectation (c k) = (\*t) ^ k * (expectation (\x. x ^ k)) / fact k" for k unfolding c_def integral_mult_right_zero integral_complex_of_real by simp then have "norm (char M t - ?t1) = norm (char M t - (CLINT x | M. (\k \ n. c k x)))" by (simp add: integ_c) also have "\ = norm ((CLINT x | M. iexp (t * x) - (\k \ n. c k x)))" unfolding char_def by (subst Bochner_Integration.integral_diff[OF integ_iexp]) (auto intro!: integ_c) also have "\ \ expectation (\x. cmod (iexp (t * x) - (\k \ n. c k x)))" by (rule integral_norm_bound) also have "\ \ expectation (\x. min (2 * \t * x\^n / fact n) (\t * x\^(Suc n) / fact (Suc n)))" (is "_ \ expectation ?f") proof (rule integral_mono) show "integrable M (\x. cmod (iexp (t * x) - (\k\n. c k x)))" by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp show "integrable M ?f" by (rule Bochner_Integration.integrable_bound[where f="\x. 2 * \t * x\ ^ n / fact n"]) (auto simp: integrable_moments power_abs[symmetric] power_mult_distrib) show "cmod (iexp (t * x) - (\k\n. c k x)) \ ?f x" for x using iexp_approx1[of "t * x" n] iexp_approx2[of "t * x" n] by (auto simp add: abs_mult field_simps c_def intro!: min.boundedI) qed also have "\ = (\t\^n / fact (Suc n)) * expectation (\x. min (2 * \x\^n * Suc n) (\t\ * \x\^Suc n))" unfolding * proof (rule Bochner_Integration.integral_mult_right) show "integrable M (\x. min (2 * \x\ ^ n * real (Suc n)) (\t\ * \x\ ^ Suc n))" by (rule Bochner_Integration.integrable_bound[where f="\x. 2 * \x\ ^ n * real (Suc n)"]) (auto simp: integrable_moments power_abs[symmetric] power_mult_distrib) qed finally show ?thesis unfolding integral_mult_right_zero . qed lemma (in real_distribution) char_approx3: fixes t assumes integrable_1: "integrable M (\x. x)" and integral_1: "expectation (\x. x) = 0" and integrable_2: "integrable M (\x. x^2)" and integral_2: "variance (\x. x) = \2" shows "cmod (char M t - (1 - t^2 * \2 / 2)) \ (t^2 / 6) * expectation (\x. min (6 * x^2) (abs t * (abs x)^3) )" proof - note real_distribution.char_approx2 [of M 2 t, simplified] have [simp]: "prob UNIV = 1" by (metis prob_space space_eq_univ) from integral_2 have [simp]: "expectation (\x. x * x) = \2" by (simp add: integral_1 numeral_eq_Suc) have 1: "k \ 2 \ integrable M (\x. x^k)" for k using assms by (auto simp: eval_nat_numeral le_Suc_eq) note char_approx1 note 2 = char_approx1 [of 2 t, OF 1, simplified] have "cmod (char M t - (\k\2. (\ * t) ^ k * (expectation (\x. x ^ k)) / (fact k))) \ t\<^sup>2 * expectation (\x. min (6 * x\<^sup>2) (\t\ * \x\ ^ 3)) / fact (3::nat)" using char_approx2 [of 2 t, OF 1] by simp also have "(\k\2. (\ * t) ^ k * expectation (\x. x ^ k) / (fact k)) = 1 - t^2 * \2 / 2" by (simp add: complex_eq_iff numeral_eq_Suc integral_1 Re_divide Im_divide) also have "fact 3 = 6" by (simp add: eval_nat_numeral) also have "t\<^sup>2 * expectation (\x. min (6 * x\<^sup>2) (\t\ * \x\ ^ 3)) / 6 = t\<^sup>2 / 6 * expectation (\x. min (6 * x\<^sup>2) (\t\ * \x\ ^ 3))" by (simp add: field_simps) finally show ?thesis . qed text \ This is a more familiar textbook formulation in terms of random variables, but we will use the previous version for the CLT. \ lemma (in prob_space) char_approx3': fixes \ :: "real measure" and X assumes rv_X [simp]: "random_variable borel X" and [simp]: "integrable M X" "integrable M (\x. (X x)^2)" "expectation X = 0" and var_X: "variance X = \2" and \_def: "\ = distr M borel X" shows "cmod (char \ t - (1 - t^2 * \2 / 2)) \ (t^2 / 6) * expectation (\x. min (6 * (X x)^2) (\t\ * \X x\^3))" using var_X unfolding \_def apply (subst integral_distr [symmetric, OF rv_X], simp) apply (intro real_distribution.char_approx3) apply (auto simp add: integrable_distr_eq integral_distr) done text \ this is the formulation in the book -- in terms of a random variable *with* the distribution, rather the distribution itself. I don't know which is more useful, though in principal we can go back and forth between them. \ lemma (in prob_space) char_approx1': fixes \ :: "real measure" and X assumes integrable_moments : "\k. k \ n \ integrable M (\x. X x ^ k)" and rv_X[measurable]: "random_variable borel X" and \_distr : "distr M borel X = \" shows "cmod (char \ t - (\k \ n. ((\ * t)^k / fact k) * expectation (\x. (X x)^k))) \ (2 * \t\^n / fact n) * expectation (\x. \X x\^n)" unfolding \_distr[symmetric] apply (subst (1 2) integral_distr [symmetric, OF rv_X], simp, simp) apply (intro real_distribution.char_approx1[of "distr M borel X" n t] real_distribution_distr rv_X) apply (auto simp: integrable_distr_eq integrable_moments) done subsection \Calculation of the Characteristic Function of the Standard Distribution\ abbreviation "std_normal_distribution \ density lborel std_normal_density" (* TODO: should this be an instance statement? *) lemma real_dist_normal_dist: "real_distribution std_normal_distribution" using prob_space_normal_density by (auto simp: real_distribution_def real_distribution_axioms_def) lemma std_normal_distribution_even_moments: fixes k :: nat shows "(LINT x|std_normal_distribution. x^(2 * k)) = fact (2 * k) / (2^k * fact k)" and "integrable std_normal_distribution (\x. x^(2 * k))" using integral_std_normal_moment_even[of k] by (subst integral_density) (auto simp: normal_density_nonneg integrable_density intro: integrable.intros std_normal_moment_even) lemma integrable_std_normal_distribution_moment: "integrable std_normal_distribution (\x. x^k)" by (auto simp: normal_density_nonneg integrable_std_normal_moment integrable_density) lemma integral_std_normal_distribution_moment_odd: "odd k \ integral\<^sup>L std_normal_distribution (\x. x^k) = 0" using integral_std_normal_moment_odd[of "(k - 1) div 2"] by (auto simp: integral_density normal_density_nonneg elim: oddE) lemma std_normal_distribution_even_moments_abs: fixes k :: nat shows "(LINT x|std_normal_distribution. \x\^(2 * k)) = fact (2 * k) / (2^k * fact k)" using integral_std_normal_moment_even[of k] by (subst integral_density) (auto simp: normal_density_nonneg power_even_abs) lemma std_normal_distribution_odd_moments_abs: fixes k :: nat shows "(LINT x|std_normal_distribution. \x\^(2 * k + 1)) = sqrt (2 / pi) * 2 ^ k * fact k" using integral_std_normal_moment_abs_odd[of k] by (subst integral_density) (auto simp: normal_density_nonneg) theorem char_std_normal_distribution: "char std_normal_distribution = (\t. complex_of_real (exp (- (t^2) / 2)))" proof (intro ext LIMSEQ_unique) fix t :: real let ?f' = "\k. (\ * t)^k / fact k * (LINT x | std_normal_distribution. x^k)" let ?f = "\n. (\k \ n. ?f' k)" show "?f \ exp (-(t^2) / 2)" proof (rule limseq_even_odd) have "(\ * complex_of_real t) ^ (2 * a) / (2 ^ a * fact a) = (- ((complex_of_real t)\<^sup>2 / 2)) ^ a / fact a" for a by (subst power_mult) (simp add: field_simps uminus_power_if power_mult) then have *: "?f (2 * n) = complex_of_real (\k < Suc n. (1 / fact k) * (- (t^2) / 2)^k)" for n :: nat unfolding of_real_sum by (intro sum.reindex_bij_witness_not_neutral[symmetric, where i="\n. n div 2" and j="\n. 2 * n" and T'="{i. i \ 2 * n \ odd i}" and S'="{}"]) (auto simp: integral_std_normal_distribution_moment_odd std_normal_distribution_even_moments) show "(\n. ?f (2 * n)) \ exp (-(t^2) / 2)" unfolding * using exp_converges[where 'a=real] by (intro tendsto_of_real LIMSEQ_Suc) (auto simp: inverse_eq_divide sums_def [symmetric]) have **: "?f (2 * n + 1) = ?f (2 * n)" for n proof - have "?f (2 * n + 1) = ?f (2 * n) + ?f' (2 * n + 1)" by simp also have "?f' (2 * n + 1) = 0" by (subst integral_std_normal_distribution_moment_odd) simp_all finally show "?f (2 * n + 1) = ?f (2 * n)" by simp qed show "(\n. ?f (2 * n + 1)) \ exp (-(t^2) / 2)" unfolding ** by fact qed have **: "(\n. x ^ n / fact n) \ 0" for x :: real using summable_LIMSEQ_zero [OF summable_exp] by (auto simp add: inverse_eq_divide) let ?F = "\n. 2 * \t\ ^ n / fact n * (LINT x|std_normal_distribution. \x\ ^ n)" show "?f \ char std_normal_distribution t" proof (rule metric_tendsto_imp_tendsto[OF limseq_even_odd]) show "(\n. ?F (2 * n)) \ 0" proof (rule Lim_transform_eventually) show "\\<^sub>F n in sequentially. 2 * ((t^2 / 2)^n / fact n) = ?F (2 * n)" unfolding std_normal_distribution_even_moments_abs by (simp add: power_mult power_divide) qed (intro tendsto_mult_right_zero **) have *: "?F (2 * n + 1) = (2 * \t\ * sqrt (2 / pi)) * ((2 * t^2)^n * fact n / fact (2 * n + 1))" for n unfolding std_normal_distribution_odd_moments_abs by (simp add: field_simps power_mult[symmetric] power_even_abs) have "norm ((2 * t\<^sup>2) ^ n * fact n / fact (2 * n + 1)) \ (2 * t\<^sup>2) ^ n / fact n" for n using mult_mono[OF _ square_fact_le_2_fact, of 1 "1 + 2 * real n" n] by (auto simp add: divide_simps intro!: mult_left_mono) then show "(\n. ?F (2 * n + 1)) \ 0" unfolding * by (intro tendsto_mult_right_zero Lim_null_comparison [OF _ ** [of "2 * t\<^sup>2"]]) auto show "\\<^sub>F n in sequentially. dist (?f n) (char std_normal_distribution t) \ dist (?F n) 0" using real_distribution.char_approx1[OF real_dist_normal_dist integrable_std_normal_distribution_moment] by (auto simp: dist_norm integral_nonneg_AE norm_minus_commute) qed qed end diff --git a/src/HOL/Probability/Sinc_Integral.thy b/src/HOL/Probability/Sinc_Integral.thy --- a/src/HOL/Probability/Sinc_Integral.thy +++ b/src/HOL/Probability/Sinc_Integral.thy @@ -1,403 +1,403 @@ (* Title: HOL/Probability/Sinc_Integral.thy Authors: Jeremy Avigad (CMU), Luke Serafin (CMU) Authors: Johannes Hölzl, TU München *) section \Integral of sinc\ theory Sinc_Integral imports Distributions begin subsection \Various preparatory integrals\ text \ Naming convention The theorem name consists of the following parts: \<^item> Kind of integral: \has_bochner_integral\ / \integrable\ / \LBINT\ \<^item> Interval: Interval (0 / infinity / open / closed) (infinity / open / closed) \<^item> Name of the occurring constants: power, exp, m (for minus), scale, sin, $\ldots$ \ lemma has_bochner_integral_I0i_power_exp_m': "has_bochner_integral lborel (\x. x^k * exp (-x) * indicator {0 ..} x::real) (fact k)" using nn_intergal_power_times_exp_Ici[of k] by (intro has_bochner_integral_nn_integral) (auto simp: nn_integral_set_ennreal split: split_indicator) lemma has_bochner_integral_I0i_power_exp_m: "has_bochner_integral lborel (\x. x^k * exp (-x) * indicator {0 <..} x::real) (fact k)" using AE_lborel_singleton[of 0] by (intro has_bochner_integral_cong_AE[THEN iffD1, OF _ _ _ has_bochner_integral_I0i_power_exp_m']) (auto split: split_indicator) lemma integrable_I0i_exp_mscale: "0 < (u::real) \ set_integrable lborel {0 <..} (\x. exp (-(x * u)))" using lborel_integrable_real_affine_iff[of u "\x. indicator {0 <..} x *\<^sub>R exp (- x)" 0] has_bochner_integral_I0i_power_exp_m[of 0] by (simp add: indicator_def zero_less_mult_iff mult_ac integrable.intros set_integrable_def) lemma LBINT_I0i_exp_mscale: "0 < (u::real) \ LBINT x=0..\. exp (-(x * u)) = 1 / u" using lborel_integral_real_affine[of u "\x. indicator {0<..} x *\<^sub>R exp (- x)" 0] has_bochner_integral_I0i_power_exp_m[of 0] by (auto simp: indicator_def zero_less_mult_iff interval_lebesgue_integral_0_infty set_lebesgue_integral_def field_simps dest!: has_bochner_integral_integral_eq) lemma LBINT_I0c_exp_mscale_sin: "LBINT x=0..t. exp (-(u * x)) * sin x = (1 / (1 + u^2)) * (1 - exp (-(u * t)) * (u * sin t + cos t))" (is "_ = ?F t") unfolding zero_ereal_def proof (subst interval_integral_FTC_finite) show "(?F has_vector_derivative exp (- (u * x)) * sin x) (at x within {min 0 t..max 0 t})" for x by (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative[symmetric] power2_eq_square) + simp: has_real_derivative_iff_has_vector_derivative[symmetric] power2_eq_square) (simp_all add: field_simps add_nonneg_eq_0_iff) qed (auto intro: continuous_at_imp_continuous_on) lemma LBINT_I0i_exp_mscale_sin: assumes "0 < x" shows "LBINT u=0..\. \exp (-u * x) * sin x\ = \sin x\ / x" proof (subst interval_integral_FTC_nonneg) let ?F = "\u. 1 / x * (1 - exp (- u * x)) * \sin x\" show "\t. (?F has_real_derivative \exp (- t * x) * sin x\) (at t)" using \0 < x\ by (auto intro!: derivative_eq_intros simp: abs_mult) show "((?F \ real_of_ereal) \ 0) (at_right 0)" using \0 < x\ by (auto simp: zero_ereal_def ereal_tendsto_simps intro!: tendsto_eq_intros) have *: "((\t. exp (- t * x)) \ 0) at_top" using \0 < x\ by (auto intro!: exp_at_bot[THEN filterlim_compose] filterlim_tendsto_pos_mult_at_top filterlim_ident simp: filterlim_uminus_at_bot mult.commute[of _ x]) show "((?F \ real_of_ereal) \ \sin x\ / x) (at_left \)" using \0 < x\ unfolding ereal_tendsto_simps by (intro filterlim_compose[OF _ *]) (auto intro!: tendsto_eq_intros filterlim_ident) qed auto lemma shows integrable_inverse_1_plus_square: "set_integrable lborel (einterval (-\) \) (\x. inverse (1 + x^2))" and LBINT_inverse_1_plus_square: "LBINT x=-\..\. inverse (1 + x^2) = pi" proof - have 1: "- (pi / 2) < x \ x * 2 < pi \ (tan has_real_derivative 1 + (tan x)\<^sup>2) (at x)" for x using cos_gt_zero_pi[of x] by (subst tan_sec) (auto intro!: DERIV_tan simp: power_inverse) have 2: "- (pi / 2) < x \ x * 2 < pi \ isCont (\x. 1 + (tan x)\<^sup>2) x" for x using cos_gt_zero_pi[of x] by auto have 3: "\- (pi / 2) < x; x * 2 < pi\ \ isCont (\x. inverse (1 + x\<^sup>2)) (tan x)" for x by (rule continuous_intros | simp add: add_nonneg_eq_0_iff)+ show "LBINT x=-\..\. inverse (1 + x^2) = pi" by (subst interval_integral_substitution_nonneg[of "-pi/2" "pi/2" tan "\x. 1 + (tan x)^2"]) (auto intro: derivative_eq_intros 1 2 3 filterlim_tan_at_right simp add: ereal_tendsto_simps filterlim_tan_at_left add_nonneg_eq_0_iff set_integrable_def) show "set_integrable lborel (einterval (-\) \) (\x. inverse (1 + x^2))" by (subst interval_integral_substitution_nonneg[of "-pi/2" "pi/2" tan "\x. 1 + (tan x)^2"]) (auto intro: derivative_eq_intros 1 2 3 filterlim_tan_at_right simp add: ereal_tendsto_simps filterlim_tan_at_left add_nonneg_eq_0_iff set_integrable_def) qed lemma shows integrable_I0i_1_div_plus_square: "interval_lebesgue_integrable lborel 0 \ (\x. 1 / (1 + x^2))" and LBINT_I0i_1_div_plus_square: "LBINT x=0..\. 1 / (1 + x^2) = pi / 2" proof - have 1: "0 < x \ x * 2 < pi \ (tan has_real_derivative 1 + (tan x)\<^sup>2) (at x)" for x using cos_gt_zero_pi[of x] by (subst tan_sec) (auto intro!: DERIV_tan simp: power_inverse) have 2: "0 < x \ x * 2 < pi \ isCont (\x. 1 + (tan x)\<^sup>2) x" for x using cos_gt_zero_pi[of x] by auto show "LBINT x=0..\. 1 / (1 + x^2) = pi / 2" by (subst interval_integral_substitution_nonneg[of "0" "pi/2" tan "\x. 1 + (tan x)^2"]) (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff set_integrable_def) show "interval_lebesgue_integrable lborel 0 \ (\x. 1 / (1 + x^2))" unfolding interval_lebesgue_integrable_def by (subst interval_integral_substitution_nonneg[of "0" "pi/2" tan "\x. 1 + (tan x)^2"]) (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff set_integrable_def) qed section \The sinc function, and the sine integral (Si)\ abbreviation sinc :: "real \ real" where "sinc \ (\x. if x = 0 then 1 else sin x / x)" lemma sinc_at_0: "((\x. sin x / x::real) \ 1) (at 0)" using DERIV_sin [of 0] by (auto simp add: has_field_derivative_def field_has_derivative_at) lemma isCont_sinc: "isCont sinc x" proof cases assume "x = 0" then show ?thesis using LIM_equal [where g = "\x. sin x / x" and a=0 and f=sinc and l=1] by (auto simp: isCont_def sinc_at_0) next assume "x \ 0" show ?thesis by (rule continuous_transform_within [where d = "abs x" and f = "\x. sin x / x"]) (auto simp add: dist_real_def \x \ 0\) qed lemma continuous_on_sinc[continuous_intros]: "continuous_on S f \ continuous_on S (\x. sinc (f x))" using continuous_on_compose[of S f sinc, OF _ continuous_at_imp_continuous_on] by (auto simp: isCont_sinc) lemma borel_measurable_sinc[measurable]: "sinc \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_at_imp_continuous_on ballI isCont_sinc) lemma sinc_AE: "AE x in lborel. sin x / x = sinc x" by (rule AE_I [where N = "{0}"], auto) definition Si :: "real \ real" where "Si t \ LBINT x=0..t. sin x / x" lemma sinc_neg [simp]: "sinc (- x) = sinc x" by auto (* TODO: Determine whether this can reasonably be eliminated by redefining Si. *) lemma Si_alt_def : "Si t = LBINT x=0..t. sinc x" proof cases assume "0 \ t" then show ?thesis using AE_lborel_singleton[of 0] by (auto simp: Si_def intro!: interval_lebesgue_integral_cong_AE) next assume "\ 0 \ t" then show ?thesis unfolding Si_def using AE_lborel_singleton[of 0] by (subst (1 2) interval_integral_endpoints_reverse) (auto simp: Si_def intro!: interval_lebesgue_integral_cong_AE) qed lemma Si_neg: assumes "T \ 0" shows "Si (- T) = - Si T" proof - have "LBINT x=ereal 0..T. -1 *\<^sub>R sinc (- x) = LBINT y= ereal (- 0)..ereal (- T). sinc y" by (rule interval_integral_substitution_finite [OF assms]) (auto intro: derivative_intros continuous_at_imp_continuous_on isCont_sinc) also have "(LBINT x=ereal 0..T. -1 *\<^sub>R sinc (- x)) = -(LBINT x=ereal 0..T. sinc x)" by (subst sinc_neg) (simp_all add: interval_lebesgue_integral_uminus) finally have *: "-(LBINT x=ereal 0..T. sinc x) = LBINT y= ereal 0..ereal (- T). sinc y" by simp show ?thesis using assms unfolding Si_alt_def by (subst zero_ereal_def)+ (auto simp add: * [symmetric]) qed lemma integrable_sinc': "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\t. sin (t * \) / t)" proof - have *: "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\t. \ * sinc (t * \))" by (intro interval_lebesgue_integrable_mult_right interval_integrable_isCont continuous_within_compose3 [OF isCont_sinc]) auto show ?thesis by (rule interval_lebesgue_integrable_cong_AE[THEN iffD1, OF _ _ _ *]) (insert AE_lborel_singleton[of 0], auto) qed (* TODO: need a better version of FTC2 *) lemma DERIV_Si: "(Si has_real_derivative sinc x) (at x)" proof - have "(at x within {min 0 (x - 1)..max 0 (x + 1)}) = at x" by (intro at_within_interior) auto moreover have "min 0 (x - 1) \ x" "x \ max 0 (x + 1)" by auto ultimately show ?thesis using interval_integral_FTC2[of "min 0 (x - 1)" 0 "max 0 (x + 1)" sinc x] by (auto simp: continuous_at_imp_continuous_on isCont_sinc Si_alt_def[abs_def] zero_ereal_def - has_field_derivative_iff_has_vector_derivative + has_real_derivative_iff_has_vector_derivative split del: if_split) qed lemma isCont_Si: "isCont Si x" using DERIV_Si by (rule DERIV_isCont) lemma borel_measurable_Si[measurable]: "Si \ borel_measurable borel" by (auto intro: isCont_Si continuous_at_imp_continuous_on borel_measurable_continuous_onI) lemma Si_at_top_LBINT: "((\t. (LBINT x=0..\. exp (-(x * t)) * (x * sin t + cos t) / (1 + x^2))) \ 0) at_top" proof - let ?F = "\x t. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2) :: real" have int: "set_integrable lborel {0<..} (\x. exp (- x) * (x + 1) :: real)" unfolding distrib_left using has_bochner_integral_I0i_power_exp_m[of 0] has_bochner_integral_I0i_power_exp_m[of 1] by (intro set_integral_add) (auto dest!: integrable.intros simp: ac_simps set_integrable_def) have "((\t::real. LBINT x:{0<..}. ?F x t) \ LBINT x::real:{0<..}. 0) at_top" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence_at_top[OF _ _ int [unfolded set_integrable_def]], simp_all del: abs_divide split: split_indicator) have *: "0 < x \ \x * sin t + cos t\ / (1 + x\<^sup>2) \ (x * 1 + 1) / 1" for x t :: real by (intro frac_le abs_triangle_ineq[THEN order_trans] add_mono) (auto simp add: abs_mult simp del: mult_1_right intro!: mult_mono) then have "1 \ t \ 0 < x \ \?F x t\ \ exp (- x) * (x + 1)" for x t :: real by (auto simp add: abs_mult times_divide_eq_right[symmetric] simp del: times_divide_eq_right intro!: mult_mono) then show "\\<^sub>F i in at_top. AE x in lborel. 0 < x \ \?F x i\ \ exp (- x) * (x + 1)" by (auto intro: eventually_mono eventually_ge_at_top[of "1::real"]) show "AE x in lborel. 0 < x \ (?F x \ 0) at_top" proof (intro always_eventually impI allI) fix x :: real assume "0 < x" show "((\t. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2)) \ 0) at_top" proof (rule Lim_null_comparison) show "\\<^sub>F t in at_top. norm (?F x t) \ exp (- (x * t)) * (x + 1)" using eventually_ge_at_top[of "1::real"] * \0 < x\ by (auto simp add: abs_mult times_divide_eq_right[symmetric] simp del: times_divide_eq_right intro!: mult_mono elim: eventually_mono) show "((\t. exp (- (x * t)) * (x + 1)) \ 0) at_top" by (auto simp: filterlim_uminus_at_top [symmetric] intro!: filterlim_tendsto_pos_mult_at_top[OF tendsto_const] \0 filterlim_ident tendsto_mult_left_zero filterlim_compose[OF exp_at_bot]) qed qed qed then show "((\t. (LBINT x=0..\. exp (-(x * t)) * (x * sin t + cos t) / (1 + x^2))) \ 0) at_top" by (simp add: interval_lebesgue_integral_0_infty set_lebesgue_integral_def) qed lemma Si_at_top_integrable: assumes "t \ 0" shows "interval_lebesgue_integrable lborel 0 \ (\x. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2))" using \0 \ t\ unfolding le_less proof assume "0 = t" then show ?thesis using integrable_I0i_1_div_plus_square by simp next assume [arith]: "0 < t" have *: "0 \ a \ 0 < x \ a / (1 + x\<^sup>2) \ a" for a x :: real using zero_le_power2[of x, arith] using divide_left_mono[of 1 "1+x\<^sup>2" a] by auto have "set_integrable lborel {0<..} (\x. (exp (- x) * x) * (sin t/t) + exp (- x) * cos t)" using has_bochner_integral_I0i_power_exp_m[of 0] has_bochner_integral_I0i_power_exp_m[of 1] by (intro set_integral_add set_integrable_mult_left) (auto dest!: integrable.intros simp: ac_simps set_integrable_def) from lborel_integrable_real_affine[OF this [unfolded set_integrable_def], of t 0] show ?thesis unfolding interval_lebesgue_integral_0_infty set_integrable_def by (rule Bochner_Integration.integrable_bound) (auto simp: field_simps * split: split_indicator) qed lemma Si_at_top: "(Si \ pi / 2) at_top" proof - have \: "\\<^sub>F t in at_top. pi / 2 - (LBINT u=0..\. exp (-(u * t)) * (u * sin t + cos t) / (1 + u^2)) = Si t" using eventually_ge_at_top[of 0] proof eventually_elim fix t :: real assume "t \ 0" have "Si t = LBINT x=0..t. sin x * (LBINT u=0..\. exp (-(u * x)))" unfolding Si_def using \0 \ t\ by (intro interval_integral_discrete_difference[where X="{0}"]) (auto simp: LBINT_I0i_exp_mscale) also have "\ = LBINT x. (LBINT u=ereal 0..\. indicator {0 <..< t} x *\<^sub>R sin x * exp (-(u * x)))" using \0\t\ by (simp add: zero_ereal_def interval_lebesgue_integral_le_eq mult_ac set_lebesgue_integral_def) also have "\ = LBINT x. (LBINT u. indicator ({0<..} \ {0 <..< t}) (u, x) *\<^sub>R (sin x * exp (-(u * x))))" apply (subst interval_integral_Ioi) unfolding set_lebesgue_integral_def by(simp_all add: indicator_times ac_simps ) also have "\ = LBINT u. (LBINT x. indicator ({0<..} \ {0 <..< t}) (u, x) *\<^sub>R (sin x * exp (-(u * x))))" proof (intro lborel_pair.Fubini_integral[symmetric] lborel_pair.Fubini_integrable) show "(\(x, y). indicator ({0<..} \ {0<..R (sin x * exp (- (y * x)))) \ borel_measurable (lborel \\<^sub>M lborel)" (is "?f \ borel_measurable _") by measurable have "AE x in lborel. indicator {0..t} x *\<^sub>R abs (sinc x) = LBINT y. norm (?f (x, y))" using AE_lborel_singleton[of 0] AE_lborel_singleton[of t] proof eventually_elim fix x :: real assume x: "x \ 0" "x \ t" have "LBINT y. \indicator ({0<..} \ {0<..R (sin x * exp (- (y * x)))\ = LBINT y. \sin x\ * exp (- (y * x)) * indicator {0<..} y * indicator {0<.. = \sin x\ * indicator {0<... exp (- (y * x)))" by (cases "x > 0") (auto simp: field_simps interval_lebesgue_integral_0_infty set_lebesgue_integral_def split: split_indicator) also have "\ = \sin x\ * indicator {0<.. 0") (auto simp add: LBINT_I0i_exp_mscale) also have "\ = indicator {0..t} x *\<^sub>R \sinc x\" using x by (simp add: field_simps split: split_indicator) finally show "indicator {0..t} x *\<^sub>R abs (sinc x) = LBINT y. norm (?f (x, y))" by simp qed moreover have "integrable lborel (\x. indicat_real {0..t} x *\<^sub>R \sinc x\)" by (auto intro!: borel_integrable_compact continuous_intros simp del: real_scaleR_def) ultimately show "integrable lborel (\x. LBINT y. norm (?f (x, y)))" by (rule integrable_cong_AE_imp[rotated 2]) simp have "0 < x \ set_integrable lborel {0<..} (\y. sin x * exp (- (y * x)))" for x :: real by (intro set_integrable_mult_right integrable_I0i_exp_mscale) then show "AE x in lborel. integrable lborel (\y. ?f (x, y))" by (intro AE_I2) (auto simp: indicator_times set_integrable_def split: split_indicator) qed also have "... = LBINT u=0..\. (LBINT x=0..t. exp (-(u * x)) * sin x)" using \0\t\ by (auto simp: interval_lebesgue_integral_def set_lebesgue_integral_def zero_ereal_def ac_simps split: split_indicator intro!: Bochner_Integration.integral_cong) also have "\ = LBINT u=0..\. 1 / (1 + u\<^sup>2) - 1 / (1 + u\<^sup>2) * (exp (- (u * t)) * (u * sin t + cos t))" by (auto simp: divide_simps LBINT_I0c_exp_mscale_sin intro!: interval_integral_cong) also have "... = pi / 2 - (LBINT u=0..\. exp (- (u * t)) * (u * sin t + cos t) / (1 + u^2))" using Si_at_top_integrable[OF \0\t\] by (simp add: integrable_I0i_1_div_plus_square LBINT_I0i_1_div_plus_square) finally show "pi / 2 - (LBINT u=0..\. exp (-(u * t)) * (u * sin t + cos t) / (1 + u^2)) = Si t" .. qed show ?thesis by (rule Lim_transform_eventually [OF _ \]) (auto intro!: tendsto_eq_intros Si_at_top_LBINT) qed subsection \The final theorems: boundedness and scalability\ lemma bounded_Si: "\B. \T. \Si T\ \ B" proof - have *: "0 \ y \ dist x y < z \ abs x \ y + z" for x y z :: real by (simp add: dist_real_def) have "eventually (\T. dist (Si T) (pi / 2) < 1) at_top" using Si_at_top by (elim tendstoD) simp then have "eventually (\T. 0 \ T \ \Si T\ \ pi / 2 + 1) at_top" using eventually_ge_at_top[of 0] by eventually_elim (insert *[of "pi/2" "Si _" 1], auto) then have "\T. 0 \ T \ (\t \ T. \Si t\ \ pi / 2 + 1)" by (auto simp add: eventually_at_top_linorder) then obtain T where T: "0 \ T" "\t. t \ T \ \Si t\ \ pi / 2 + 1" by auto moreover have "t \ - T \ \Si t\ \ pi / 2 + 1" for t using T(1) T(2)[of "-t"] Si_neg[of "- t"] by simp moreover have "\M. \t. -T \ t \ t \ T \ \Si t\ \ M" by (rule isCont_bounded) (auto intro!: isCont_Si continuous_intros \0 \ T\) then obtain M where M: "\t. -T \ t \ t \ T \ \Si t\ \ M" by auto ultimately show ?thesis by (intro exI[of _ "max M (pi/2 + 1)"]) (meson le_max_iff_disj linorder_not_le order_le_less) qed lemma LBINT_I0c_sin_scale_divide: assumes "T \ 0" shows "LBINT t=0..T. sin (t * \) / t = sgn \ * Si (T * \\\)" proof - { assume "0 < \" have "(LBINT t=ereal 0..T. sin (t * \) / t) = (LBINT t=ereal 0..T. \ *\<^sub>R sinc (t * \))" by (rule interval_integral_discrete_difference[of "{0}"]) auto also have "\ = (LBINT t=ereal (0 * \)..T * \. sinc t)" apply (rule interval_integral_substitution_finite [OF assms]) apply (subst mult.commute, rule DERIV_subset) by (auto intro!: derivative_intros continuous_at_imp_continuous_on isCont_sinc) also have "\ = (LBINT t=ereal (0 * \)..T * \. sin t / t)" by (rule interval_integral_discrete_difference[of "{0}"]) auto finally have "(LBINT t=ereal 0..T. sin (t * \) / t) = (LBINT t=ereal 0..T * \. sin t / t)" by simp hence "LBINT x. indicator {0<..) / x = LBINT x. indicator {0<..} x * sin x / x" using assms \0 < \\ unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def by (auto simp: ac_simps set_lebesgue_integral_def) } note aux1 = this { assume "0 > \" have [simp]: "integrable lborel (\x. sin (x * \) * indicator {0<..] assms by (simp add: interval_lebesgue_integrable_def set_integrable_def ac_simps) have "(LBINT t=ereal 0..T. sin (t * -\) / t) = (LBINT t=ereal 0..T. -\ *\<^sub>R sinc (t * -\))" by (rule interval_integral_discrete_difference[of "{0}"]) auto also have "\ = (LBINT t=ereal (0 * -\)..T * -\. sinc t)" apply (rule interval_integral_substitution_finite [OF assms]) apply (subst mult.commute, rule DERIV_subset) by (auto intro!: derivative_intros continuous_at_imp_continuous_on isCont_sinc) also have "\ = (LBINT t=ereal (0 * -\)..T * -\. sin t / t)" by (rule interval_integral_discrete_difference[of "{0}"]) auto finally have "(LBINT t=ereal 0..T. sin (t * -\) / t) = (LBINT t=ereal 0..T * -\. sin t / t)" by simp hence "LBINT x. indicator {0<..) / x = - (LBINT x. indicator {0<..<- (T * \)} x * sin x / x)" using assms \0 > \\ unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def by (auto simp add: field_simps mult_le_0_iff set_lebesgue_integral_def split: if_split_asm) } note aux2 = this show ?thesis using assms unfolding Si_def interval_lebesgue_integral_def set_lebesgue_integral_def sgn_real_def einterval_eq zero_ereal_def by (auto simp: aux1 aux2) qed end diff --git a/src/HOL/Topological_Spaces.thy b/src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy +++ b/src/HOL/Topological_Spaces.thy @@ -1,3814 +1,3819 @@ (* Title: HOL/Topological_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Topological Spaces\ theory Topological_Spaces imports Main begin named_theorems continuous_intros "structural introduction rules for continuity" subsection \Topological space\ class "open" = fixes "open" :: "'a set \ bool" class topological_space = "open" + assumes open_UNIV [simp, intro]: "open UNIV" assumes open_Int [intro]: "open S \ open T \ open (S \ T)" assumes open_Union [intro]: "\S\K. open S \ open (\K)" begin definition closed :: "'a set \ bool" where "closed S \ open (- S)" lemma open_empty [continuous_intros, intro, simp]: "open {}" using open_Union [of "{}"] by simp lemma open_Un [continuous_intros, intro]: "open S \ open T \ open (S \ T)" using open_Union [of "{S, T}"] by simp lemma open_UN [continuous_intros, intro]: "\x\A. open (B x) \ open (\x\A. B x)" using open_Union [of "B ` A"] by simp lemma open_Inter [continuous_intros, intro]: "finite S \ \T\S. open T \ open (\S)" by (induction set: finite) auto lemma open_INT [continuous_intros, intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" using open_Inter [of "B ` A"] by simp lemma openI: assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" shows "open S" proof - have "open (\{T. open T \ T \ S})" by auto moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) ultimately show "open S" by simp qed lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" by (auto intro: openI) lemma closed_empty [continuous_intros, intro, simp]: "closed {}" unfolding closed_def by simp lemma closed_Un [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_UNIV [continuous_intros, intro, simp]: "closed UNIV" unfolding closed_def by simp lemma closed_Int [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_INT [continuous_intros, intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" unfolding closed_def by auto lemma closed_Inter [continuous_intros, intro]: "\S\K. closed S \ closed (\K)" unfolding closed_def uminus_Inf by auto lemma closed_Union [continuous_intros, intro]: "finite S \ \T\S. closed T \ closed (\S)" by (induct set: finite) auto lemma closed_UN [continuous_intros, intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" using closed_Union [of "B ` A"] by simp lemma open_closed: "open S \ closed (- S)" by (simp add: closed_def) lemma closed_open: "closed S \ open (- S)" by (rule closed_def) lemma open_Diff [continuous_intros, intro]: "open S \ closed T \ open (S - T)" by (simp add: closed_open Diff_eq open_Int) lemma closed_Diff [continuous_intros, intro]: "closed S \ open T \ closed (S - T)" by (simp add: open_closed Diff_eq closed_Int) lemma open_Compl [continuous_intros, intro]: "closed S \ open (- S)" by (simp add: closed_open) lemma closed_Compl [continuous_intros, intro]: "open S \ closed (- S)" by (simp add: open_closed) lemma open_Collect_neg: "closed {x. P x} \ open {x. \ P x}" unfolding Collect_neg_eq by (rule open_Compl) lemma open_Collect_conj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Int[OF assms] by (simp add: Int_def) lemma open_Collect_disj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Un[OF assms] by (simp add: Un_def) lemma open_Collect_ex: "(\i. open {x. P i x}) \ open {x. \i. P i x}" using open_UN[of UNIV "\i. {x. P i x}"] unfolding Collect_ex_eq by simp lemma open_Collect_imp: "closed {x. P x} \ open {x. Q x} \ open {x. P x \ Q x}" unfolding imp_conv_disj by (intro open_Collect_disj open_Collect_neg) lemma open_Collect_const: "open {x. P}" by (cases P) auto lemma closed_Collect_neg: "open {x. P x} \ closed {x. \ P x}" unfolding Collect_neg_eq by (rule closed_Compl) lemma closed_Collect_conj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Int[OF assms] by (simp add: Int_def) lemma closed_Collect_disj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Un[OF assms] by (simp add: Un_def) lemma closed_Collect_all: "(\i. closed {x. P i x}) \ closed {x. \i. P i x}" using closed_INT[of UNIV "\i. {x. P i x}"] by (simp add: Collect_all_eq) lemma closed_Collect_imp: "open {x. P x} \ closed {x. Q x} \ closed {x. P x \ Q x}" unfolding imp_conv_disj by (intro closed_Collect_disj closed_Collect_neg) lemma closed_Collect_const: "closed {x. P}" by (cases P) auto end subsection \Hausdorff and other separation properties\ class t0_space = topological_space + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" class t1_space = topological_space + assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" instance t1_space \ t0_space by standard (fast dest: t1_space) context t1_space begin lemma separation_t1: "x \ y \ (\U. open U \ x \ U \ y \ U)" using t1_space[of x y] by blast lemma closed_singleton [iff]: "closed {a}" proof - let ?T = "\{S. open S \ a \ S}" have "open ?T" by (simp add: open_Union) also have "?T = - {a}" by (auto simp add: set_eq_iff separation_t1) finally show "closed {a}" by (simp only: closed_def) qed lemma closed_insert [continuous_intros, simp]: assumes "closed S" shows "closed (insert a S)" proof - from closed_singleton assms have "closed ({a} \ S)" by (rule closed_Un) then show "closed (insert a S)" by simp qed lemma finite_imp_closed: "finite S \ closed S" by (induct pred: finite) simp_all end text \T2 spaces are also known as Hausdorff spaces.\ class t2_space = topological_space + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" instance t2_space \ t1_space by standard (fast dest: hausdorff) lemma (in t2_space) separation_t2: "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" using hausdorff [of x y] by blast lemma (in t0_space) separation_t0: "x \ y \ (\U. open U \ \ (x \ U \ y \ U))" using t0_space [of x y] by blast text \A classical separation axiom for topological space, the T3 axiom -- also called regularity: if a point is not in a closed set, then there are open sets separating them.\ class t3_space = t2_space + assumes t3_space: "closed S \ y \ S \ \U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" text \A classical separation axiom for topological space, the T4 axiom -- also called normality: if two closed sets are disjoint, then there are open sets separating them.\ class t4_space = t2_space + assumes t4_space: "closed S \ closed T \ S \ T = {} \ \U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" text \T4 is stronger than T3, and weaker than metric.\ instance t4_space \ t3_space proof fix S and y::'a assume "closed S" "y \ S" then show "\U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" using t4_space[of "{y}" S] by auto qed text \A perfect space is a topological space with no isolated points.\ class perfect_space = topological_space + assumes not_open_singleton: "\ open {x}" lemma (in perfect_space) UNIV_not_singleton: "UNIV \ {x}" for x::'a by (metis (no_types) open_UNIV not_open_singleton) subsection \Generators for toplogies\ inductive generate_topology :: "'a set set \ 'a set \ bool" for S :: "'a set set" where UNIV: "generate_topology S UNIV" | Int: "generate_topology S (a \ b)" if "generate_topology S a" and "generate_topology S b" | UN: "generate_topology S (\K)" if "(\k. k \ K \ generate_topology S k)" | Basis: "generate_topology S s" if "s \ S" hide_fact (open) UNIV Int UN Basis lemma generate_topology_Union: "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" using generate_topology.UN [of "K ` I"] by auto lemma topological_space_generate_topology: "class.topological_space (generate_topology S)" by standard (auto intro: generate_topology.intros) subsection \Order topologies\ class order_topology = order + "open" + assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" begin subclass topological_space unfolding open_generated_order by (rule topological_space_generate_topology) lemma open_greaterThan [continuous_intros, simp]: "open {a <..}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_lessThan [continuous_intros, simp]: "open {..< a}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_greaterThanLessThan [continuous_intros, simp]: "open {a <..< b}" unfolding greaterThanLessThan_eq by (simp add: open_Int) end class linorder_topology = linorder + order_topology lemma closed_atMost [continuous_intros, simp]: "closed {..a}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeast [continuous_intros, simp]: "closed {a..}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeastAtMost [continuous_intros, simp]: "closed {a..b}" for a b :: "'a::linorder_topology" proof - have "{a .. b} = {a ..} \ {.. b}" by auto then show ?thesis by (simp add: closed_Int) qed lemma (in order) less_separate: assumes "x < y" shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" proof (cases "\z. x < z \ z < y") case True then obtain z where "x < z \ z < y" .. then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" by auto then show ?thesis by blast next case False with \x < y\ have "x \ {..< y}" "y \ {x <..}" "{x <..} \ {..< y} = {}" by auto then show ?thesis by blast qed instance linorder_topology \ t2_space proof fix x y :: 'a show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" using less_separate [of x y] less_separate [of y x] by (elim neqE; metis open_lessThan open_greaterThan Int_commute) qed lemma (in linorder_topology) open_right: assumes "open S" "x \ S" and gt_ex: "x < y" shows "\b>x. {x ..< b} \ S" using assms unfolding open_generated_order proof induct case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a > x" "{x ..< a} \ A" "b > x" "{x ..< b} \ B" by auto then show ?case by (auto intro!: exI[of _ "min a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] gt_ex) qed lemma (in linorder_topology) open_left: assumes "open S" "x \ S" and lt_ex: "y < x" shows "\b S" using assms unfolding open_generated_order proof induction case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a < x" "{a <.. x} \ A" "b < x" "{b <.. x} \ B" by auto then show ?case by (auto intro!: exI[of _ "max a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] lt_ex) qed subsection \Setup some topologies\ subsubsection \Boolean is an order topology\ class discrete_topology = topological_space + assumes open_discrete: "\A. open A" instance discrete_topology < t2_space proof fix x y :: 'a assume "x \ y" then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (intro exI[of _ "{_}"]) (auto intro!: open_discrete) qed instantiation bool :: linorder_topology begin definition open_bool :: "bool set \ bool" where "open_bool = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_bool_def) end instance bool :: discrete_topology proof fix A :: "bool set" have *: "{False <..} = {True}" "{..< True} = {False}" by auto have "A = UNIV \ A = {} \ A = {False <..} \ A = {..< True}" using subset_UNIV[of A] unfolding UNIV_bool * by blast then show "open A" by auto qed instantiation nat :: linorder_topology begin definition open_nat :: "nat set \ bool" where "open_nat = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_nat_def) end instance nat :: discrete_topology proof fix A :: "nat set" have "open {n}" for n :: nat proof (cases n) case 0 moreover have "{0} = {..<1::nat}" by auto ultimately show ?thesis by auto next case (Suc n') then have "{n} = {.. {n' <..}" by auto with Suc show ?thesis by (auto intro: open_lessThan open_greaterThan) qed then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed instantiation int :: linorder_topology begin definition open_int :: "int set \ bool" where "open_int = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_int_def) end instance int :: discrete_topology proof fix A :: "int set" have "{.. {i-1 <..} = {i}" for i :: int by auto then have "open {i}" for i :: int using open_Int[OF open_lessThan[of "i + 1"] open_greaterThan[of "i - 1"]] by auto then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed subsubsection \Topological filters\ definition (in topological_space) nhds :: "'a \ 'a filter" where "nhds a = (INF S\{S. open S \ a \ S}. principal S)" definition (in topological_space) at_within :: "'a \ 'a set \ 'a filter" ("at (_)/ within (_)" [1000, 60] 60) where "at a within s = inf (nhds a) (principal (s - {a}))" abbreviation (in topological_space) at :: "'a \ 'a filter" ("at") where "at x \ at x within (CONST UNIV)" abbreviation (in order_topology) at_right :: "'a \ 'a filter" where "at_right x \ at x within {x <..}" abbreviation (in order_topology) at_left :: "'a \ 'a filter" where "at_left x \ at x within {..< x}" lemma (in topological_space) nhds_generated_topology: "open = generate_topology T \ nhds x = (INF S\{S\T. x \ S}. principal S)" unfolding nhds_def proof (safe intro!: antisym INF_greatest) fix S assume "generate_topology T S" "x \ S" then show "(INF S\{S \ T. x \ S}. principal S) \ principal S" by induct (auto intro: INF_lower order_trans simp: inf_principal[symmetric] simp del: inf_principal) qed (auto intro!: INF_lower intro: generate_topology.intros) lemma (in topological_space) eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def by (subst eventually_INF_base) (auto simp: eventually_principal) lemma eventually_eventually: "eventually (\y. eventually P (nhds y)) (nhds x) = eventually P (nhds x)" by (auto simp: eventually_nhds) lemma (in topological_space) eventually_nhds_in_open: "open s \ x \ s \ eventually (\y. y \ s) (nhds x)" by (subst eventually_nhds) blast lemma (in topological_space) eventually_nhds_x_imp_x: "eventually P (nhds x) \ P x" by (subst (asm) eventually_nhds) blast lemma (in topological_space) nhds_neq_bot [simp]: "nhds a \ bot" by (simp add: trivial_limit_def eventually_nhds) lemma (in t1_space) t1_space_nhds: "x \ y \ (\\<^sub>F x in nhds x. x \ y)" by (drule t1_space) (auto simp: eventually_nhds) lemma (in topological_space) nhds_discrete_open: "open {x} \ nhds x = principal {x}" by (auto simp: nhds_def intro!: antisym INF_greatest INF_lower2[of "{x}"]) lemma (in discrete_topology) nhds_discrete: "nhds x = principal {x}" by (simp add: nhds_discrete_open open_discrete) lemma (in discrete_topology) at_discrete: "at x within S = bot" unfolding at_within_def nhds_discrete by simp lemma (in discrete_topology) tendsto_discrete: "filterlim (f :: 'b \ 'a) (nhds y) F \ eventually (\x. f x = y) F" by (auto simp: nhds_discrete filterlim_principal) lemma (in topological_space) at_within_eq: "at x within s = (INF S\{S. open S \ x \ S}. principal (S \ s - {x}))" unfolding nhds_def at_within_def by (subst INF_inf_const2[symmetric]) (auto simp: Diff_Int_distrib) lemma (in topological_space) eventually_at_filter: "eventually P (at a within s) \ eventually (\x. x \ a \ x \ s \ P x) (nhds a)" by (simp add: at_within_def eventually_inf_principal imp_conjL[symmetric] conj_commute) lemma (in topological_space) at_le: "s \ t \ at x within s \ at x within t" unfolding at_within_def by (intro inf_mono) auto lemma (in topological_space) eventually_at_topological: "eventually P (at a within s) \ (\S. open S \ a \ S \ (\x\S. x \ a \ x \ s \ P x))" by (simp add: eventually_nhds eventually_at_filter) lemma (in topological_space) at_within_open: "a \ S \ open S \ at a within S = at a" unfolding filter_eq_iff eventually_at_topological by (metis open_Int Int_iff UNIV_I) lemma (in topological_space) at_within_open_NO_MATCH: "a \ s \ open s \ NO_MATCH UNIV s \ at a within s = at a" by (simp only: at_within_open) lemma (in topological_space) at_within_open_subset: "a \ S \ open S \ S \ T \ at a within T = at a" by (metis at_le at_within_open dual_order.antisym subset_UNIV) lemma (in topological_space) at_within_nhd: assumes "x \ S" "open S" "T \ S - {x} = U \ S - {x}" shows "at x within T = at x within U" unfolding filter_eq_iff eventually_at_filter proof (intro allI eventually_subst) have "eventually (\x. x \ S) (nhds x)" using \x \ S\ \open S\ by (auto simp: eventually_nhds) then show "\\<^sub>F n in nhds x. (n \ x \ n \ T \ P n) = (n \ x \ n \ U \ P n)" for P by eventually_elim (insert \T \ S - {x} = U \ S - {x}\, blast) qed lemma (in topological_space) at_within_empty [simp]: "at a within {} = bot" unfolding at_within_def by simp lemma (in topological_space) at_within_union: "at x within (S \ T) = sup (at x within S) (at x within T)" unfolding filter_eq_iff eventually_sup eventually_at_filter by (auto elim!: eventually_rev_mp) lemma (in topological_space) at_eq_bot_iff: "at a = bot \ open {a}" unfolding trivial_limit_def eventually_at_topological by (metis UNIV_I empty_iff is_singletonE is_singletonI' singleton_iff) lemma (in perfect_space) at_neq_bot [simp]: "at a \ bot" by (simp add: at_eq_bot_iff not_open_singleton) lemma (in order_topology) nhds_order: "nhds x = inf (INF a\{x <..}. principal {..< a}) (INF a\{..< x}. principal {a <..})" proof - have 1: "{S \ range lessThan \ range greaterThan. x \ S} = (\a. {..< a}) ` {x <..} \ (\a. {a <..}) ` {..< x}" by auto show ?thesis by (simp only: nhds_generated_topology[OF open_generated_order] INF_union 1 INF_image comp_def) qed lemma (in topological_space) filterlim_at_within_If: assumes "filterlim f G (at x within (A \ {x. P x}))" and "filterlim g G (at x within (A \ {x. \P x}))" shows "filterlim (\x. if P x then f x else g x) G (at x within A)" proof (rule filterlim_If) note assms(1) also have "at x within (A \ {x. P x}) = inf (nhds x) (principal (A \ Collect P - {x}))" by (simp add: at_within_def) also have "A \ Collect P - {x} = (A - {x}) \ Collect P" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal (Collect P))" by (simp add: at_within_def inf_assoc) finally show "filterlim f G (inf (at x within A) (principal (Collect P)))" . next note assms(2) also have "at x within (A \ {x. \ P x}) = inf (nhds x) (principal (A \ {x. \ P x} - {x}))" by (simp add: at_within_def) also have "A \ {x. \ P x} - {x} = (A - {x}) \ {x. \ P x}" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal {x. \ P x})" by (simp add: at_within_def inf_assoc) finally show "filterlim g G (inf (at x within A) (principal {x. \ P x}))" . qed lemma (in topological_space) filterlim_at_If: assumes "filterlim f G (at x within {x. P x})" and "filterlim g G (at x within {x. \P x})" shows "filterlim (\x. if P x then f x else g x) G (at x)" using assms by (intro filterlim_at_within_If) simp_all lemma (in linorder_topology) at_within_order: assumes "UNIV \ {x}" shows "at x within s = inf (INF a\{x <..}. principal ({..< a} \ s - {x})) (INF a\{..< x}. principal ({a <..} \ s - {x}))" proof (cases "{x <..} = {}" "{..< x} = {}" rule: case_split [case_product case_split]) case True_True have "UNIV = {..< x} \ {x} \ {x <..}" by auto with assms True_True show ?thesis by auto qed (auto simp del: inf_principal simp: at_within_def nhds_order Int_Diff inf_principal[symmetric] INF_inf_const2 inf_sup_aci[where 'a="'a filter"]) lemma (in linorder_topology) at_left_eq: "y < x \ at_left x = (INF a\{..< x}. principal {a <..< x})" by (subst at_within_order) (auto simp: greaterThan_Int_greaterThan greaterThanLessThan_eq[symmetric] min.absorb2 INF_constant intro!: INF_lower2 inf_absorb2) lemma (in linorder_topology) eventually_at_left: "y < x \ eventually P (at_left x) \ (\by>b. y < x \ P y)" unfolding at_left_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma (in linorder_topology) at_right_eq: "x < y \ at_right x = (INF a\{x <..}. principal {x <..< a})" by (subst at_within_order) (auto simp: lessThan_Int_lessThan greaterThanLessThan_eq[symmetric] max.absorb2 INF_constant Int_commute intro!: INF_lower2 inf_absorb1) lemma (in linorder_topology) eventually_at_right: "x < y \ eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" unfolding at_right_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma eventually_at_right_less: "\\<^sub>F y in at_right (x::'a::{linorder_topology, no_top}). x < y" using gt_ex[of x] eventually_at_right[of x] by auto lemma trivial_limit_at_right_top: "at_right (top::_::{order_top,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_bot: "at_left (bot::_::{order_bot,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_real [simp]: "\ trivial_limit (at_left x)" for x :: "'a::{no_bot,dense_order,linorder_topology}" using lt_ex [of x] by safe (auto simp add: trivial_limit_def eventually_at_left dest: dense) lemma trivial_limit_at_right_real [simp]: "\ trivial_limit (at_right x)" for x :: "'a::{no_top,dense_order,linorder_topology}" using gt_ex[of x] by safe (auto simp add: trivial_limit_def eventually_at_right dest: dense) lemma (in linorder_topology) at_eq_sup_left_right: "at x = sup (at_left x) (at_right x)" by (auto simp: eventually_at_filter filter_eq_iff eventually_sup elim: eventually_elim2 eventually_mono) lemma (in linorder_topology) eventually_at_split: "eventually P (at x) \ eventually P (at_left x) \ eventually P (at_right x)" by (subst at_eq_sup_left_right) (simp add: eventually_sup) lemma (in order_topology) eventually_at_leftI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_left b)" using assms unfolding eventually_at_topological by (intro exI[of _ "{a<..}"]) auto lemma (in order_topology) eventually_at_rightI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_right a)" using assms unfolding eventually_at_topological by (intro exI[of _ "{.. (\S. open S \ x \ S \ (\x. f x \ S \ P x))" unfolding eventually_filtercomap eventually_nhds by auto lemma eventually_filtercomap_at_topological: "eventually P (filtercomap f (at A within B)) \ (\S. open S \ A \ S \ (\x. f x \ S \ B - {A} \ P x))" (is "?lhs = ?rhs") unfolding at_within_def filtercomap_inf eventually_inf_principal filtercomap_principal eventually_filtercomap_nhds eventually_principal by blast lemma eventually_at_right_field: "eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_ub[rule_format, of x] by (auto simp: eventually_at_right) lemma eventually_at_left_field: "eventually P (at_left x) \ (\by>b. y < x \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_lb[rule_format, of x] by (auto simp: eventually_at_left) subsubsection \Tendsto\ abbreviation (in topological_space) tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "\" 55) where "(f \ l) F \ filterlim f (nhds l) F" definition (in t2_space) Lim :: "'f filter \ ('f \ 'a) \ 'a" where "Lim A f = (THE l. (f \ l) A)" lemma (in topological_space) tendsto_eq_rhs: "(f \ x) F \ x = y \ (f \ y) F" by simp named_theorems tendsto_intros "introduction rules for tendsto" setup \ Global_Theory.add_thms_dynamic (\<^binding>\tendsto_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\tendsto_intros\ |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm]))) \ context topological_space begin lemma tendsto_def: "(f \ l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" unfolding nhds_def filterlim_INF filterlim_principal by auto lemma tendsto_cong: "(f \ c) F \ (g \ c) F" if "eventually (\x. f x = g x) F" by (rule filterlim_cong [OF refl refl that]) lemma tendsto_mono: "F \ F' \ (f \ l) F' \ (f \ l) F" unfolding tendsto_def le_filter_def by fast lemma tendsto_ident_at [tendsto_intros, simp, intro]: "((\x. x) \ a) (at a within s)" by (auto simp: tendsto_def eventually_at_topological) lemma tendsto_const [tendsto_intros, simp, intro]: "((\x. k) \ k) F" by (simp add: tendsto_def) lemma filterlim_at: "(LIM x F. f x :> at b within s) \ eventually (\x. f x \ s \ f x \ b) F \ (f \ b) F" by (simp add: at_within_def filterlim_inf filterlim_principal conj_commute) lemma (in -) assumes "filterlim f (nhds L) F" shows tendsto_imp_filterlim_at_right: "eventually (\x. f x > L) F \ filterlim f (at_right L) F" and tendsto_imp_filterlim_at_left: "eventually (\x. f x < L) F \ filterlim f (at_left L) F" using assms by (auto simp: filterlim_at elim: eventually_mono) lemma filterlim_at_withinI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ A - {c}) F" shows "filterlim f (at c within A) F" using assms by (simp add: filterlim_at) lemma filterlim_atI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ c) F" shows "filterlim f (at c) F" using assms by (intro filterlim_at_withinI) simp_all lemma topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) \ (f \ l) F" by (auto simp: tendsto_def) lemma topological_tendstoD: "(f \ l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" by (auto simp: tendsto_def) lemma tendsto_bot [simp]: "(f \ a) bot" by (simp add: tendsto_def) lemma tendsto_eventually: "eventually (\x. f x = l) net \ ((\x. f x) \ l) net" by (rule topological_tendstoI) (auto elim: eventually_mono) (* Contributed by Dominique Unruh *) lemma tendsto_principal_singleton[simp]: shows "(f \ f x) (principal {x})" unfolding tendsto_def eventually_principal by simp end lemma (in topological_space) filterlim_within_subset: "filterlim f l (at x within S) \ T \ S \ filterlim f l (at x within T)" by (blast intro: filterlim_mono at_le) lemmas tendsto_within_subset = filterlim_within_subset lemma (in order_topology) order_tendsto_iff: "(f \ x) F \ (\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" by (auto simp: nhds_order filterlim_inf filterlim_INF filterlim_principal) lemma (in order_topology) order_tendstoI: "(\a. a < y \ eventually (\x. a < f x) F) \ (\a. y < a \ eventually (\x. f x < a) F) \ (f \ y) F" by (auto simp: order_tendsto_iff) lemma (in order_topology) order_tendstoD: assumes "(f \ y) F" shows "a < y \ eventually (\x. a < f x) F" and "y < a \ eventually (\x. f x < a) F" using assms by (auto simp: order_tendsto_iff) lemma (in linorder_topology) tendsto_max[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. max (X x) (Y x)) \ max x y) net" proof (rule order_tendstoI) fix a assume "a < max x y" then show "eventually (\x. a < max (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: less_max_iff_disj elim: eventually_mono) next fix a assume "max x y < a" then show "eventually (\x. max (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: eventually_conj_iff) qed lemma (in linorder_topology) tendsto_min[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. min (X x) (Y x)) \ min x y) net" proof (rule order_tendstoI) fix a assume "a < min x y" then show "eventually (\x. a < min (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: eventually_conj_iff) next fix a assume "min x y < a" then show "eventually (\x. min (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: min_less_iff_disj elim: eventually_mono) qed lemma (in order_topology) assumes "a < b" shows at_within_Icc_at_right: "at a within {a..b} = at_right a" and at_within_Icc_at_left: "at b within {a..b} = at_left b" using order_tendstoD(2)[OF tendsto_ident_at assms, of "{a<..}"] using order_tendstoD(1)[OF tendsto_ident_at assms, of "{.. x < b \ at x within {a..b} = at x" by (rule at_within_open_subset[where S="{a<.. bot" and "(f \ a) F" and "(f \ b) F" shows "a = b" proof (rule ccontr) assume "a \ b" obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" using hausdorff [OF \a \ b\] by fast have "eventually (\x. f x \ U) F" using \(f \ a) F\ \open U\ \a \ U\ by (rule topological_tendstoD) moreover have "eventually (\x. f x \ V) F" using \(f \ b) F\ \open V\ \b \ V\ by (rule topological_tendstoD) ultimately have "eventually (\x. False) F" proof eventually_elim case (elim x) then have "f x \ U \ V" by simp with \U \ V = {}\ show ?case by simp qed with \\ trivial_limit F\ show "False" by (simp add: trivial_limit_def) qed lemma (in t2_space) tendsto_const_iff: fixes a b :: 'a assumes "\ trivial_limit F" shows "((\x. a) \ b) F \ a = b" by (auto intro!: tendsto_unique [OF assms tendsto_const]) lemma (in t2_space) tendsto_unique': assumes "F \ bot" shows "\\<^sub>\\<^sub>1l. (f \ l) F" using Uniq_def assms local.tendsto_unique by fastforce lemma Lim_in_closed_set: assumes "closed S" "eventually (\x. f(x) \ S) F" "F \ bot" "(f \ l) F" shows "l \ S" proof (rule ccontr) assume "l \ S" with \closed S\ have "open (- S)" "l \ - S" by (simp_all add: open_Compl) with assms(4) have "eventually (\x. f x \ - S) F" by (rule topological_tendstoD) with assms(2) have "eventually (\x. False) F" by (rule eventually_elim2) simp with assms(3) show "False" by (simp add: eventually_False) qed lemma (in t3_space) nhds_closed: assumes "x \ A" and "open A" shows "\A'. x \ A' \ closed A' \ A' \ A \ eventually (\y. y \ A') (nhds x)" proof - from assms have "\U V. open U \ open V \ x \ U \ - A \ V \ U \ V = {}" by (intro t3_space) auto then obtain U V where UV: "open U" "open V" "x \ U" "-A \ V" "U \ V = {}" by auto have "eventually (\y. y \ U) (nhds x)" using \open U\ and \x \ U\ by (intro eventually_nhds_in_open) hence "eventually (\y. y \ -V) (nhds x)" by eventually_elim (use UV in auto) with UV show ?thesis by (intro exI[of _ "-V"]) auto qed lemma (in order_topology) increasing_tendsto: assumes bdd: "eventually (\n. f n \ l) F" and en: "\x. x < l \ eventually (\n. x < f n) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) decreasing_tendsto: assumes bdd: "eventually (\n. l \ f n) F" and en: "\x. l < x \ eventually (\n. f n < x) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) tendsto_sandwich: assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" assumes lim: "(f \ c) net" "(h \ c) net" shows "(g \ c) net" proof (rule order_tendstoI) fix a show "a < c \ eventually (\x. a < g x) net" using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) next fix a show "c < a \ eventually (\x. g x < a) net" using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) qed lemma (in t1_space) limit_frequently_eq: assumes "F \ bot" and "frequently (\x. f x = c) F" and "(f \ d) F" shows "d = c" proof (rule ccontr) assume "d \ c" from t1_space[OF this] obtain U where "open U" "d \ U" "c \ U" by blast with assms have "eventually (\x. f x \ U) F" unfolding tendsto_def by blast then have "eventually (\x. f x \ c) F" by eventually_elim (insert \c \ U\, blast) with assms(2) show False unfolding frequently_def by contradiction qed lemma (in t1_space) tendsto_imp_eventually_ne: assumes "(f \ c) F" "c \ c'" shows "eventually (\z. f z \ c') F" proof (cases "F=bot") case True thus ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ eventually (\z. f z \ c') F" then have "frequently (\z. f z = c') F" by (simp add: frequently_def) from limit_frequently_eq[OF False this \(f \ c) F\] and \c \ c'\ show False by contradiction qed qed lemma (in linorder_topology) tendsto_le: assumes F: "\ trivial_limit F" and x: "(f \ x) F" and y: "(g \ y) F" and ev: "eventually (\x. g x \ f x) F" shows "y \ x" proof (rule ccontr) assume "\ y \ x" with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" by (auto simp: not_le) then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" using x y by (auto intro: order_tendstoD) with ev have "eventually (\x. False) F" by eventually_elim (insert xy, fastforce) with F show False by (simp add: eventually_False) qed lemma (in linorder_topology) tendsto_lowerbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" using F x tendsto_const ev by (rule tendsto_le) lemma (in linorder_topology) tendsto_upperbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" by (rule tendsto_le [OF F tendsto_const x ev]) lemma filterlim_at_within_not_equal: fixes f::"'a \ 'b::t2_space" assumes "filterlim f (at a within s) F" shows "eventually (\w. f w\s \ f w \b) F" proof (cases "a=b") case True then show ?thesis using assms by (simp add: filterlim_at) next case False from hausdorff[OF this] obtain U V where UV:"open U" "open V" "a \ U" "b \ V" "U \ V = {}" by auto have "(f \ a) F" using assms filterlim_at by auto then have "\\<^sub>F x in F. f x \ U" using UV unfolding tendsto_def by auto moreover have "\\<^sub>F x in F. f x \ s \ f x\a" using assms filterlim_at by auto ultimately show ?thesis apply eventually_elim using UV by auto qed subsubsection \Rules about \<^const>\Lim\\ lemma tendsto_Lim: "\ trivial_limit net \ (f \ l) net \ Lim net f = l" unfolding Lim_def using tendsto_unique [of net f] by auto lemma Lim_ident_at: "\ trivial_limit (at x within s) \ Lim (at x within s) (\x. x) = x" by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto lemma eventually_Lim_ident_at: "(\\<^sub>F y in at x within X. P (Lim (at x within X) (\x. x)) y) \ (\\<^sub>F y in at x within X. P x y)" for x::"'a::t2_space" by (cases "at x within X = bot") (auto simp: Lim_ident_at) lemma filterlim_at_bot_at_right: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" and P: "eventually P at_bot" shows "filterlim f at_bot (at_right a)" proof - from P obtain x where x: "\y. y \ x \ P y" unfolding eventually_at_bot_linorder by auto show ?thesis proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) fix z assume "z \ x" with x have "P z" by auto have "eventually (\x. x \ g z) (at_right a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_right[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. f x \ z) (at_right a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_at_top_at_left: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" and P: "eventually P at_top" shows "filterlim f at_top (at_left a)" proof - from P obtain x where x: "\y. x \ y \ P y" unfolding eventually_at_top_linorder by auto show ?thesis proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) fix z assume "x \ z" with x have "P z" by auto have "eventually (\x. g z \ x) (at_left a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_left[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. z \ f x) (at_left a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_split_at: "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (rule filterlim_sup) lemma filterlim_at_split: "filterlim f F (at x) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) lemma eventually_nhds_top: fixes P :: "'a :: {order_top,linorder_topology} \ bool" and b :: 'a assumes "b < top" shows "eventually P (nhds top) \ (\bz. b < z \ P z))" unfolding eventually_nhds proof safe fix S :: "'a set" assume "open S" "top \ S" note open_left[OF this \b < top\] moreover assume "\s\S. P s" ultimately show "\bz>b. P z" by (auto simp: subset_eq Ball_def) next fix b assume "b < top" "\z>b. P z" then show "\S. open S \ top \ S \ (\xa\S. P xa)" by (intro exI[of _ "{b <..}"]) auto qed lemma tendsto_at_within_iff_tendsto_nhds: "(g \ g l) (at l within S) \ (g \ g l) (inf (nhds l) (principal S))" unfolding tendsto_def eventually_at_filter eventually_inf_principal by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) subsection \Limits on sequences\ abbreviation (in topological_space) LIMSEQ :: "[nat \ 'a, 'a] \ bool" ("((_)/ \ (_))" [60, 60] 60) where "X \ L \ (X \ L) sequentially" abbreviation (in t2_space) lim :: "(nat \ 'a) \ 'a" where "lim X \ Lim sequentially X" definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where "convergent X = (\L. X \ L)" lemma lim_def: "lim X = (THE L. X \ L)" unfolding Lim_def .. lemma lim_explicit: "f \ f0 \ (\S. open S \ f0 \ S \ (\N. \n\N. f n \ S))" unfolding tendsto_def eventually_sequentially by auto subsection \Monotone sequences and subsequences\ text \ Definition of monotonicity. The use of disjunction here complicates proofs considerably. One alternative is to add a Boolean argument to indicate the direction. Another is to develop the notions of increasing and decreasing first. \ definition monoseq :: "(nat \ 'a::order) \ bool" where "monoseq X \ (\m. \n\m. X m \ X n) \ (\m. \n\m. X n \ X m)" abbreviation incseq :: "(nat \ 'a::order) \ bool" where "incseq X \ mono X" lemma incseq_def: "incseq X \ (\m. \n\m. X n \ X m)" unfolding mono_def .. abbreviation decseq :: "(nat \ 'a::order) \ bool" where "decseq X \ antimono X" lemma decseq_def: "decseq X \ (\m. \n\m. X n \ X m)" unfolding antimono_def .. subsubsection \Definition of subsequence.\ (* For compatibility with the old "subseq" *) lemma strict_mono_leD: "strict_mono r \ m \ n \ r m \ r n" by (erule (1) monoD [OF strict_mono_mono]) lemma strict_mono_id: "strict_mono id" by (simp add: strict_mono_def) lemma incseq_SucI: "(\n. X n \ X (Suc n)) \ incseq X" using lift_Suc_mono_le[of X] by (auto simp: incseq_def) lemma incseqD: "incseq f \ i \ j \ f i \ f j" by (auto simp: incseq_def) lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" using incseqD[of A i "Suc i"] by auto lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" by (auto intro: incseq_SucI dest: incseq_SucD) lemma incseq_const[simp, intro]: "incseq (\x. k)" unfolding incseq_def by auto lemma decseq_SucI: "(\n. X (Suc n) \ X n) \ decseq X" using order.lift_Suc_mono_le[OF dual_order, of X] by (auto simp: decseq_def) lemma decseqD: "decseq f \ i \ j \ f j \ f i" by (auto simp: decseq_def) lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" using decseqD[of A i "Suc i"] by auto lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" by (auto intro: decseq_SucI dest: decseq_SucD) lemma decseq_const[simp, intro]: "decseq (\x. k)" unfolding decseq_def by auto lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" unfolding monoseq_def incseq_def decseq_def .. lemma monoseq_Suc: "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. lemma monoI1: "\m. \n \ m. X m \ X n \ monoseq X" by (simp add: monoseq_def) lemma monoI2: "\m. \n \ m. X n \ X m \ monoseq X" by (simp add: monoseq_def) lemma mono_SucI1: "\n. X n \ X (Suc n) \ monoseq X" by (simp add: monoseq_Suc) lemma mono_SucI2: "\n. X (Suc n) \ X n \ monoseq X" by (simp add: monoseq_Suc) lemma monoseq_minus: fixes a :: "nat \ 'a::ordered_ab_group_add" assumes "monoseq a" shows "monoseq (\ n. - a n)" proof (cases "\m. \n \ m. a m \ a n") case True then have "\m. \n \ m. - a n \ - a m" by auto then show ?thesis by (rule monoI2) next case False then have "\m. \n \ m. - a m \ - a n" using \monoseq a\[unfolded monoseq_def] by auto then show ?thesis by (rule monoI1) qed subsubsection \Subsequence (alternative definition, (e.g. Hoskins)\ lemma strict_mono_Suc_iff: "strict_mono f \ (\n. f n < f (Suc n))" proof (intro iffI strict_monoI) assume *: "\n. f n < f (Suc n)" fix m n :: nat assume "m < n" thus "f m < f n" by (induction rule: less_Suc_induct) (use * in auto) qed (auto simp: strict_mono_def) lemma strict_mono_add: "strict_mono (\n::'a::linordered_semidom. n + k)" by (auto simp: strict_mono_def) text \For any sequence, there is a monotonic subsequence.\ lemma seq_monosub: fixes s :: "nat \ 'a::linorder" shows "\f. strict_mono f \ monoseq (\n. (s (f n)))" proof (cases "\n. \p>n. \m\p. s m \ s p") case True then have "\f. \n. (\m\f n. s m \ s (f n)) \ f n < f (Suc n)" by (intro dependent_nat_choice) (auto simp: conj_commute) then obtain f :: "nat \ nat" where f: "strict_mono f" and mono: "\n m. f n \ m \ s m \ s (f n)" by (auto simp: strict_mono_Suc_iff) then have "incseq f" unfolding strict_mono_Suc_iff incseq_Suc_iff by (auto intro: less_imp_le) then have "monoseq (\n. s (f n))" by (auto simp add: incseq_def intro!: mono monoI2) with f show ?thesis by auto next case False then obtain N where N: "p > N \ \m>p. s p < s m" for p by (force simp: not_le le_less) have "\f. \n. N < f n \ f n < f (Suc n) \ s (f n) \ s (f (Suc n))" proof (intro dependent_nat_choice) fix x assume "N < x" with N[of x] show "\y>N. x < y \ s x \ s y" by (auto intro: less_trans) qed auto then show ?thesis by (auto simp: monoseq_iff incseq_Suc_iff strict_mono_Suc_iff) qed lemma seq_suble: assumes sf: "strict_mono (f :: nat \ nat)" shows "n \ f n" proof (induct n) case 0 show ?case by simp next case (Suc n) with sf [unfolded strict_mono_Suc_iff, rule_format, of n] have "n < f (Suc n)" by arith then show ?case by arith qed lemma eventually_subseq: "strict_mono r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" unfolding eventually_sequentially by (metis seq_suble le_trans) lemma not_eventually_sequentiallyD: assumes "\ eventually P sequentially" shows "\r::nat\nat. strict_mono r \ (\n. \ P (r n))" proof - from assms have "\n. \m\n. \ P m" unfolding eventually_sequentially by (simp add: not_less) then obtain r where "\n. r n \ n" "\n. \ P (r n)" by (auto simp: choice_iff) then show ?thesis by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] simp: less_eq_Suc_le strict_mono_Suc_iff) qed lemma sequentially_offset: assumes "eventually (\i. P i) sequentially" shows "eventually (\i. P (i + k)) sequentially" using assms by (rule eventually_sequentially_seg [THEN iffD2]) lemma seq_offset_neg: "(f \ l) sequentially \ ((\i. f(i - k)) \ l) sequentially" apply (erule filterlim_compose) apply (simp add: filterlim_def le_sequentially eventually_filtermap eventually_sequentially, arith) done lemma filterlim_subseq: "strict_mono f \ filterlim f sequentially sequentially" unfolding filterlim_iff by (metis eventually_subseq) lemma strict_mono_o: "strict_mono r \ strict_mono s \ strict_mono (r \ s)" unfolding strict_mono_def by simp lemma strict_mono_compose: "strict_mono r \ strict_mono s \ strict_mono (\x. r (s x))" using strict_mono_o[of r s] by (simp add: o_def) lemma incseq_imp_monoseq: "incseq X \ monoseq X" by (simp add: incseq_def monoseq_def) lemma decseq_imp_monoseq: "decseq X \ monoseq X" by (simp add: decseq_def monoseq_def) lemma decseq_eq_incseq: "decseq X = incseq (\n. - X n)" for X :: "nat \ 'a::ordered_ab_group_add" by (simp add: decseq_def incseq_def) lemma INT_decseq_offset: assumes "decseq F" shows "(\i. F i) = (\i\{n..}. F i)" proof safe fix x i assume x: "x \ (\i\{n..}. F i)" show "x \ F i" proof cases from x have "x \ F n" by auto also assume "i \ n" with \decseq F\ have "F n \ F i" unfolding decseq_def by simp finally show ?thesis . qed (insert x, simp) qed auto lemma LIMSEQ_const_iff: "(\n. k) \ l \ k = l" for k l :: "'a::t2_space" using trivial_limit_sequentially by (rule tendsto_const_iff) lemma LIMSEQ_SUP: "incseq X \ X \ (SUP i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro increasing_tendsto) (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) lemma LIMSEQ_INF: "decseq X \ X \ (INF i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro decreasing_tendsto) (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) lemma LIMSEQ_ignore_initial_segment: "f \ a \ (\n. f (n + k)) \ a" unfolding tendsto_def by (subst eventually_sequentially_seg[where k=k]) lemma LIMSEQ_offset: "(\n. f (n + k)) \ a \ f \ a" unfolding tendsto_def by (subst (asm) eventually_sequentially_seg[where k=k]) lemma LIMSEQ_Suc: "f \ l \ (\n. f (Suc n)) \ l" by (drule LIMSEQ_ignore_initial_segment [where k="Suc 0"]) simp lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) \ l \ f \ l" by (rule LIMSEQ_offset [where k="Suc 0"]) simp lemma LIMSEQ_lessThan_iff_atMost: shows "(\n. f {.. x \ (\n. f {..n}) \ x" apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost) done lemma (in t2_space) LIMSEQ_Uniq: "\\<^sub>\\<^sub>1l. X \ l" by (simp add: tendsto_unique') lemma (in t2_space) LIMSEQ_unique: "X \ a \ X \ b \ a = b" using trivial_limit_sequentially by (rule tendsto_unique) lemma LIMSEQ_le_const: "X \ x \ \N. \n\N. a \ X n \ a \ x" for a x :: "'a::linorder_topology" by (simp add: eventually_at_top_linorder tendsto_lowerbound) lemma LIMSEQ_le: "X \ x \ Y \ y \ \N. \n\N. X n \ Y n \ x \ y" for x y :: "'a::linorder_topology" using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) lemma LIMSEQ_le_const2: "X \ x \ \N. \n\N. X n \ a \ x \ a" for a x :: "'a::linorder_topology" by (rule LIMSEQ_le[of X x "\n. a"]) auto lemma Lim_bounded: "f \ l \ \n\M. f n \ C \ l \ C" for l :: "'a::linorder_topology" by (intro LIMSEQ_le_const2) auto lemma Lim_bounded2: fixes f :: "nat \ 'a::linorder_topology" assumes lim:"f \ l" and ge: "\n\N. f n \ C" shows "l \ C" using ge by (intro tendsto_le[OF trivial_limit_sequentially lim tendsto_const]) (auto simp: eventually_sequentially) lemma lim_mono: fixes X Y :: "nat \ 'a::linorder_topology" assumes "\n. N \ n \ X n \ Y n" and "X \ x" and "Y \ y" shows "x \ y" using assms(1) by (intro LIMSEQ_le[OF assms(2,3)]) auto lemma Sup_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "a \ Sup s" by (metis Lim_bounded assms complete_lattice_class.Sup_upper) lemma Inf_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "Inf s \ a" by (metis Lim_bounded2 assms complete_lattice_class.Inf_lower) lemma SUP_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes inc: "incseq X" and l: "X \ l" shows "(SUP n. X n) = l" using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma INF_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes dec: "decseq X" and l: "X \ l" shows "(INF n. X n) = l" using LIMSEQ_INF[OF dec] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma convergentD: "convergent X \ \L. X \ L" by (simp add: convergent_def) lemma convergentI: "X \ L \ convergent X" by (auto simp add: convergent_def) lemma convergent_LIMSEQ_iff: "convergent X \ X \ lim X" by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) lemma convergent_const: "convergent (\n. c)" by (rule convergentI) (rule tendsto_const) lemma monoseq_le: "monoseq a \ a \ x \ (\n. a n \ x) \ (\m. \n\m. a m \ a n) \ (\n. x \ a n) \ (\m. \n\m. a n \ a m)" for x :: "'a::linorder_topology" by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) lemma LIMSEQ_subseq_LIMSEQ: "X \ L \ strict_mono f \ (X \ f) \ L" unfolding comp_def by (rule filterlim_compose [of X, OF _ filterlim_subseq]) lemma convergent_subseq_convergent: "convergent X \ strict_mono f \ convergent (X \ f)" by (auto simp: convergent_def intro: LIMSEQ_subseq_LIMSEQ) lemma limI: "X \ L \ lim X = L" by (rule tendsto_Lim) (rule trivial_limit_sequentially) lemma lim_le: "convergent f \ (\n. f n \ x) \ lim f \ x" for x :: "'a::linorder_topology" using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) lemma lim_const [simp]: "lim (\m. a) = a" by (simp add: limI) subsubsection \Increasing and Decreasing Series\ lemma incseq_le: "incseq X \ X \ L \ X n \ L" for L :: "'a::linorder_topology" by (metis incseq_def LIMSEQ_le_const) lemma decseq_ge: "decseq X \ X \ L \ L \ X n" for L :: "'a::linorder_topology" by (metis decseq_def LIMSEQ_le_const2) subsection \First countable topologies\ class first_countable_topology = topological_space + assumes first_countable_basis: "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" lemma (in first_countable_topology) countable_basis_at_decseq: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ (A i)" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" proof atomize_elim from first_countable_basis[of x] obtain A :: "nat \ 'a set" where nhds: "\i. open (A i)" "\i. x \ A i" and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto define F where "F n = (\i\n. A i)" for n show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" proof (safe intro!: exI[of _ F]) fix i show "open (F i)" using nhds(1) by (auto simp: F_def) show "x \ F i" using nhds(2) by (auto simp: F_def) next fix S assume "open S" "x \ S" from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto moreover have "\j. i \ j \ F j \ F i" by (simp add: Inf_superset_mono F_def image_mono) ultimately show "eventually (\i. F i \ S) sequentially" by (auto simp: eventually_sequentially) qed qed lemma (in first_countable_topology) nhds_countable: obtains X :: "nat \ 'a set" where "decseq X" "\n. open (X n)" "\n. x \ X n" "nhds x = (INF n. principal (X n))" proof - from first_countable_basis obtain A :: "nat \ 'a set" where *: "\n. x \ A n" "\n. open (A n)" "\S. open S \ x \ S \ \i. A i \ S" by metis show thesis proof show "decseq (\n. \i\n. A i)" by (simp add: antimono_iff_le_Suc atMost_Suc) show "x \ (\i\n. A i)" "\n. open (\i\n. A i)" for n using * by auto with * show "nhds x = (INF n. principal (\i\n. A i))" unfolding nhds_def apply (intro INF_eq) apply fastforce apply blast done qed qed lemma (in first_countable_topology) countable_basis: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" proof atomize_elim obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by (rule countable_basis_at_decseq) blast have "eventually (\n. F n \ S) sequentially" if "\n. F n \ A n" "open S" "x \ S" for F S using *(3)[of S] that by (auto elim: eventually_mono simp: subset_eq) with * show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F \ x)" by (intro exI[of _ A]) (auto simp: tendsto_def) qed lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within: assumes "\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (inf (nhds a) (principal s))" proof (rule ccontr) obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. a \ A i" "\F. \n. F n \ A n \ F \ a" by (rule countable_basis) blast assume "\ ?thesis" with * have "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" unfolding eventually_inf_principal eventually_nhds by (intro choice) fastforce then obtain F where F: "\n. F n \ s" and "\n. F n \ A n" and F': "\n. \ P (F n)" by blast with * have "F \ a" by auto then have "eventually (\n. P (F n)) sequentially" using assms F by simp then show False by (simp add: F') qed lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially: "eventually P (inf (nhds a) (principal s)) \ (\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially)" proof (safe intro!: sequentially_imp_eventually_nhds_within) assume "eventually P (inf (nhds a) (principal s))" then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" by (auto simp: eventually_inf_principal eventually_nhds) moreover fix f assume "\n. f n \ s" "f \ a" ultimately show "eventually (\n. P (f n)) sequentially" by (auto dest!: topological_tendstoD elim: eventually_mono) qed lemma (in first_countable_topology) eventually_nhds_iff_sequentially: "eventually P (nhds a) \ (\f. f \ a \ eventually (\n. P (f n)) sequentially)" using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp (*Thanks to Sébastien Gouëzel*) lemma Inf_as_limit: fixes A::"'a::{linorder_topology, first_countable_topology, complete_linorder} set" assumes "A \ {}" shows "\u. (\n. u n \ A) \ u \ Inf A" proof (cases "Inf A \ A") case True show ?thesis by (rule exI[of _ "\n. Inf A"], auto simp add: True) next case False obtain y where "y \ A" using assms by auto then have "Inf A < y" using False Inf_lower less_le by auto obtain F :: "nat \ 'a set" where F: "\i. open (F i)" "\i. Inf A \ F i" "\u. (\n. u n \ F n) \ u \ Inf A" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ F n \ z \ A)" have "\z. z \ U \ z \ A" if "Inf A \ U" "open U" for U proof - obtain b where "b > Inf A" "{Inf A .. U" using open_right[OF \open U\ \Inf A \ U\ \Inf A < y\] by auto obtain z where "z < b" "z \ A" using \Inf A < b\ Inf_less_iff by auto then have "z \ {Inf A ..z \ A\ \{Inf A .. U\ by auto qed then have *: "u n \ F n \ u n \ A" for n using \Inf A \ F n\ \open (F n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ Inf A" using F(3) by simp then show ?thesis using * by auto qed lemma tendsto_at_iff_sequentially: "(f \ a) (at x within s) \ (\X. (\i. X i \ s - {x}) \ X \ x \ ((f \ X) \ a))" for f :: "'a::first_countable_topology \ _" unfolding filterlim_def[of _ "nhds a"] le_filter_def eventually_filtermap at_within_def eventually_nhds_within_iff_sequentially comp_def by metis lemma approx_from_above_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x < y" shows "\u. (\n. u n > x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z > x)" have "\z. z \ U \ x < z" if "x \ U" "open U" for U using open_right[OF \open U\ \x \ U\ \x < y\] by (meson atLeastLessThan_iff dense less_imp_le subset_eq) then have *: "u n \ A n \ x < u n" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed lemma approx_from_below_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x > y" shows "\u. (\n. u n < x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z < x)" have "\z. z \ U \ z < x" if "x \ U" "open U" for U using open_left[OF \open U\ \x \ U\ \x > y\] by (meson dense greaterThanAtMost_iff less_imp_le subset_eq) then have *: "u n \ A n \ u n < x" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed subsection \Function limit at a point\ abbreviation LIM :: "('a::topological_space \ 'b::topological_space) \ 'a \ 'b \ bool" ("((_)/ \(_)/\ (_))" [60, 0, 60] 60) where "f \a\ L \ (f \ L) (at a)" lemma tendsto_within_open: "a \ S \ open S \ (f \ l) (at a within S) \ (f \a\ l)" by (simp add: tendsto_def at_within_open[where S = S]) lemma tendsto_within_open_NO_MATCH: "a \ S \ NO_MATCH UNIV S \ open S \ (f \ l)(at a within S) \ (f \ l)(at a)" for f :: "'a::topological_space \ 'b::topological_space" using tendsto_within_open by blast lemma LIM_const_not_eq[tendsto_intros]: "k \ L \ \ (\x. k) \a\ L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] lemma LIM_const_eq: "(\x. k) \a\ L \ k = L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemma LIM_unique: "f \a\ L \ f \a\ M \ L = M" for a :: "'a::perfect_space" and L M :: "'b::t2_space" using at_neq_bot by (rule tendsto_unique) lemma LIM_Uniq: "\\<^sub>\\<^sub>1L::'b::t2_space. f \a\ L" for a :: "'a::perfect_space" by (auto simp add: Uniq_def LIM_unique) text \Limits are equal for functions equal except at limit point.\ lemma LIM_equal: "\x. x \ a \ f x = g x \ (f \a\ l) \ (g \a\ l)" by (simp add: tendsto_def eventually_at_topological) lemma LIM_cong: "a = b \ (\x. x \ b \ f x = g x) \ l = m \ (f \a\ l) \ (g \b\ m)" by (simp add: LIM_equal) lemma tendsto_cong_limit: "(f \ l) F \ k = l \ (f \ k) F" by simp lemma tendsto_at_iff_tendsto_nhds: "g \l\ g l \ (g \ g l) (nhds l)" unfolding tendsto_def eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma tendsto_compose: "g \l\ g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) lemma tendsto_compose_eventually: "g \l\ m \ (f \ l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) \ m) F" by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) lemma LIM_compose_eventually: assumes "f \a\ b" and "g \b\ c" and "eventually (\x. f x \ b) (at a)" shows "(\x. g (f x)) \a\ c" using assms(2,1,3) by (rule tendsto_compose_eventually) lemma tendsto_compose_filtermap: "((g \ f) \ T) F \ (g \ T) (filtermap f F)" by (simp add: filterlim_def filtermap_filtermap comp_def) lemma tendsto_compose_at: assumes f: "(f \ y) F" and g: "(g \ z) (at y)" and fg: "eventually (\w. f w = y \ g y = z) F" shows "((g \ f) \ z) F" proof - have "(\\<^sub>F a in F. f a \ y) \ g y = z" using fg by force moreover have "(g \ z) (filtermap f F) \ \ (\\<^sub>F a in F. f a \ y)" by (metis (no_types) filterlim_atI filterlim_def tendsto_mono f g) ultimately show ?thesis by (metis (no_types) f filterlim_compose filterlim_filtermap g tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap) qed subsubsection \Relation of \LIM\ and \LIMSEQ\\ lemma (in first_countable_topology) sequentially_imp_eventually_within: "(\f. (\n. f n \ s \ f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a within s)" unfolding at_within_def by (intro sequentially_imp_eventually_nhds_within) auto lemma (in first_countable_topology) sequentially_imp_eventually_at: "(\f. (\n. f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a)" using sequentially_imp_eventually_within [where s=UNIV] by simp lemma LIMSEQ_SEQ_conv: "(\S. (\n. S n \ a) \ S \ a \ (\n. X (S n)) \ L) \ X \a\ L" (is "?lhs=?rhs") for a :: "'a::first_countable_topology" and L :: "'b::topological_space" proof assume ?lhs then show ?rhs by (simp add: sequentially_imp_eventually_within tendsto_def) next assume ?rhs then show ?lhs using tendsto_compose_eventually eventuallyI by blast qed lemma sequentially_imp_eventually_at_left: fixes a :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "b < a" and *: "\f. (\n. b < f n) \ (\n. f n < a) \ incseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_left a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {..< a} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ b < X (s n)) \ (X (s n) \ X (s (Suc n)) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. b < X n \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(1) [OF X(2) b]) then show "\x. \ P (X x) \ b < X x" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ b < X n \ X x < X n \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(1)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ b < X n) \ (X x \ X n \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "b < X (s n)" and "X (s n) < a" and "incseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_left_sequentially: fixes a b :: "'b::{linorder_topology,first_countable_topology}" assumes "b < a" assumes *: "\S. (\n. S n < a) \ (\n. b < S n) \ incseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_left a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_left) lemma sequentially_imp_eventually_at_right: fixes a b :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "a < b" assumes *: "\f. (\n. a < f n) \ (\n. f n < b) \ decseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_right a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {a <..} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ X (s n) < b) \ (X (s (Suc n)) \ X (s n) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. X n < b \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(2) [OF X(2) b]) then show "\x. \ P (X x) \ X x < b" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ X n < b \ X n < X x \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(2)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ X n < b) \ (X n \ X x \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "a < X (s n)" and "X (s n) < b" and "decseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_right_sequentially: fixes a :: "_ :: {linorder_topology, first_countable_topology}" assumes "a < b" and *: "\S. (\n. a < S n) \ (\n. S n < b) \ decseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_right a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_right) subsection \Continuity\ subsubsection \Continuity on a set\ definition continuous_on :: "'a set \ ('a::topological_space \ 'b::topological_space) \ bool" where "continuous_on s f \ (\x\s. (f \ f x) (at x within s))" lemma continuous_on_cong [cong]: "s = t \ (\x. x \ t \ f x = g x) \ continuous_on s f \ continuous_on t g" unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_at_filter) lemma continuous_on_cong_simp: "s = t \ (\x. x \ t =simp=> f x = g x) \ continuous_on s f \ continuous_on t g" unfolding simp_implies_def by (rule continuous_on_cong) lemma continuous_on_topological: "continuous_on s f \ (\x\s. \B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_on_def tendsto_def eventually_at_topological by metis lemma continuous_on_open_invariant: "continuous_on s f \ (\B. open B \ (\A. open A \ A \ s = f -` B \ s))" proof safe fix B :: "'b set" assume "continuous_on s f" "open B" then have "\x\f -` B \ s. (\A. open A \ x \ A \ s \ A \ f -` B)" by (auto simp: continuous_on_topological subset_eq Ball_def imp_conjL) then obtain A where "\x\f -` B \ s. open (A x) \ x \ A x \ s \ A x \ f -` B" unfolding bchoice_iff .. then show "\A. open A \ A \ s = f -` B \ s" by (intro exI[of _ "\x\f -` B \ s. A x"]) auto next assume B: "\B. open B \ (\A. open A \ A \ s = f -` B \ s)" show "continuous_on s f" unfolding continuous_on_topological proof safe fix x B assume "x \ s" "open B" "f x \ B" with B obtain A where A: "open A" "A \ s = f -` B \ s" by auto with \x \ s\ \f x \ B\ show "\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)" by (intro exI[of _ A]) auto qed qed lemma continuous_on_open_vimage: "open s \ continuous_on s f \ (\B. open B \ open (f -` B \ s))" unfolding continuous_on_open_invariant by (metis open_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary continuous_imp_open_vimage: assumes "continuous_on s f" "open s" "open B" "f -` B \ s" shows "open (f -` B)" by (metis assms continuous_on_open_vimage le_iff_inf) corollary open_vimage[continuous_intros]: assumes "open s" and "continuous_on UNIV f" shows "open (f -` s)" using assms by (simp add: continuous_on_open_vimage [OF open_UNIV]) lemma continuous_on_closed_invariant: "continuous_on s f \ (\B. closed B \ (\A. closed A \ A \ s = f -` B \ s))" proof - have *: "(\A. P A \ Q (- A)) \ (\A. P A) \ (\A. Q A)" for P Q :: "'b set \ bool" by (metis double_compl) show ?thesis unfolding continuous_on_open_invariant by (intro *) (auto simp: open_closed[symmetric]) qed lemma continuous_on_closed_vimage: "closed s \ continuous_on s f \ (\B. closed B \ closed (f -` B \ s))" unfolding continuous_on_closed_invariant by (metis closed_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary closed_vimage_Int[continuous_intros]: assumes "closed s" and "continuous_on t f" and t: "closed t" shows "closed (f -` s \ t)" using assms by (simp add: continuous_on_closed_vimage [OF t]) corollary closed_vimage[continuous_intros]: assumes "closed s" and "continuous_on UNIV f" shows "closed (f -` s)" using closed_vimage_Int [OF assms] by simp lemma continuous_on_empty [simp]: "continuous_on {} f" by (simp add: continuous_on_def) lemma continuous_on_sing [simp]: "continuous_on {x} f" by (simp add: continuous_on_def at_within_def) lemma continuous_on_open_Union: "(\s. s \ S \ open s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f" unfolding continuous_on_def by safe (metis open_Union at_within_open UnionI) lemma continuous_on_open_UN: "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" by (rule continuous_on_open_Union) auto lemma continuous_on_open_Un: "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" using continuous_on_open_Union [of "{s,t}"] by auto lemma continuous_on_closed_Un: "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib) lemma continuous_on_closed_Union: assumes "finite I" "\i. i \ I \ closed (U i)" "\i. i \ I \ continuous_on (U i) f" shows "continuous_on (\ i \ I. U i) f" using assms by (induction I) (auto intro!: continuous_on_closed_Un) lemma continuous_on_If: assumes closed: "closed s" "closed t" and cont: "continuous_on s f" "continuous_on t g" and P: "\x. x \ s \ \ P x \ f x = g x" "\x. x \ t \ P x \ f x = g x" shows "continuous_on (s \ t) (\x. if P x then f x else g x)" (is "continuous_on _ ?h") proof- from P have "\x\s. f x = ?h x" "\x\t. g x = ?h x" by auto with cont have "continuous_on s ?h" "continuous_on t ?h" by simp_all with closed show ?thesis by (rule continuous_on_closed_Un) qed lemma continuous_on_cases: "closed s \ closed t \ continuous_on s f \ continuous_on t g \ \x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_If) auto lemma continuous_on_id[continuous_intros,simp]: "continuous_on s (\x. x)" unfolding continuous_on_def by fast lemma continuous_on_id'[continuous_intros,simp]: "continuous_on s id" unfolding continuous_on_def id_def by fast lemma continuous_on_const[continuous_intros,simp]: "continuous_on s (\x. c)" unfolding continuous_on_def by auto lemma continuous_on_subset: "continuous_on s f \ t \ s \ continuous_on t f" unfolding continuous_on_def by (metis subset_eq tendsto_within_subset) lemma continuous_on_compose[continuous_intros]: "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g \ f)" unfolding continuous_on_topological by simp metis lemma continuous_on_compose2: "continuous_on t g \ continuous_on s f \ f ` s \ t \ continuous_on s (\x. g (f x))" using continuous_on_compose[of s f g] continuous_on_subset by (force simp add: comp_def) lemma continuous_on_generate_topology: assumes *: "open = generate_topology X" and **: "\B. B \ X \ \C. open C \ C \ A = f -` B \ A" shows "continuous_on A f" unfolding continuous_on_open_invariant proof safe fix B :: "'a set" assume "open B" then show "\C. open C \ C \ A = f -` B \ A" unfolding * proof induct case (UN K) then obtain C where "\k. k \ K \ open (C k)" "\k. k \ K \ C k \ A = f -` k \ A" by metis then show ?case by (intro exI[of _ "\k\K. C k"]) blast qed (auto intro: **) qed lemma continuous_onI_mono: fixes f :: "'a::linorder_topology \ 'b::{dense_order,linorder_topology}" assumes "open (f`A)" and mono: "\x y. x \ A \ y \ A \ x \ y \ f x \ f y" shows "continuous_on A f" proof (rule continuous_on_generate_topology[OF open_generated_order], safe) have monoD: "\x y. x \ A \ y \ A \ f x < f y \ x < y" by (auto simp: not_le[symmetric] mono) have "\x. x \ A \ f x < b \ a < x" if a: "a \ A" and fa: "f a < b" for a b proof - obtain y where "f a < y" "{f a ..< y} \ f`A" using open_right[OF \open (f`A)\, of "f a" b] a fa by auto obtain z where z: "f a < z" "z < min b y" using dense[of "f a" "min b y"] \f a < y\ \f a < b\ by auto then obtain c where "z = f c" "c \ A" using \{f a ..< y} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {.. A" for b by (intro exI[of _ "(\x\{x\A. f x < b}. {..< x})"]) (auto intro: le_less_trans[OF mono] less_imp_le) have "\x. x \ A \ b < f x \ x < a" if a: "a \ A" and fa: "b < f a" for a b proof - note a fa moreover obtain y where "y < f a" "{y <.. f a} \ f`A" using open_left[OF \open (f`A)\, of "f a" b] a fa by auto then obtain z where z: "max b y < z" "z < f a" using dense[of "max b y" "f a"] \y < f a\ \b < f a\ by auto then obtain c where "z = f c" "c \ A" using \{y <.. f a} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {b <..} \ A" for b by (intro exI[of _ "(\x\{x\A. b < f x}. {x <..})"]) (auto intro: less_le_trans[OF _ mono] less_imp_le) qed lemma continuous_on_IccI: "\(f \ f a) (at_right a); (f \ f b) (at_left b); (\x. a < x \ x < b \ f \x\ f x); a < b\ \ continuous_on {a .. b} f" for a::"'a::linorder_topology" using at_within_open[of _ "{a<.. f a) (at_right a)" and continuous_on_Icc_at_leftD: "(f \ f b) (at_left b)" using assms by (auto simp: at_within_Icc_at_right at_within_Icc_at_left continuous_on_def dest: bspec[where x=a] bspec[where x=b]) lemma continuous_on_discrete [simp]: "continuous_on A (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_on_def at_discrete) subsubsection \Continuity at a point\ definition continuous :: "'a::t2_space filter \ ('a \ 'b::topological_space) \ bool" where "continuous F f \ (f \ f (Lim F (\x. x))) F" lemma continuous_bot[continuous_intros, simp]: "continuous bot f" unfolding continuous_def by auto lemma continuous_trivial_limit: "trivial_limit net \ continuous net f" by simp lemma continuous_within: "continuous (at x within s) f \ (f \ f x) (at x within s)" by (cases "trivial_limit (at x within s)") (auto simp add: Lim_ident_at continuous_def) lemma continuous_within_topological: "continuous (at x within s) f \ (\B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_within tendsto_def eventually_at_topological by metis lemma continuous_within_compose[continuous_intros]: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (g \ f)" by (simp add: continuous_within_topological) metis lemma continuous_within_compose2: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (\x. g (f x))" using continuous_within_compose[of x s f g] by (simp add: comp_def) lemma continuous_at: "continuous (at x) f \ f \x\ f x" using continuous_within[of x UNIV f] by simp lemma continuous_ident[continuous_intros, simp]: "continuous (at x within S) (\x. x)" unfolding continuous_within by (rule tendsto_ident_at) lemma continuous_id[continuous_intros, simp]: "continuous (at x within S) id" by (simp add: id_def) lemma continuous_const[continuous_intros, simp]: "continuous F (\x. c)" unfolding continuous_def by (rule tendsto_const) lemma continuous_on_eq_continuous_within: "continuous_on s f \ (\x\s. continuous (at x within s) f)" unfolding continuous_on_def continuous_within .. lemma continuous_discrete [simp]: "continuous (at x within A) (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_def at_discrete) abbreviation isCont :: "('a::t2_space \ 'b::topological_space) \ 'a \ bool" where "isCont f a \ continuous (at a) f" lemma isCont_def: "isCont f a \ f \a\ f a" by (rule continuous_at) lemma isContD: "isCont f x \ f \x\ f x" by (simp add: isCont_def) lemma isCont_cong: assumes "eventually (\x. f x = g x) (nhds x)" shows "isCont f x \ isCont g x" proof - from assms have [simp]: "f x = g x" by (rule eventually_nhds_x_imp_x) from assms have "eventually (\x. f x = g x) (at x)" by (auto simp: eventually_at_filter elim!: eventually_mono) with assms have "isCont f x \ isCont g x" unfolding isCont_def by (intro filterlim_cong) (auto elim!: eventually_mono) with assms show ?thesis by simp qed lemma continuous_at_imp_continuous_at_within: "isCont f x \ continuous (at x within s) f" by (auto intro: tendsto_mono at_le simp: continuous_at continuous_within) lemma continuous_on_eq_continuous_at: "open s \ continuous_on s f \ (\x\s. isCont f x)" by (simp add: continuous_on_def continuous_at at_within_open[of _ s]) lemma continuous_within_open: "a \ A \ open A \ continuous (at a within A) f \ isCont f a" by (simp add: at_within_open_NO_MATCH) lemma continuous_at_imp_continuous_on: "\x\s. isCont f x \ continuous_on s f" by (auto intro: continuous_at_imp_continuous_at_within simp: continuous_on_eq_continuous_within) lemma isCont_o2: "isCont f a \ isCont g (f a) \ isCont (\x. g (f x)) a" unfolding isCont_def by (rule tendsto_compose) lemma continuous_at_compose[continuous_intros]: "isCont f a \ isCont g (f a) \ isCont (g \ f) a" unfolding o_def by (rule isCont_o2) lemma isCont_tendsto_compose: "isCont g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding isCont_def by (rule tendsto_compose) lemma continuous_on_tendsto_compose: assumes f_cont: "continuous_on s f" and g: "(g \ l) F" and l: "l \ s" and ev: "\\<^sub>Fx in F. g x \ s" shows "((\x. f (g x)) \ f l) F" proof - from f_cont l have f: "(f \ f l) (at l within s)" by (simp add: continuous_on_def) have i: "((\x. if g x = l then f l else f (g x)) \ f l) F" by (rule filterlim_If) (auto intro!: filterlim_compose[OF f] eventually_conj tendsto_mono[OF _ g] simp: filterlim_at eventually_inf_principal eventually_mono[OF ev]) show ?thesis by (rule filterlim_cong[THEN iffD1[OF _ i]]) auto qed lemma continuous_within_compose3: "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" using continuous_at_imp_continuous_at_within continuous_within_compose2 by blast lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (nhds a))" then obtain S where "open S" "a \ S" "\x\S. P (f x)" by (auto simp: eventually_filtermap eventually_nhds) then show "eventually P (nhds (f a))" unfolding eventually_nhds by (intro exI[of _ "f`S"]) (auto intro!: open_map) qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) lemma continuous_at_split: "continuous (at x) f \ continuous (at_left x) f \ continuous (at_right x) f" for x :: "'a::linorder_topology" by (simp add: continuous_within filterlim_at_split) lemma continuous_on_max [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. max (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_max) lemma continuous_on_min [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_min) lemma continuous_max [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (max (f x) (g x)))" by (simp add: tendsto_max continuous_def) lemma continuous_min [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (min (f x) (g x)))" by (simp add: tendsto_min continuous_def) text \ The following open/closed Collect lemmas are ported from Sébastien Gouëzel's \Ergodic_Theory\. \ lemma open_Collect_neq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x \ g x}" proof (rule openI) fix t assume "t \ {x. f x \ g x}" then obtain U V where *: "open U" "open V" "f t \ U" "g t \ V" "U \ V = {}" by (auto simp add: separation_t2) with open_vimage[OF \open U\ f] open_vimage[OF \open V\ g] show "\T. open T \ t \ T \ T \ {x. f x \ g x}" by (intro exI[of _ "f -` U \ g -` V"]) auto qed lemma closed_Collect_eq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x = g x}" using open_Collect_neq[OF f g] by (simp add: closed_def Collect_neg_eq) lemma open_Collect_less: fixes f g :: "'a::topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x < g x}" proof (rule openI) fix t assume t: "t \ {x. f x < g x}" show "\T. open T \ t \ T \ T \ {x. f x < g x}" proof (cases "\z. f t < z \ z < g t") case True then obtain z where "f t < z \ z < g t" by blast then show ?thesis using open_vimage[OF _ f, of "{..< z}"] open_vimage[OF _ g, of "{z <..}"] by (intro exI[of _ "f -` {.. g -` {z<..}"]) auto next case False then have *: "{g t ..} = {f t <..}" "{..< g t} = {.. f t}" using t by (auto intro: leI) show ?thesis using open_vimage[OF _ f, of "{..< g t}"] open_vimage[OF _ g, of "{f t <..}"] t apply (intro exI[of _ "f -` {..< g t} \ g -` {f t<..}"]) apply (simp add: open_Int) apply (auto simp add: *) done qed qed lemma closed_Collect_le: fixes f g :: "'a :: topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x \ g x}" using open_Collect_less [OF g f] by (simp add: closed_def Collect_neg_eq[symmetric] not_le) subsubsection \Open-cover compactness\ context topological_space begin definition compact :: "'a set \ bool" where compact_eq_Heine_Borel: (* This name is used for backwards compatibility *) "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" lemma compactI: assumes "\C. \t\C. open t \ s \ \C \ \C'. C' \ C \ finite C' \ s \ \C'" shows "compact s" unfolding compact_eq_Heine_Borel using assms by metis lemma compact_empty[simp]: "compact {}" by (auto intro!: compactI) lemma compactE: (*related to COMPACT_IMP_HEINE_BOREL in HOL Light*) assumes "compact S" "S \ \\" "\B. B \ \ \ open B" obtains \' where "\' \ \" "finite \'" "S \ \\'" by (meson assms compact_eq_Heine_Borel) lemma compactE_image: assumes "compact S" and opn: "\T. T \ C \ open (f T)" and S: "S \ (\c\C. f c)" obtains C' where "C' \ C" and "finite C'" and "S \ (\c\C'. f c)" apply (rule compactE[OF \compact S\ S]) using opn apply force by (metis finite_subset_image) lemma compact_Int_closed [intro]: assumes "compact S" and "closed T" shows "compact (S \ T)" proof (rule compactI) fix C assume C: "\c\C. open c" assume cover: "S \ T \ \C" from C \closed T\ have "\c\C \ {- T}. open c" by auto moreover from cover have "S \ \(C \ {- T})" by auto ultimately have "\D\C \ {- T}. finite D \ S \ \D" using \compact S\ unfolding compact_eq_Heine_Borel by auto then obtain D where "D \ C \ {- T} \ finite D \ S \ \D" .. then show "\D\C. finite D \ S \ T \ \D" by (intro exI[of _ "D - {-T}"]) auto qed lemma compact_diff: "\compact S; open T\ \ compact(S - T)" by (simp add: Diff_eq compact_Int_closed open_closed) lemma inj_setminus: "inj_on uminus (A::'a set set)" by (auto simp: inj_on_def) subsection \Finite intersection property\ lemma compact_fip: "compact U \ (\A. (\a\A. closed a) \ (\B \ A. finite B \ U \ \B \ {}) \ U \ \A \ {})" (is "_ \ ?R") proof (safe intro!: compact_eq_Heine_Borel[THEN iffD2]) fix A assume "compact U" assume A: "\a\A. closed a" "U \ \A = {}" assume fin: "\B \ A. finite B \ U \ \B \ {}" from A have "(\a\uminus`A. open a) \ U \ \(uminus`A)" by auto with \compact U\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B)" unfolding compact_eq_Heine_Borel by (metis subset_image_iff) with fin[THEN spec, of B] show False by (auto dest: finite_imageD intro: inj_setminus) next fix A assume ?R assume "\a\A. open a" "U \ \A" then have "U \ \(uminus`A) = {}" "\a\uminus`A. closed a" by auto with \?R\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B) = {}" by (metis subset_image_iff) then show "\T\A. finite T \ U \ \T" by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) qed lemma compact_imp_fip: assumes "compact S" and "\T. T \ F \ closed T" and "\F'. finite F' \ F' \ F \ S \ (\F') \ {}" shows "S \ (\F) \ {}" using assms unfolding compact_fip by auto lemma compact_imp_fip_image: assumes "compact s" and P: "\i. i \ I \ closed (f i)" and Q: "\I'. finite I' \ I' \ I \ (s \ (\i\I'. f i) \ {})" shows "s \ (\i\I. f i) \ {}" proof - from P have "\i \ f ` I. closed i" by blast moreover have "\A. finite A \ A \ f ` I \ (s \ (\A) \ {})" by (metis Q finite_subset_image) ultimately show "s \ (\(f ` I)) \ {}" by (metis \compact s\ compact_imp_fip) qed end lemma (in t2_space) compact_imp_closed: assumes "compact s" shows "closed s" unfolding closed_def proof (rule openI) fix y assume "y \ - s" let ?C = "\x\s. {u. open u \ x \ u \ eventually (\y. y \ u) (nhds y)}" have "s \ \?C" proof fix x assume "x \ s" with \y \ - s\ have "x \ y" by clarsimp then have "\u v. open u \ open v \ x \ u \ y \ v \ u \ v = {}" by (rule hausdorff) with \x \ s\ show "x \ \?C" unfolding eventually_nhds by auto qed then obtain D where "D \ ?C" and "finite D" and "s \ \D" by (rule compactE [OF \compact s\]) auto from \D \ ?C\ have "\x\D. eventually (\y. y \ x) (nhds y)" by auto with \finite D\ have "eventually (\y. y \ \D) (nhds y)" by (simp add: eventually_ball_finite) with \s \ \D\ have "eventually (\y. y \ s) (nhds y)" by (auto elim!: eventually_mono) then show "\t. open t \ y \ t \ t \ - s" by (simp add: eventually_nhds subset_eq) qed lemma compact_continuous_image: assumes f: "continuous_on s f" and s: "compact s" shows "compact (f ` s)" proof (rule compactI) fix C assume "\c\C. open c" and cover: "f`s \ \C" with f have "\c\C. \A. open A \ A \ s = f -` c \ s" unfolding continuous_on_open_invariant by blast then obtain A where A: "\c\C. open (A c) \ A c \ s = f -` c \ s" unfolding bchoice_iff .. with cover have "\c. c \ C \ open (A c)" "s \ (\c\C. A c)" by (fastforce simp add: subset_eq set_eq_iff)+ from compactE_image[OF s this] obtain D where "D \ C" "finite D" "s \ (\c\D. A c)" . with A show "\D \ C. finite D \ f`s \ \D" by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ qed lemma continuous_on_inv: fixes f :: "'a::topological_space \ 'b::t2_space" assumes "continuous_on s f" and "compact s" and "\x\s. g (f x) = x" shows "continuous_on (f ` s) g" unfolding continuous_on_topological proof (clarsimp simp add: assms(3)) fix x :: 'a and B :: "'a set" assume "x \ s" and "open B" and "x \ B" have 1: "\x\s. f x \ f ` (s - B) \ x \ s - B" using assms(3) by (auto, metis) have "continuous_on (s - B) f" using \continuous_on s f\ Diff_subset by (rule continuous_on_subset) moreover have "compact (s - B)" using \open B\ and \compact s\ unfolding Diff_eq by (intro compact_Int_closed closed_Compl) ultimately have "compact (f ` (s - B))" by (rule compact_continuous_image) then have "closed (f ` (s - B))" by (rule compact_imp_closed) then have "open (- f ` (s - B))" by (rule open_Compl) moreover have "f x \ - f ` (s - B)" using \x \ s\ and \x \ B\ by (simp add: 1) moreover have "\y\s. f y \ - f ` (s - B) \ y \ B" by (simp add: 1) ultimately show "\A. open A \ f x \ A \ (\y\s. f y \ A \ y \ B)" by fast qed lemma continuous_on_inv_into: fixes f :: "'a::topological_space \ 'b::t2_space" assumes s: "continuous_on s f" "compact s" and f: "inj_on f s" shows "continuous_on (f ` s) (the_inv_into s f)" by (rule continuous_on_inv[OF s]) (auto simp: the_inv_into_f_f[OF f]) lemma (in linorder_topology) compact_attains_sup: assumes "compact S" "S \ {}" shows "\s\S. \t\S. t \ s" proof (rule classical) assume "\ (\s\S. \t\S. t \ s)" then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" by (metis not_le) then have "\s. s\S \ open {..< t s}" "S \ (\s\S. {..< t s})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" by (metis compactE_image) with \S \ {}\ have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" by (auto intro!: Max_in) with C have "S \ {..< Max (t`C)}" by (auto intro: less_le_trans simp: subset_eq) with t Max \C \ S\ show ?thesis by fastforce qed lemma (in linorder_topology) compact_attains_inf: assumes "compact S" "S \ {}" shows "\s\S. \t\S. s \ t" proof (rule classical) assume "\ (\s\S. \t\S. s \ t)" then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" by (metis not_le) then have "\s. s\S \ open {t s <..}" "S \ (\s\S. {t s <..})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" by (metis compactE_image) with \S \ {}\ have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" by (auto intro!: Min_in) with C have "S \ {Min (t`C) <..}" by (auto intro: le_less_trans simp: subset_eq) with t Min \C \ S\ show ?thesis by fastforce qed lemma continuous_attains_sup: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto lemma continuous_attains_inf: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto subsection \Connectedness\ context topological_space begin definition "connected S \ \ (\A B. open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S \ {} \ B \ S \ {})" lemma connectedI: "(\A B. open A \ open B \ A \ U \ {} \ B \ U \ {} \ A \ B \ U = {} \ U \ A \ B \ False) \ connected U" by (auto simp: connected_def) lemma connected_empty [simp]: "connected {}" by (auto intro!: connectedI) lemma connected_sing [simp]: "connected {x}" by (auto intro!: connectedI) lemma connectedD: "connected A \ open U \ open V \ U \ V \ A = {} \ A \ U \ V \ U \ A = {} \ V \ A = {}" by (auto simp: connected_def) end lemma connected_closed: "connected s \ \ (\A B. closed A \ closed B \ s \ A \ B \ A \ B \ s = {} \ A \ s \ {} \ B \ s \ {})" apply (simp add: connected_def del: ex_simps, safe) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: closed_def [symmetric]) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: open_closed [symmetric]) done lemma connected_closedD: "\connected s; A \ B \ s = {}; s \ A \ B; closed A; closed B\ \ A \ s = {} \ B \ s = {}" by (simp add: connected_closed) lemma connected_Union: assumes cs: "\s. s \ S \ connected s" and ne: "\S \ {}" shows "connected(\S)" proof (rule connectedI) fix A B assume A: "open A" and B: "open B" and Alap: "A \ \S \ {}" and Blap: "B \ \S \ {}" and disj: "A \ B \ \S = {}" and cover: "\S \ A \ B" have disjs:"\s. s \ S \ A \ B \ s = {}" using disj by auto obtain sa where sa: "sa \ S" "A \ sa \ {}" using Alap by auto obtain sb where sb: "sb \ S" "B \ sb \ {}" using Blap by auto obtain x where x: "\s. s \ S \ x \ s" using ne by auto then have "x \ \S" using \sa \ S\ by blast then have "x \ A \ x \ B" using cover by auto then show False using cs [unfolded connected_def] by (metis A B IntI Sup_upper sa sb disjs x cover empty_iff subset_trans) qed lemma connected_Un: "connected s \ connected t \ s \ t \ {} \ connected (s \ t)" using connected_Union [of "{s,t}"] by auto lemma connected_diff_open_from_closed: assumes st: "s \ t" and tu: "t \ u" and s: "open s" and t: "closed t" and u: "connected u" and ts: "connected (t - s)" shows "connected(u - s)" proof (rule connectedI) fix A B assume AB: "open A" "open B" "A \ (u - s) \ {}" "B \ (u - s) \ {}" and disj: "A \ B \ (u - s) = {}" and cover: "u - s \ A \ B" then consider "A \ (t - s) = {}" | "B \ (t - s) = {}" using st ts tu connectedD [of "t-s" "A" "B"] by auto then show False proof cases case 1 then have "(A - t) \ (B \ s) \ u = {}" using disj st by auto moreover have "u \ (A - t) \ (B \ s)" using 1 cover by auto ultimately show False using connectedD [of u "A - t" "B \ s"] AB s t 1 u by auto next case 2 then have "(A \ s) \ (B - t) \ u = {}" using disj st by auto moreover have "u \ (A \ s) \ (B - t)" using 2 cover by auto ultimately show False using connectedD [of u "A \ s" "B - t"] AB s t 2 u by auto qed qed lemma connected_iff_const: fixes S :: "'a::topological_space set" shows "connected S \ (\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c))" proof safe fix P :: "'a \ bool" assume "connected S" "continuous_on S P" then have "\b. \A. open A \ A \ S = P -` {b} \ S" unfolding continuous_on_open_invariant by (simp add: open_discrete) from this[of True] this[of False] obtain t f where "open t" "open f" and *: "f \ S = P -` {False} \ S" "t \ S = P -` {True} \ S" by meson then have "t \ S = {} \ f \ S = {}" by (intro connectedD[OF \connected S\]) auto then show "\c. \s\S. P s = c" proof (rule disjE) assume "t \ S = {}" then show ?thesis unfolding * by (intro exI[of _ False]) auto next assume "f \ S = {}" then show ?thesis unfolding * by (intro exI[of _ True]) auto qed next assume P: "\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c)" show "connected S" proof (rule connectedI) fix A B assume *: "open A" "open B" "A \ S \ {}" "B \ S \ {}" "A \ B \ S = {}" "S \ A \ B" have "continuous_on S (\x. x \ A)" unfolding continuous_on_open_invariant proof safe fix C :: "bool set" have "C = UNIV \ C = {True} \ C = {False} \ C = {}" using subset_UNIV[of C] unfolding UNIV_bool by auto with * show "\T. open T \ T \ S = (\x. x \ A) -` C \ S" by (intro exI[of _ "(if True \ C then A else {}) \ (if False \ C then B else {})"]) auto qed from P[rule_format, OF this] obtain c where "\s. s \ S \ (s \ A) = c" by blast with * show False by (cases c) auto qed qed lemma connectedD_const: "connected S \ continuous_on S P \ \c. \s\S. P s = c" for P :: "'a::topological_space \ bool" by (auto simp: connected_iff_const) lemma connectedI_const: "(\P::'a::topological_space \ bool. continuous_on S P \ \c. \s\S. P s = c) \ connected S" by (auto simp: connected_iff_const) lemma connected_local_const: assumes "connected A" "a \ A" "b \ A" and *: "\a\A. eventually (\b. f a = f b) (at a within A)" shows "f a = f b" proof - obtain S where S: "\a. a \ A \ a \ S a" "\a. a \ A \ open (S a)" "\a x. a \ A \ x \ S a \ x \ A \ f a = f x" using * unfolding eventually_at_topological by metis let ?P = "\b\{b\A. f a = f b}. S b" and ?N = "\b\{b\A. f a \ f b}. S b" have "?P \ A = {} \ ?N \ A = {}" using \connected A\ S \a\A\ by (intro connectedD) (auto, metis) then show "f a = f b" proof assume "?N \ A = {}" then have "\x\A. f a = f x" using S(1) by auto with \b\A\ show ?thesis by auto next assume "?P \ A = {}" then show ?thesis using \a \ A\ S(1)[of a] by auto qed qed lemma (in linorder_topology) connectedD_interval: assumes "connected U" and xy: "x \ U" "y \ U" and "x \ z" "z \ y" shows "z \ U" proof - have eq: "{.. {z<..} = - {z}" by auto have "\ connected U" if "z \ U" "x < z" "z < y" using xy that apply (simp only: connected_def simp_thms) apply (rule_tac exI[of _ "{..< z}"]) apply (rule_tac exI[of _ "{z <..}"]) apply (auto simp add: eq) done with assms show "z \ U" by (metis less_le) qed lemma (in linorder_topology) not_in_connected_cases: assumes conn: "connected S" assumes nbdd: "x \ S" assumes ne: "S \ {}" obtains "bdd_above S" "\y. y \ S \ x \ y" | "bdd_below S" "\y. y \ S \ x \ y" proof - obtain s where "s \ S" using ne by blast { assume "s \ x" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \s \ S\ \y \ S\ \s \ x\ \x \ y\] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_above S" by (rule local.bdd_aboveI) note this wit } moreover { assume "x \ s" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \y \ S\ \s \ S\ \x \ y\ \s \ x\ ] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_below S" by (rule bdd_belowI) note this wit } ultimately show ?thesis by (meson le_cases that) qed lemma connected_continuous_image: assumes *: "continuous_on s f" and "connected s" shows "connected (f ` s)" proof (rule connectedI_const) fix P :: "'b \ bool" assume "continuous_on (f ` s) P" then have "continuous_on s (P \ f)" by (rule continuous_on_compose[OF *]) from connectedD_const[OF \connected s\ this] show "\c. \s\f ` s. P s = c" by auto qed section \Linear Continuum Topologies\ class linear_continuum_topology = linorder_topology + linear_continuum begin lemma Inf_notin_open: assumes A: "open A" and bnd: "\a\A. x < a" shows "Inf A \ A" proof assume "Inf A \ A" then obtain b where "b < Inf A" "{b <.. Inf A} \ A" using open_left[of A "Inf A" x] assms by auto with dense[of b "Inf A"] obtain c where "c < Inf A" "c \ A" by (auto simp: subset_eq) then show False using cInf_lower[OF \c \ A\] bnd by (metis not_le less_imp_le bdd_belowI) qed lemma Sup_notin_open: assumes A: "open A" and bnd: "\a\A. a < x" shows "Sup A \ A" proof assume "Sup A \ A" with assms obtain b where "Sup A < b" "{Sup A ..< b} \ A" using open_right[of A "Sup A" x] by auto with dense[of "Sup A" b] obtain c where "Sup A < c" "c \ A" by (auto simp: subset_eq) then show False using cSup_upper[OF \c \ A\] bnd by (metis less_imp_le not_le bdd_aboveI) qed end instance linear_continuum_topology \ perfect_space proof fix x :: 'a obtain y where "x < y \ y < x" using ex_gt_or_lt [of x] .. with Inf_notin_open[of "{x}" y] Sup_notin_open[of "{x}" y] show "\ open {x}" by auto qed lemma connectedI_interval: fixes U :: "'a :: linear_continuum_topology set" assumes *: "\x y z. x \ U \ y \ U \ x \ z \ z \ y \ z \ U" shows "connected U" proof (rule connectedI) { fix A B assume "open A" "open B" "A \ B \ U = {}" "U \ A \ B" fix x y assume "x < y" "x \ A" "y \ B" "x \ U" "y \ U" let ?z = "Inf (B \ {x <..})" have "x \ ?z" "?z \ y" using \y \ B\ \x < y\ by (auto intro: cInf_lower cInf_greatest) with \x \ U\ \y \ U\ have "?z \ U" by (rule *) moreover have "?z \ B \ {x <..}" using \open B\ by (intro Inf_notin_open) auto ultimately have "?z \ A" using \x \ ?z\ \A \ B \ U = {}\ \x \ A\ \U \ A \ B\ by auto have "\b\B. b \ A \ b \ U" if "?z < y" proof - obtain a where "?z < a" "{?z ..< a} \ A" using open_right[OF \open A\ \?z \ A\ \?z < y\] by auto moreover obtain b where "b \ B" "x < b" "b < min a y" using cInf_less_iff[of "B \ {x <..}" "min a y"] \?z < a\ \?z < y\ \x < y\ \y \ B\ by auto moreover have "?z \ b" using \b \ B\ \x < b\ by (intro cInf_lower) auto moreover have "b \ U" using \x \ ?z\ \?z \ b\ \b < min a y\ by (intro *[OF \x \ U\ \y \ U\]) (auto simp: less_imp_le) ultimately show ?thesis by (intro bexI[of _ b]) auto qed then have False using \?z \ y\ \?z \ A\ \y \ B\ \y \ U\ \A \ B \ U = {}\ unfolding le_less by blast } note not_disjoint = this fix A B assume AB: "open A" "open B" "U \ A \ B" "A \ B \ U = {}" moreover assume "A \ U \ {}" then obtain x where x: "x \ U" "x \ A" by auto moreover assume "B \ U \ {}" then obtain y where y: "y \ U" "y \ B" by auto moreover note not_disjoint[of B A y x] not_disjoint[of A B x y] ultimately show False by (cases x y rule: linorder_cases) auto qed lemma connected_iff_interval: "connected U \ (\x\U. \y\U. \z. x \ z \ z \ y \ z \ U)" for U :: "'a::linear_continuum_topology set" by (auto intro: connectedI_interval dest: connectedD_interval) lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)" by (simp add: connected_iff_interval) lemma connected_Ioi[simp]: "connected {a<..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Ici[simp]: "connected {a..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Iio[simp]: "connected {.. A" "b \ A" shows "{a <..< b} \ A" using connectedD_interval[OF assms] by (simp add: subset_eq Ball_def less_imp_le) lemma connected_contains_Icc: fixes A :: "'a::linorder_topology set" assumes "connected A" "a \ A" "b \ A" shows "{a..b} \ A" proof fix x assume "x \ {a..b}" then have "x = a \ x = b \ x \ {a<.. A" using assms connected_contains_Ioo[of A a b] by auto qed subsection \Intermediate Value Theorem\ lemma IVT': fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes y: "f a \ y" "y \ f b" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT2': fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" assumes y: "f b \ y" "y \ f a" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f a \ y \ y \ f b \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT') (auto intro: continuous_at_imp_continuous_on) lemma IVT2: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f b \ y \ y \ f a \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT2') (auto intro: continuous_at_imp_continuous_on) lemma continuous_inj_imp_mono: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes x: "a < x" "x < b" and cont: "continuous_on {a..b} f" and inj: "inj_on f {a..b}" shows "(f a < f x \ f x < f b) \ (f b < f x \ f x < f a)" proof - note I = inj_on_eq_iff[OF inj] { assume "f x < f a" "f x < f b" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f x < f s" using IVT'[of f x "min (f a) (f b)" b] IVT2'[of f x "min (f a) (f b)" a] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } moreover { assume "f a < f x" "f b < f x" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f s < f x" using IVT'[of f a "max (f a) (f b)" x] IVT2'[of f b "max (f a) (f b)" x] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } ultimately show ?thesis using I[of a x] I[of x b] x less_trans[OF x] by (auto simp add: le_less less_imp_neq neq_iff) qed lemma continuous_at_Sup_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "f (Sup S) \ (SUP s\S. f s)" proof cases assume "Sup S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_mono S \mono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cSUP_upper[OF \c \ S\ bdd_above_image_mono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \mono f\[THEN monoD] cSup_upper S) lemma continuous_at_Sup_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Sup S)" proof cases assume "Sup S \ S" then show ?thesis by (intro cINF_lower) (auto intro: bdd_below_image_antimono S \antimono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cINF_lower[OF bdd_below_image_antimono, of f S c] \c \ S\ show False by (auto simp: assms) qed qed qed (intro cINF_greatest \antimono f\[THEN antimonoD] cSup_upper S) lemma continuous_at_Inf_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Inf S)" proof cases assume "Inf S \ S" then show ?thesis by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S \mono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cINF_lower[OF bdd_below_image_mono[of f] \c \ S\] show False by (auto simp: assms) qed qed qed (intro cINF_greatest \mono f\[THEN monoD] cInf_lower \bdd_below S\ \S \ {}\) lemma continuous_at_Inf_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "f (Inf S) \ (SUP s\S. f s)" proof cases assume "Inf S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S \antimono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cSUP_upper[OF \c \ S\ bdd_above_image_antimono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \antimono f\[THEN antimonoD] cInf_lower S) subsection \Uniform spaces\ class uniformity = fixes uniformity :: "('a \ 'a) filter" begin abbreviation uniformity_on :: "'a set \ ('a \ 'a) filter" where "uniformity_on s \ inf uniformity (principal (s\s))" end lemma uniformity_Abort: "uniformity = Filter.abstract_filter (\u. Code.abort (STR ''uniformity is not executable'') (\u. uniformity))" by simp class open_uniformity = "open" + uniformity + assumes open_uniformity: "\U. open U \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" begin subclass topological_space by standard (force elim: eventually_mono eventually_elim2 simp: split_beta' open_uniformity)+ end class uniform_space = open_uniformity + assumes uniformity_refl: "eventually E uniformity \ E (x, x)" and uniformity_sym: "eventually E uniformity \ eventually (\(x, y). E (y, x)) uniformity" and uniformity_trans: "eventually E uniformity \ \D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" begin lemma uniformity_bot: "uniformity \ bot" using uniformity_refl by auto lemma uniformity_trans': "eventually E uniformity \ eventually (\((x, y), (y', z)). y = y' \ E (x, z)) (uniformity \\<^sub>F uniformity)" by (drule uniformity_trans) (auto simp add: eventually_prod_same) lemma uniformity_transE: assumes "eventually E uniformity" obtains D where "eventually D uniformity" "\x y z. D (x, y) \ D (y, z) \ E (x, z)" using uniformity_trans [OF assms] by auto lemma eventually_nhds_uniformity: "eventually P (nhds x) \ eventually (\(x', y). x' = x \ P y) uniformity" (is "_ \ ?N P x") unfolding eventually_nhds proof safe assume *: "?N P x" have "?N (?N P) x" if "?N P x" for x proof - from that obtain D where ev: "eventually D uniformity" and D: "D (a, b) \ D (b, c) \ case (a, c) of (x', y) \ x' = x \ P y" for a b c by (rule uniformity_transE) simp from ev show ?thesis by eventually_elim (insert ev D, force elim: eventually_mono split: prod.split) qed then have "open {x. ?N P x}" by (simp add: open_uniformity) then show "\S. open S \ x \ S \ (\x\S. P x)" by (intro exI[of _ "{x. ?N P x}"]) (auto dest: uniformity_refl simp: *) qed (force simp add: open_uniformity elim: eventually_mono) subsubsection \Totally bounded sets\ definition totally_bounded :: "'a set \ bool" where "totally_bounded S \ (\E. eventually E uniformity \ (\X. finite X \ (\s\S. \x\X. E (x, s))))" lemma totally_bounded_empty[iff]: "totally_bounded {}" by (auto simp add: totally_bounded_def) lemma totally_bounded_subset: "totally_bounded S \ T \ S \ totally_bounded T" by (fastforce simp add: totally_bounded_def) lemma totally_bounded_Union[intro]: assumes M: "finite M" "\S. S \ M \ totally_bounded S" shows "totally_bounded (\M)" unfolding totally_bounded_def proof safe fix E assume "eventually E uniformity" with M obtain X where "\S\M. finite (X S) \ (\s\S. \x\X S. E (x, s))" by (metis totally_bounded_def) with \finite M\ show "\X. finite X \ (\s\\M. \x\X. E (x, s))" by (intro exI[of _ "\S\M. X S"]) force qed subsubsection \Cauchy filter\ definition cauchy_filter :: "'a filter \ bool" where "cauchy_filter F \ F \\<^sub>F F \ uniformity" definition Cauchy :: "(nat \ 'a) \ bool" where Cauchy_uniform: "Cauchy X = cauchy_filter (filtermap X sequentially)" lemma Cauchy_uniform_iff: "Cauchy X \ (\P. eventually P uniformity \ (\N. \n\N. \m\N. P (X n, X m)))" unfolding Cauchy_uniform cauchy_filter_def le_filter_def eventually_prod_same eventually_filtermap eventually_sequentially proof safe let ?U = "\P. eventually P uniformity" { fix P assume "?U P" "\P. ?U P \ (\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y)))" then obtain Q N where "\n. n \ N \ Q (X n)" "\x y. Q x \ Q y \ P (x, y)" by metis then show "\N. \n\N. \m\N. P (X n, X m)" by blast next fix P assume "?U P" and P: "\P. ?U P \ (\N. \n\N. \m\N. P (X n, X m))" then obtain Q where "?U Q" and Q: "\x y z. Q (x, y) \ Q (y, z) \ P (x, z)" by (auto elim: uniformity_transE) then have "?U (\x. Q x \ (\(x, y). Q (y, x)) x)" unfolding eventually_conj_iff by (simp add: uniformity_sym) from P[rule_format, OF this] obtain N where N: "\n m. n \ N \ m \ N \ Q (X n, X m) \ Q (X m, X n)" by auto show "\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y))" proof (safe intro!: exI[of _ "\x. \n\N. Q (x, X n) \ Q (X n, x)"] exI[of _ N] N) fix x y assume "\n\N. Q (x, X n) \ Q (X n, x)" "\n\N. Q (y, X n) \ Q (X n, y)" then have "Q (x, X N)" "Q (X N, y)" by auto then show "P (x, y)" by (rule Q) qed } qed lemma nhds_imp_cauchy_filter: assumes *: "F \ nhds x" shows "cauchy_filter F" proof - have "F \\<^sub>F F \ nhds x \\<^sub>F nhds x" by (intro prod_filter_mono *) also have "\ \ uniformity" unfolding le_filter_def eventually_nhds_uniformity eventually_prod_same proof safe fix P assume "eventually P uniformity" then obtain Ql where ev: "eventually Ql uniformity" and "Ql (x, y) \ Ql (y, z) \ P (x, z)" for x y z by (rule uniformity_transE) simp with ev[THEN uniformity_sym] show "\Q. eventually (\(x', y). x' = x \ Q y) uniformity \ (\x y. Q x \ Q y \ P (x, y))" by (rule_tac exI[of _ "\y. Ql (y, x) \ Ql (x, y)"]) (fastforce elim: eventually_elim2) qed finally show ?thesis by (simp add: cauchy_filter_def) qed lemma LIMSEQ_imp_Cauchy: "X \ x \ Cauchy X" unfolding Cauchy_uniform filterlim_def by (intro nhds_imp_cauchy_filter) lemma Cauchy_subseq_Cauchy: assumes "Cauchy X" "strict_mono f" shows "Cauchy (X \ f)" unfolding Cauchy_uniform comp_def filtermap_filtermap[symmetric] cauchy_filter_def by (rule order_trans[OF _ \Cauchy X\[unfolded Cauchy_uniform cauchy_filter_def]]) (intro prod_filter_mono filtermap_mono filterlim_subseq[OF \strict_mono f\, unfolded filterlim_def]) lemma convergent_Cauchy: "convergent X \ Cauchy X" unfolding convergent_def by (erule exE, erule LIMSEQ_imp_Cauchy) definition complete :: "'a set \ bool" where complete_uniform: "complete S \ (\F \ principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x))" +lemma (in uniform_space) cauchy_filter_complete_converges: + assumes "cauchy_filter F" "complete A" "F \ principal A" "F \ bot" + shows "\c. F \ nhds c" + using assms unfolding complete_uniform by blast + end subsubsection \Uniformly continuous functions\ definition uniformly_continuous_on :: "'a set \ ('a::uniform_space \ 'b::uniform_space) \ bool" where uniformly_continuous_on_uniformity: "uniformly_continuous_on s f \ (LIM (x, y) (uniformity_on s). (f x, f y) :> uniformity)" lemma uniformly_continuous_onD: "uniformly_continuous_on s f \ eventually E uniformity \ eventually (\(x, y). x \ s \ y \ s \ E (f x, f y)) uniformity" by (simp add: uniformly_continuous_on_uniformity filterlim_iff eventually_inf_principal split_beta' mem_Times_iff imp_conjL) lemma uniformly_continuous_on_const[continuous_intros]: "uniformly_continuous_on s (\x. c)" by (auto simp: uniformly_continuous_on_uniformity filterlim_iff uniformity_refl) lemma uniformly_continuous_on_id[continuous_intros]: "uniformly_continuous_on s (\x. x)" by (auto simp: uniformly_continuous_on_uniformity filterlim_def) lemma uniformly_continuous_on_compose[continuous_intros]: "uniformly_continuous_on s g \ uniformly_continuous_on (g`s) f \ uniformly_continuous_on s (\x. f (g x))" using filterlim_compose[of "\(x, y). (f x, f y)" uniformity "uniformity_on (g`s)" "\(x, y). (g x, g y)" "uniformity_on s"] by (simp add: split_beta' uniformly_continuous_on_uniformity filterlim_inf filterlim_principal eventually_inf_principal mem_Times_iff) lemma uniformly_continuous_imp_continuous: assumes f: "uniformly_continuous_on s f" shows "continuous_on s f" by (auto simp: filterlim_iff eventually_at_filter eventually_nhds_uniformity continuous_on_def elim: eventually_mono dest!: uniformly_continuous_onD[OF f]) section \Product Topology\ subsection \Product is a topological space\ instantiation prod :: (topological_space, topological_space) topological_space begin definition open_prod_def[code del]: "open (S :: ('a \ 'b) set) \ (\x\S. \A B. open A \ open B \ x \ A \ B \ A \ B \ S)" lemma open_prod_elim: assumes "open S" and "x \ S" obtains A B where "open A" and "open B" and "x \ A \ B" and "A \ B \ S" using assms unfolding open_prod_def by fast lemma open_prod_intro: assumes "\x. x \ S \ \A B. open A \ open B \ x \ A \ B \ A \ B \ S" shows "open S" using assms unfolding open_prod_def by fast instance proof show "open (UNIV :: ('a \ 'b) set)" unfolding open_prod_def by auto next fix S T :: "('a \ 'b) set" assume "open S" "open T" show "open (S \ T)" proof (rule open_prod_intro) fix x assume x: "x \ S \ T" from x have "x \ S" by simp obtain Sa Sb where A: "open Sa" "open Sb" "x \ Sa \ Sb" "Sa \ Sb \ S" using \open S\ and \x \ S\ by (rule open_prod_elim) from x have "x \ T" by simp obtain Ta Tb where B: "open Ta" "open Tb" "x \ Ta \ Tb" "Ta \ Tb \ T" using \open T\ and \x \ T\ by (rule open_prod_elim) let ?A = "Sa \ Ta" and ?B = "Sb \ Tb" have "open ?A \ open ?B \ x \ ?A \ ?B \ ?A \ ?B \ S \ T" using A B by (auto simp add: open_Int) then show "\A B. open A \ open B \ x \ A \ B \ A \ B \ S \ T" by fast qed next fix K :: "('a \ 'b) set set" assume "\S\K. open S" then show "open (\K)" unfolding open_prod_def by fast qed end declare [[code abort: "open :: ('a::topological_space \ 'b::topological_space) set \ bool"]] lemma open_Times: "open S \ open T \ open (S \ T)" unfolding open_prod_def by auto lemma fst_vimage_eq_Times: "fst -` S = S \ UNIV" by auto lemma snd_vimage_eq_Times: "snd -` S = UNIV \ S" by auto lemma open_vimage_fst: "open S \ open (fst -` S)" by (simp add: fst_vimage_eq_Times open_Times) lemma open_vimage_snd: "open S \ open (snd -` S)" by (simp add: snd_vimage_eq_Times open_Times) lemma closed_vimage_fst: "closed S \ closed (fst -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_fst) lemma closed_vimage_snd: "closed S \ closed (snd -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_snd) lemma closed_Times: "closed S \ closed T \ closed (S \ T)" proof - have "S \ T = (fst -` S) \ (snd -` T)" by auto then show "closed S \ closed T \ closed (S \ T)" by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) qed lemma subset_fst_imageI: "A \ B \ S \ y \ B \ A \ fst ` S" unfolding image_def subset_eq by force lemma subset_snd_imageI: "A \ B \ S \ x \ A \ B \ snd ` S" unfolding image_def subset_eq by force lemma open_image_fst: assumes "open S" shows "open (fst ` S)" proof (rule openI) fix x assume "x \ fst ` S" then obtain y where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \y \ B\ have "A \ fst ` S" by (rule subset_fst_imageI) with \open A\ \x \ A\ have "open A \ x \ A \ A \ fst ` S" by simp then show "\T. open T \ x \ T \ T \ fst ` S" .. qed lemma open_image_snd: assumes "open S" shows "open (snd ` S)" proof (rule openI) fix y assume "y \ snd ` S" then obtain x where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \x \ A\ have "B \ snd ` S" by (rule subset_snd_imageI) with \open B\ \y \ B\ have "open B \ y \ B \ B \ snd ` S" by simp then show "\T. open T \ y \ T \ T \ snd ` S" .. qed lemma nhds_prod: "nhds (a, b) = nhds a \\<^sub>F nhds b" unfolding nhds_def proof (subst prod_filter_INF, auto intro!: antisym INF_greatest simp: principal_prod_principal) fix S T assume "open S" "a \ S" "open T" "b \ T" then show "(INF x \ {S. open S \ (a, b) \ S}. principal x) \ principal (S \ T)" by (intro INF_lower) (auto intro!: open_Times) next fix S' assume "open S'" "(a, b) \ S'" then obtain S T where "open S" "a \ S" "open T" "b \ T" "S \ T \ S'" by (auto elim: open_prod_elim) then show "(INF x \ {S. open S \ a \ S}. INF y \ {S. open S \ b \ S}. principal (x \ y)) \ principal S'" by (auto intro!: INF_lower2) qed subsubsection \Continuity of operations\ lemma tendsto_fst [tendsto_intros]: assumes "(f \ a) F" shows "((\x. fst (f x)) \ fst a) F" proof (rule topological_tendstoI) fix S assume "open S" and "fst a \ S" then have "open (fst -` S)" and "a \ fst -` S" by (simp_all add: open_vimage_fst) with assms have "eventually (\x. f x \ fst -` S) F" by (rule topological_tendstoD) then show "eventually (\x. fst (f x) \ S) F" by simp qed lemma tendsto_snd [tendsto_intros]: assumes "(f \ a) F" shows "((\x. snd (f x)) \ snd a) F" proof (rule topological_tendstoI) fix S assume "open S" and "snd a \ S" then have "open (snd -` S)" and "a \ snd -` S" by (simp_all add: open_vimage_snd) with assms have "eventually (\x. f x \ snd -` S) F" by (rule topological_tendstoD) then show "eventually (\x. snd (f x) \ S) F" by simp qed lemma tendsto_Pair [tendsto_intros]: assumes "(f \ a) F" and "(g \ b) F" shows "((\x. (f x, g x)) \ (a, b)) F" unfolding nhds_prod using assms by (rule filterlim_Pair) lemma continuous_fst[continuous_intros]: "continuous F f \ continuous F (\x. fst (f x))" unfolding continuous_def by (rule tendsto_fst) lemma continuous_snd[continuous_intros]: "continuous F f \ continuous F (\x. snd (f x))" unfolding continuous_def by (rule tendsto_snd) lemma continuous_Pair[continuous_intros]: "continuous F f \ continuous F g \ continuous F (\x. (f x, g x))" unfolding continuous_def by (rule tendsto_Pair) lemma continuous_on_fst[continuous_intros]: "continuous_on s f \ continuous_on s (\x. fst (f x))" unfolding continuous_on_def by (auto intro: tendsto_fst) lemma continuous_on_snd[continuous_intros]: "continuous_on s f \ continuous_on s (\x. snd (f x))" unfolding continuous_on_def by (auto intro: tendsto_snd) lemma continuous_on_Pair[continuous_intros]: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. (f x, g x))" unfolding continuous_on_def by (auto intro: tendsto_Pair) lemma continuous_on_swap[continuous_intros]: "continuous_on A prod.swap" by (simp add: prod.swap_def continuous_on_fst continuous_on_snd continuous_on_Pair continuous_on_id) lemma continuous_on_swap_args: assumes "continuous_on (A\B) (\(x,y). d x y)" shows "continuous_on (B\A) (\(x,y). d y x)" proof - have "(\(x,y). d y x) = (\(x,y). d x y) \ prod.swap" by force then show ?thesis by (metis assms continuous_on_compose continuous_on_swap product_swap) qed lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" by (fact continuous_fst) lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" by (fact continuous_snd) lemma isCont_Pair [simp]: "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" by (fact continuous_Pair) lemma continuous_on_compose_Pair: assumes f: "continuous_on (Sigma A B) (\(a, b). f a b)" assumes g: "continuous_on C g" assumes h: "continuous_on C h" assumes subset: "\c. c \ C \ g c \ A" "\c. c \ C \ h c \ B (g c)" shows "continuous_on C (\c. f (g c) (h c))" using continuous_on_compose2[OF f continuous_on_Pair[OF g h]] subset by auto subsubsection \Connectedness of products\ proposition connected_Times: assumes S: "connected S" and T: "connected T" shows "connected (S \ T)" proof (rule connectedI_const) fix P::"'a \ 'b \ bool" assume P[THEN continuous_on_compose2, continuous_intros]: "continuous_on (S \ T) P" have "continuous_on S (\s. P (s, t))" if "t \ T" for t by (auto intro!: continuous_intros that) from connectedD_const[OF S this] obtain c1 where c1: "\s t. t \ T \ s \ S \ P (s, t) = c1 t" by metis moreover have "continuous_on T (\t. P (s, t))" if "s \ S" for s by (auto intro!: continuous_intros that) from connectedD_const[OF T this] obtain c2 where "\s t. t \ T \ s \ S \ P (s, t) = c2 s" by metis ultimately show "\c. \s\S \ T. P s = c" by auto qed corollary connected_Times_eq [simp]: "connected (S \ T) \ S = {} \ T = {} \ connected S \ connected T" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof cases assume "S \ {} \ T \ {}" moreover have "connected (fst ` (S \ T))" "connected (snd ` (S \ T))" using continuous_on_fst continuous_on_snd continuous_on_id by (blast intro: connected_continuous_image [OF _ L])+ ultimately show ?thesis by auto qed auto qed (auto simp: connected_Times) subsubsection \Separation axioms\ instance prod :: (t0_space, t0_space) t0_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ (x \ U) \ (y \ U)" by (fast dest: t0_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t1_space, t1_space) t1_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ x \ U \ y \ U" by (fast dest: t1_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t2_space, t2_space) t2_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (fast dest: hausdorff elim: open_vimage_fst open_vimage_snd) qed lemma isCont_swap[continuous_intros]: "isCont prod.swap a" using continuous_on_eq_continuous_within continuous_on_swap by blast lemma open_diagonal_complement: "open {(x,y) |x y. x \ (y::('a::t2_space))}" proof - have "open {(x, y). x \ (y::'a)}" unfolding split_def by (intro open_Collect_neq continuous_intros) also have "{(x, y). x \ (y::'a)} = {(x, y) |x y. x \ (y::'a)}" by auto finally show ?thesis . qed lemma closed_diagonal: "closed {y. \ x::('a::t2_space). y = (x,x)}" proof - have "{y. \ x::'a. y = (x,x)} = UNIV - {(x,y) | x y. x \ y}" by auto then show ?thesis using open_diagonal_complement closed_Diff by auto qed lemma open_superdiagonal: "open {(x,y) | x y. x > (y::'a::{linorder_topology})}" proof - have "open {(x, y). x > (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x > (y::'a)} = {(x, y) |x y. x > (y::'a)}" by auto finally show ?thesis . qed lemma closed_subdiagonal: "closed {(x,y) | x y. x \ (y::'a::{linorder_topology})}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x > (y::'a)}" by auto then show ?thesis using open_superdiagonal closed_Diff by auto qed lemma open_subdiagonal: "open {(x,y) | x y. x < (y::'a::{linorder_topology})}" proof - have "open {(x, y). x < (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x < (y::'a)} = {(x, y) |x y. x < (y::'a)}" by auto finally show ?thesis . qed lemma closed_superdiagonal: "closed {(x,y) | x y. x \ (y::('a::{linorder_topology}))}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x < y}" by auto then show ?thesis using open_subdiagonal closed_Diff by auto qed end