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,7591 +1,7592 @@ (* 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 integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide: fixes f :: "_ \ 'b :: real_normed_field" assumes "f integrable_on s" shows "(\x. f x / of_real c) integrable_on s" by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by (rule integral_unique) auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "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 by (metis Diff_iff Un_iff indicator_def that) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" 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 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] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False 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) 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_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] thm derf_exp 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) 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 has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) subsection \More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" 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_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "finite \" and int: "\S. S \ \ \ (f has_integral (i S)) S" and neg: "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - let ?\ = "((\(a,b). a \ b) ` {(a,b). a \ \ \ b \ {y. y \ \ \ a \ y}})" have "((\x. if x \ \\ then f x else 0) has_integral sum i \) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (\ \ \)" by (simp add: \) moreover have "{(a, b). a \ \ \ b \ {y \ \. a \ y}} \ \ \ \" by auto ultimately show "finite ?\" by (blast intro: finite_subset[of _ "\ \ \"]) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ \\ then f x else 0) = (\A\\. if x \ A then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix S assume "S \ \" "x \ S" moreover then have "\b\\. x \ b \ b = S" using that by blast ultimately show "f x = (\A\\. if x \ A then f x else 0)" by (simp add: sum.delta[OF \]) qed next show "((\x. \A\\. if x \ A then f x else 0) has_integral (\A\\. i A)) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum [OF \]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" 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, hide_lams) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "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 + 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 by blast + 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 (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" 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] 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 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]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (insert a, simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" { fix k :: nat have "(f k has_integral F k) {0..c}" proof (cases "inverse (of_nat (Suc k)) \ c") case True { fix x assume x: "x \ inverse (1 + real k)" have "0 < inverse (1 + real k)" by simp also note x finally have "x > 0" . } note x = this hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const simp: has_field_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed } note has_integral_f = this have integral_f: "integral {0..c} (f k) = F k" for k using has_integral_f[of k] by (rule integral_unique) have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) also note x finally have "inverse (real (Suc (Suc k))) \ x" . } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" by (rule has_integral_on_superset) (insert False, simp_all add: f_def) with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (insert assms, simp add: integral_f F_def field_split_simps) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) from conjunct2[OF *] and this have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) with conjunct1[OF *] show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Improper_Integral.thy b/src/HOL/Analysis/Improper_Integral.thy --- a/src/HOL/Analysis/Improper_Integral.thy +++ b/src/HOL/Analysis/Improper_Integral.thy @@ -1,2268 +1,2273 @@ (* Title: HOL/Analysis/Improper_Integral.thy Author: LC Paulson (ported from HOL Light) *) section \Continuity of the indefinite integral; improper integral theorem\ theory "Improper_Integral" imports Equivalence_Lebesgue_Henstock_Integration begin subsection \Equiintegrability\ text\The definition here only really makes sense for an elementary set. We just use compact intervals in applications below.\ definition\<^marker>\tag important\ equiintegrable_on (infixr "equiintegrable'_on" 46) where "F equiintegrable_on I \ (\f \ F. f integrable_on I) \ (\e > 0. \\. gauge \ \ (\f \. f \ F \ \ tagged_division_of I \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < e))" lemma equiintegrable_on_integrable: "\F equiintegrable_on I; f \ F\ \ f integrable_on I" using equiintegrable_on_def by metis lemma equiintegrable_on_sing [simp]: "{f} equiintegrable_on cbox a b \ f integrable_on cbox a b" by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def) lemma equiintegrable_on_subset: "\F equiintegrable_on I; G \ F\ \ G equiintegrable_on I" unfolding equiintegrable_on_def Ball_def by (erule conj_forward imp_forward all_forward ex_forward | blast)+ lemma equiintegrable_on_Un: assumes "F equiintegrable_on I" "G equiintegrable_on I" shows "(F \ G) equiintegrable_on I" unfolding equiintegrable_on_def proof (intro conjI impI allI) show "\f\F \ G. f integrable_on I" using assms unfolding equiintegrable_on_def by blast show "\\. gauge \ \ (\f \. f \ F \ G \ \ tagged_division_of I \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \)" if "\ > 0" for \ proof - obtain \1 where "gauge \1" and \1: "\f \. f \ F \ \ tagged_division_of I \ \1 fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \" using assms \\ > 0\ unfolding equiintegrable_on_def by auto obtain \2 where "gauge \2" and \2: "\f \. f \ G \ \ tagged_division_of I \ \2 fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \" using assms \\ > 0\ unfolding equiintegrable_on_def by auto have "gauge (\x. \1 x \ \2 x)" using \gauge \1\ \gauge \2\ by blast moreover have "\f \. f \ F \ G \ \ tagged_division_of I \ (\x. \1 x \ \2 x) fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \" using \1 \2 by (auto simp: fine_Int) ultimately show ?thesis by (intro exI conjI) assumption+ qed qed lemma equiintegrable_on_insert: assumes "f integrable_on cbox a b" "F equiintegrable_on cbox a b" shows "(insert f F) equiintegrable_on cbox a b" by (metis assms equiintegrable_on_Un equiintegrable_on_sing insert_is_Un) lemma equiintegrable_cmul: assumes F: "F equiintegrable_on I" shows "(\c \ {-k..k}. \f \ F. {(\x. c *\<^sub>R f x)}) equiintegrable_on I" unfolding equiintegrable_on_def proof (intro conjI impI allI ballI) show "f integrable_on I" if "f \ (\c\{- k..k}. \f\F. {\x. c *\<^sub>R f x})" for f :: "'a \ 'b" using that assms equiintegrable_on_integrable integrable_cmul by blast show "\\. gauge \ \ (\f \. f \ (\c\{- k..k}. \f\F. {\x. c *\<^sub>R f x}) \ \ tagged_division_of I \ \ fine \ \ norm ((\(x, K)\\. content K *\<^sub>R f x) - integral I f) < \)" if "\ > 0" for \ proof - obtain \ where "gauge \" and \: "\f \. \f \ F; \ tagged_division_of I; \ fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \ / (\k\ + 1)" using assms \\ > 0\ unfolding equiintegrable_on_def by (metis add.commute add.right_neutral add_strict_mono divide_pos_pos norm_eq_zero real_norm_def zero_less_norm_iff zero_less_one) moreover have "norm ((\(x, K)\\. content K *\<^sub>R c *\<^sub>R (f x)) - integral I (\x. c *\<^sub>R f x)) < \" if c: "c \ {- k..k}" and "f \ F" "\ tagged_division_of I" "\ fine \" for \ c f proof - have "norm ((\x\\. case x of (x, K) \ content K *\<^sub>R c *\<^sub>R f x) - integral I (\x. c *\<^sub>R f x)) = \c\ * norm ((\x\\. case x of (x, K) \ content K *\<^sub>R f x) - integral I f)" by (simp add: algebra_simps scale_sum_right case_prod_unfold flip: norm_scaleR) also have "\ \ (\k\ + 1) * norm ((\x\\. case x of (x, K) \ content K *\<^sub>R f x) - integral I f)" using c by (auto simp: mult_right_mono) also have "\ < (\k\ + 1) * (\ / (\k\ + 1))" by (rule mult_strict_left_mono) (use \ less_eq_real_def that in auto) also have "\ = \" by auto finally show ?thesis . qed ultimately show ?thesis by (rule_tac x="\" in exI) auto qed qed lemma equiintegrable_add: assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I" shows "(\f \ F. \g \ G. {(\x. f x + g x)}) equiintegrable_on I" unfolding equiintegrable_on_def proof (intro conjI impI allI ballI) show "f integrable_on I" if "f \ (\f\F. \g\G. {\x. f x + g x})" for f using that equiintegrable_on_integrable assms by (auto intro: integrable_add) show "\\. gauge \ \ (\f \. f \ (\f\F. \g\G. {\x. f x + g x}) \ \ tagged_division_of I \ \ fine \ \ norm ((\(x, K)\\. content K *\<^sub>R f x) - integral I f) < \)" if "\ > 0" for \ proof - obtain \1 where "gauge \1" and \1: "\f \. \f \ F; \ tagged_division_of I; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) < \/2" using assms \\ > 0\ unfolding equiintegrable_on_def by (meson half_gt_zero_iff) obtain \2 where "gauge \2" and \2: "\g \. \g \ G; \ tagged_division_of I; \2 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R g x) - integral I g) < \/2" using assms \\ > 0\ unfolding equiintegrable_on_def by (meson half_gt_zero_iff) have "gauge (\x. \1 x \ \2 x)" using \gauge \1\ \gauge \2\ by blast moreover have "norm ((\(x,K) \ \. content K *\<^sub>R h x) - integral I h) < \" if h: "h \ (\f\F. \g\G. {\x. f x + g x})" and \: "\ tagged_division_of I" and fine: "(\x. \1 x \ \2 x) fine \" for h \ proof - obtain f g where "f \ F" "g \ G" and heq: "h = (\x. f x + g x)" using h by blast then have int: "f integrable_on I" "g integrable_on I" using F G equiintegrable_on_def by blast+ have "norm ((\(x,K) \ \. content K *\<^sub>R h x) - integral I h) = norm ((\(x,K) \ \. content K *\<^sub>R f x + content K *\<^sub>R g x) - (integral I f + integral I g))" by (simp add: heq algebra_simps integral_add int) also have "\ = norm (((\(x,K) \ \. content K *\<^sub>R f x) - integral I f + (\(x,K) \ \. content K *\<^sub>R g x) - integral I g))" by (simp add: sum.distrib algebra_simps case_prod_unfold) also have "\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral I f) + norm ((\(x,K) \ \. content K *\<^sub>R g x) - integral I g)" by (metis (mono_tags) add_diff_eq norm_triangle_ineq) also have "\ < \/2 + \/2" using \1 [OF \f \ F\ \] \2 [OF \g \ G\ \] fine by (simp add: fine_Int) finally show ?thesis by simp qed ultimately show ?thesis by meson qed qed lemma equiintegrable_minus: assumes "F equiintegrable_on I" shows "(\f \ F. {(\x. - f x)}) equiintegrable_on I" by (force intro: equiintegrable_on_subset [OF equiintegrable_cmul [OF assms, of 1]]) lemma equiintegrable_diff: assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I" shows "(\f \ F. \g \ G. {(\x. f x - g x)}) equiintegrable_on I" by (rule equiintegrable_on_subset [OF equiintegrable_add [OF F equiintegrable_minus [OF G]]]) auto lemma equiintegrable_sum: fixes F :: "('a::euclidean_space \ 'b::euclidean_space) set" assumes "F equiintegrable_on cbox a b" shows "(\I \ Collect finite. \c \ {c. (\i \ I. c i \ 0) \ sum c I = 1}. \f \ I \ F. {(\x. sum (\i::'j. c i *\<^sub>R f i x) I)}) equiintegrable_on cbox a b" (is "?G equiintegrable_on _") unfolding equiintegrable_on_def proof (intro conjI impI allI ballI) show "f integrable_on cbox a b" if "f \ ?G" for f using that assms by (auto simp: equiintegrable_on_def intro!: integrable_sum integrable_cmul) show "\\. gauge \ \ (\g \. g \ ?G \ \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g x) - integral (cbox a b) g) < \)" if "\ > 0" for \ proof - obtain \ where "gauge \" and \: "\f \. \f \ F; \ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral (cbox a b) f) < \ / 2" using assms \\ > 0\ unfolding equiintegrable_on_def by (meson half_gt_zero_iff) moreover have "norm ((\(x,K) \ \. content K *\<^sub>R g x) - integral (cbox a b) g) < \" if g: "g \ ?G" and \: "\ tagged_division_of cbox a b" and fine: "\ fine \" for \ g proof - obtain I c f where "finite I" and 0: "\i::'j. i \ I \ 0 \ c i" and 1: "sum c I = 1" and f: "f \ I \ F" and geq: "g = (\x. \i\I. c i *\<^sub>R f i x)" using g by auto have fi_int: "f i integrable_on cbox a b" if "i \ I" for i by (metis Pi_iff assms equiintegrable_on_def f that) have *: "integral (cbox a b) (\x. c i *\<^sub>R f i x) = (\(x, K)\\. integral K (\x. c i *\<^sub>R f i x))" if "i \ I" for i proof - have "f i integrable_on cbox a b" by (metis Pi_iff assms equiintegrable_on_def f that) then show ?thesis by (intro \ integrable_cmul integral_combine_tagged_division_topdown) qed have "finite \" using \ by blast have swap: "(\(x,K)\\. content K *\<^sub>R (\i\I. c i *\<^sub>R f i x)) = (\i\I. c i *\<^sub>R (\(x,K)\\. content K *\<^sub>R f i x))" by (simp add: scale_sum_right case_prod_unfold algebra_simps) (rule sum.swap) have "norm ((\(x, K)\\. content K *\<^sub>R g x) - integral (cbox a b) g) = norm ((\i\I. c i *\<^sub>R ((\(x,K)\\. content K *\<^sub>R f i x) - integral (cbox a b) (f i))))" unfolding geq swap by (simp add: scaleR_right.sum algebra_simps integral_sum fi_int integrable_cmul \finite I\ sum_subtractf flip: sum_diff) also have "\ \ (\i\I. c i * \ / 2)" proof (rule sum_norm_le) show "norm (c i *\<^sub>R ((\(xa, K)\\. content K *\<^sub>R f i xa) - integral (cbox a b) (f i))) \ c i * \ / 2" if "i \ I" for i proof - have "norm ((\(x, K)\\. content K *\<^sub>R f i x) - integral (cbox a b) (f i)) \ \/2" using \ [OF _ \ fine, of "f i"] funcset_mem [OF f] that by auto then show ?thesis using that by (auto simp: 0 mult.assoc intro: mult_left_mono) qed qed also have "\ < \" using 1 \\ > 0\ by (simp add: flip: sum_divide_distrib sum_distrib_right) finally show ?thesis . qed ultimately show ?thesis by (rule_tac x="\" in exI) auto qed qed corollary equiintegrable_sum_real: fixes F :: "(real \ 'b::euclidean_space) set" assumes "F equiintegrable_on {a..b}" shows "(\I \ Collect finite. \c \ {c. (\i \ I. c i \ 0) \ sum c I = 1}. \f \ I \ F. {(\x. sum (\i. c i *\<^sub>R f i x) I)}) equiintegrable_on {a..b}" using equiintegrable_sum [of F a b] assms by auto text\ Basic combining theorems for the interval of integration.\ lemma equiintegrable_on_null [simp]: "content(cbox a b) = 0 \ F equiintegrable_on cbox a b" - apply (auto simp: equiintegrable_on_def) - by (metis gauge_trivial norm_eq_zero sum_content_null) + unfolding equiintegrable_on_def + by (metis diff_zero gauge_trivial integrable_on_null integral_null norm_zero sum_content_null) text\ Main limit theorem for an equiintegrable sequence.\ theorem equiintegrable_limit: fixes g :: "'a :: euclidean_space \ 'b :: banach" assumes feq: "range f equiintegrable_on cbox a b" and to_g: "\x. x \ cbox a b \ (\n. f n x) \ g x" shows "g integrable_on cbox a b \ (\n. integral (cbox a b) (f n)) \ integral (cbox a b) g" proof - have "Cauchy (\n. integral(cbox a b) (f n))" proof (clarsimp simp add: Cauchy_def) fix e::real assume "0 < e" then have e3: "0 < e/3" by simp then obtain \ where "gauge \" and \: "\n \. \\ tagged_division_of cbox a b; \ fine \\ \ norm((\(x,K) \ \. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/3" using feq unfolding equiintegrable_on_def by (meson image_eqI iso_tuple_UNIV_I) obtain \ where \: "\ tagged_division_of (cbox a b)" and "\ fine \" "finite \" by (meson \gauge \\ fine_division_exists tagged_division_of_finite) with \ have \T: "\n. dist ((\(x,K)\\. content K *\<^sub>R f n x)) (integral (cbox a b) (f n)) < e/3" by (force simp: dist_norm) have "(\n. \(x,K)\\. content K *\<^sub>R f n x) \ (\(x,K)\\. content K *\<^sub>R g x)" using \ to_g by (auto intro!: tendsto_sum tendsto_scaleR) then have "Cauchy (\n. \(x,K)\\. content K *\<^sub>R f n x)" by (meson convergent_eq_Cauchy) with e3 obtain M where M: "\m n. \m\M; n\M\ \ dist (\(x,K)\\. content K *\<^sub>R f m x) (\(x,K)\\. content K *\<^sub>R f n x) < e/3" unfolding Cauchy_def by blast have "\m n. \m\M; n\M; dist (\(x,K)\\. content K *\<^sub>R f m x) (\(x,K)\\. content K *\<^sub>R f n x) < e/3\ \ dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e" by (metis \T dist_commute dist_triangle_third [OF _ _ \T]) then show "\M. \m\M. \n\M. dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e" using M by auto qed then obtain L where L: "(\n. integral (cbox a b) (f n)) \ L" by (meson convergent_eq_Cauchy) have "(g has_integral L) (cbox a b)" proof (clarsimp simp: has_integral) fix e::real assume "0 < e" then have e2: "0 < e/2" by simp then obtain \ where "gauge \" and \: "\n \. \\ tagged_division_of cbox a b; \ fine \\ \ norm((\(x,K)\\. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/2" using feq unfolding equiintegrable_on_def by (meson image_eqI iso_tuple_UNIV_I) moreover have "norm ((\(x,K)\\. content K *\<^sub>R g x) - L) < e" if "\ tagged_division_of cbox a b" "\ fine \" for \ proof - have "norm ((\(x,K)\\. content K *\<^sub>R g x) - L) \ e/2" proof (rule Lim_norm_ubound) show "(\n. (\(x,K)\\. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \ (\(x,K)\\. content K *\<^sub>R g x) - L" using to_g that L by (intro tendsto_diff tendsto_sum) (auto simp: tag_in_interval tendsto_scaleR) show "\\<^sub>F n in sequentially. norm ((\(x,K) \ \. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \ e/2" by (intro eventuallyI less_imp_le \ that) qed auto with \0 < e\ show ?thesis by linarith qed ultimately show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R g x) - L) < e)" by meson qed with L show ?thesis by (simp add: \(\n. integral (cbox a b) (f n)) \ L\ has_integral_integrable_integral) qed lemma equiintegrable_reflect: assumes "F equiintegrable_on cbox a b" shows "(\f. f \ uminus) ` F equiintegrable_on cbox (-b) (-a)" proof - - have "\\. gauge \ \ + have \
: "\\. gauge \ \ (\f \. f \ (\f. f \ uminus) ` F \ \ tagged_division_of cbox (- b) (- a) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral (cbox (- b) (- a)) f) < e)" if "gauge \" and \: "\f \. \f \ F; \ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral (cbox a b) f) < e" for e \ proof (intro exI, safe) show "gauge (\x. uminus ` \ (-x))" by (metis \gauge \\ gauge_reflect) show "norm ((\(x,K) \ \. content K *\<^sub>R (f \ uminus) x) - integral (cbox (- b) (- a)) (f \ uminus)) < e" if "f \ F" and tag: "\ tagged_division_of cbox (- b) (- a)" and fine: "(\x. uminus ` \ (- x)) fine \" for f \ proof - have 1: "(\(x,K). (- x, uminus ` K)) ` \ tagged_partial_division_of cbox a b" if "\ tagged_partial_division_of cbox (- b) (- a)" proof - have "- y \ cbox a b" if "\x K. (x,K) \ \ \ x \ K \ K \ cbox (- b) (- a) \ (\a b. K = cbox a b)" "(x, Y) \ \" "y \ Y" for x Y y proof - have "y \ uminus ` cbox a b" using that by auto then show "- y \ cbox a b" by force qed with that show ?thesis by (fastforce simp: tagged_partial_division_of_def interior_negations image_iff) qed have 2: "\K. (\x. (x,K) \ (\(x,K). (- x, uminus ` K)) ` \) \ x \ K" if "\{K. \x. (x,K) \ \} = cbox (- b) (- a)" "x \ cbox a b" for x proof - have xm: "x \ uminus ` \{A. \a. (a, A) \ \}" by (simp add: that) then obtain a X where "-x \ X" "(a, X) \ \" by auto then show ?thesis by (metis (no_types, lifting) add.inverse_inverse image_iff pair_imageI) qed have 3: "\x X y. \\ tagged_partial_division_of cbox (- b) (- a); (x, X) \ \; y \ X\ \ - y \ cbox a b" by (metis (no_types, lifting) equation_minus_iff imageE subsetD tagged_partial_division_ofD(3) uminus_interval_vector) have tag': "(\(x,K). (- x, uminus ` K)) ` \ tagged_division_of cbox a b" using tag by (auto simp: tagged_division_of_def dest: 1 2 3) have fine': "\ fine (\(x,K). (- x, uminus ` K)) ` \" using fine by (fastforce simp: fine_def) have inj: "inj_on (\(x,K). (- x, uminus ` K)) \" unfolding inj_on_def by force have eq: "content (uminus ` I) = content I" if I: "(x, I) \ \" and fnz: "f (- x) \ 0" for x I proof - obtain a b where "I = cbox a b" using tag I that by (force simp: tagged_division_of_def tagged_partial_division_of_def) then show ?thesis using content_image_affinity_cbox [of "-1" 0] by auto qed have "(\(x,K) \ (\(x,K). (- x, uminus ` K)) ` \. content K *\<^sub>R f x) = (\(x,K) \ \. content K *\<^sub>R f (- x))" - apply (simp add: sum.reindex [OF inj]) - apply (auto simp: eq intro!: sum.cong) - done + by (auto simp add: eq sum.reindex [OF inj] intro!: sum.cong) then show ?thesis using \ [OF \f \ F\ tag' fine'] integral_reflect by (metis (mono_tags, lifting) Henstock_Kurzweil_Integration.integral_cong comp_apply split_def sum.cong) qed qed - then show ?thesis + show ?thesis using assms apply (auto simp: equiintegrable_on_def) - apply (rule integrable_eq) - by auto + subgoal for f + by (metis (mono_tags, lifting) comp_apply integrable_eq integrable_reflect) + using \
by fastforce qed subsection\Subinterval restrictions for equiintegrable families\ text\First, some technical lemmas about minimizing a "flat" part of a sum over a division.\ lemma lemma0: assumes "i \ Basis" shows "content (cbox u v) / (interval_upperbound (cbox u v) \ i - interval_lowerbound (cbox u v) \ i) = (if content (cbox u v) = 0 then 0 else \j \ Basis - {i}. interval_upperbound (cbox u v) \ j - interval_lowerbound (cbox u v) \ j)" proof (cases "content (cbox u v) = 0") case True then show ?thesis by simp next case False then show ?thesis using prod.subset_diff [of "{i}" Basis] assms by (force simp: content_cbox_if divide_simps split: if_split_asm) qed lemma content_division_lemma1: assumes div: "\ division_of S" and S: "S \ cbox a b" and i: "i \ Basis" and mt: "\K. K \ \ \ content K \ 0" and disj: "(\K \ \. K \ {x. x \ i = a \ i} \ {}) \ (\K \ \. K \ {x. x \ i = b \ i} \ {})" shows "(b \ i - a \ i) * (\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ content(cbox a b)" (is "?lhs \ ?rhs") proof - have "finite \" using div by blast define extend where "extend \ \K. cbox (\j \ Basis. if j = i then (a \ i) *\<^sub>R i else (interval_lowerbound K \ j) *\<^sub>R j) (\j \ Basis. if j = i then (b \ i) *\<^sub>R i else (interval_upperbound K \ j) *\<^sub>R j)" have div_subset_cbox: "\K. K \ \ \ K \ cbox a b" using S div by auto have "\K. K \ \ \ K \ {}" using div by blast + have extend_cbox: "\K. K \ \ \ \a b. extend K = cbox a b" + using extend_def by blast have extend: "extend K \ {}" "extend K \ cbox a b" if K: "K \ \" for K proof - obtain u v where K: "K = cbox u v" "K \ {}" "K \ cbox a b" using K cbox_division_memE [OF _ div] by (meson div_subset_cbox) - with i show "extend K \ {}" "extend K \ cbox a b" - apply (auto simp: extend_def subset_box box_ne_empty) - by fastforce + with i show "extend K \ cbox a b" + by (auto simp: extend_def subset_box box_ne_empty) + have "a \ i \ b \ i" + using K by (metis bot.extremum_uniqueI box_ne_empty(1) i) + with K show "extend K \ {}" + by (simp add: extend_def i box_ne_empty) qed have int_extend_disjoint: "interior(extend K1) \ interior(extend K2) = {}" if K: "K1 \ \" "K2 \ \" "K1 \ K2" for K1 K2 proof - obtain u v where K1: "K1 = cbox u v" "K1 \ {}" "K1 \ cbox a b" using K cbox_division_memE [OF _ div] by (meson div_subset_cbox) obtain w z where K2: "K2 = cbox w z" "K2 \ {}" "K2 \ cbox a b" using K cbox_division_memE [OF _ div] by (meson div_subset_cbox) have cboxes: "cbox u v \ \" "cbox w z \ \" "cbox u v \ cbox w z" using K1 K2 that by auto with div have "interior (cbox u v) \ interior (cbox w z) = {}" by blast moreover have "\x. x \ box u v \ x \ box w z" if "x \ interior (extend K1)" "x \ interior (extend K2)" for x proof - have "a \ i < x \ i" "x \ i < b \ i" and ux: "\k. k \ Basis - {i} \ u \ k < x \ k" and xv: "\k. k \ Basis - {i} \ x \ k < v \ k" and wx: "\k. k \ Basis - {i} \ w \ k < x \ k" and xz: "\k. k \ Basis - {i} \ x \ k < z \ k" using that K1 K2 i by (auto simp: extend_def box_ne_empty mem_box) have "box u v \ {}" "box w z \ {}" using cboxes interior_cbox by (auto simp: content_eq_0_interior dest: mt) then obtain q s where q: "\k. k \ Basis \ w \ k < q \ k \ q \ k < z \ k" and s: "\k. k \ Basis \ u \ k < s \ k \ s \ k < v \ k" by (meson all_not_in_conv mem_box(1)) show ?thesis using disj proof assume "\K\\. K \ {x. x \ i = a \ i} \ {}" then have uva: "(cbox u v) \ {x. x \ i = a \ i} \ {}" and wza: "(cbox w z) \ {x. x \ i = a \ i} \ {}" using cboxes by (auto simp: content_eq_0_interior) then obtain r t where "r \ i = a \ i" and r: "\k. k \ Basis \ w \ k \ r \ k \ r \ k \ z \ k" and "t \ i = a \ i" and t: "\k. k \ Basis \ u \ k \ t \ k \ t \ k \ v \ k" by (fastforce simp: mem_box) have u: "u \ i < q \ i" using i K2(1) K2(3) \t \ i = a \ i\ q s t [OF i] by (force simp: subset_box) have w: "w \ i < s \ i" using i K1(1) K1(3) \r \ i = a \ i\ s r [OF i] by (force simp: subset_box) - let ?x = "(\j \ Basis. if j = i then min (q \ i) (s \ i) *\<^sub>R i else (x \ j) *\<^sub>R j)" + define \ where "\ \ (\j \ Basis. if j = i then min (q \ i) (s \ i) *\<^sub>R i else (x \ j) *\<^sub>R j)" + have [simp]: "\ \ j = (if j = i then min (q \ j) (s \ j) else x \ j)" if "j \ Basis" for j + unfolding \_def + by (intro sum_if_inner that \i \ Basis\) show ?thesis proof (intro exI conjI) - show "?x \ box u v" - using \i \ Basis\ s apply (clarsimp simp: mem_box) - apply (subst sum_if_inner; simp)+ - apply (fastforce simp: u ux xv) - done - show "?x \ box w z" - using \i \ Basis\ q apply (clarsimp simp: mem_box) - apply (subst sum_if_inner; simp)+ - apply (fastforce simp: w wx xz) - done + have "min (q \ i) (s \ i) < v \ i" + using i s by fastforce + with \i \ Basis\ s u ux xv + show "\ \ box u v" + by (force simp: mem_box) + have "min (q \ i) (s \ i) < z \ i" + using i q by force + with \i \ Basis\ q w wx xz + show "\ \ box w z" + by (force simp: mem_box) qed next assume "\K\\. K \ {x. x \ i = b \ i} \ {}" then have uva: "(cbox u v) \ {x. x \ i = b \ i} \ {}" and wza: "(cbox w z) \ {x. x \ i = b \ i} \ {}" using cboxes by (auto simp: content_eq_0_interior) then obtain r t where "r \ i = b \ i" and r: "\k. k \ Basis \ w \ k \ r \ k \ r \ k \ z \ k" and "t \ i = b \ i" and t: "\k. k \ Basis \ u \ k \ t \ k \ t \ k \ v \ k" by (fastforce simp: mem_box) have z: "s \ i < z \ i" using K1(1) K1(3) \r \ i = b \ i\ r [OF i] i s by (force simp: subset_box) have v: "q \ i < v \ i" using K2(1) K2(3) \t \ i = b \ i\ t [OF i] i q by (force simp: subset_box) - let ?x = "(\j \ Basis. if j = i then max (q \ i) (s \ i) *\<^sub>R i else (x \ j) *\<^sub>R j)" + define \ where "\ \ (\j \ Basis. if j = i then max (q \ i) (s \ i) *\<^sub>R i else (x \ j) *\<^sub>R j)" + have [simp]: "\ \ j = (if j = i then max (q \ j) (s \ j) else x \ j)" if "j \ Basis" for j + unfolding \_def + by (intro sum_if_inner that \i \ Basis\) show ?thesis proof (intro exI conjI) - show "?x \ box u v" - using \i \ Basis\ s apply (clarsimp simp: mem_box) - apply (subst sum_if_inner; simp)+ - apply (fastforce simp: v ux xv) - done - show "?x \ box w z" - using \i \ Basis\ q apply (clarsimp simp: mem_box) - apply (subst sum_if_inner; simp)+ - apply (fastforce simp: z wx xz) - done + show "\ \ box u v" + using \i \ Basis\ s by (force simp: mem_box ux v xv) + show "\ \ box w z" + using \i \ Basis\ q by (force simp: mem_box wx xz z) qed qed qed ultimately show ?thesis by auto qed - have "?lhs = (\K\\. (b \ i - a \ i) * content K / (interval_upperbound K \ i - interval_lowerbound K \ i))" - by (simp add: sum_distrib_left) + define interv_diff where "interv_diff \ \K. \i::'a. interval_upperbound K \ i - interval_lowerbound K \ i" + have "?lhs = (\K\\. (b \ i - a \ i) * content K / (interv_diff K i))" + by (simp add: sum_distrib_left interv_diff_def) also have "\ = sum (content \ extend) \" proof (rule sum.cong [OF refl]) fix K assume "K \ \" then obtain u v where K: "K = cbox u v" "cbox u v \ {}" "K \ cbox a b" using cbox_division_memE [OF _ div] div_subset_cbox by metis then have uv: "u \ i < v \ i" using mt [OF \K \ \\] \i \ Basis\ content_eq_0 by fastforce have "insert i (Basis \ -{i}) = Basis" using \i \ Basis\ by auto - then have "(b \ i - a \ i) * content K / (interval_upperbound K \ i - interval_lowerbound K \ i) - = (b \ i - a \ i) * (\i \ insert i (Basis \ -{i}). v \ i - u \ i) / (interval_upperbound (cbox u v) \ i - interval_lowerbound (cbox u v) \ i)" + then have "(b \ i - a \ i) * content K / (interv_diff K i) + = (b \ i - a \ i) * (\i \ insert i (Basis \ -{i}). v \ i - u \ i) / (interv_diff (cbox u v) i)" using K box_ne_empty(1) content_cbox by fastforce also have "... = (\x\Basis. if x = i then b \ x - a \ x else (interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \ x)" - using \i \ Basis\ K uv by (simp add: prod.If_cases) (simp add: algebra_simps) + using \i \ Basis\ K uv by (simp add: prod.If_cases interv_diff_def) (simp add: algebra_simps) also have "... = (\k\Basis. - (\j\Basis. if j = i then (b \ i - a \ i) *\<^sub>R i else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \ j) *\<^sub>R j) \ k)" + (\j\Basis. if j = i then (b \ i - a \ i) *\<^sub>R i + else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \ j) *\<^sub>R j) \ k)" using \i \ Basis\ by (subst prod.cong [OF refl sum_if_inner]; simp) also have "... = (\k\Basis. (\j\Basis. if j = i then (b \ i) *\<^sub>R i else (interval_upperbound (cbox u v) \ j) *\<^sub>R j) \ k - (\j\Basis. if j = i then (a \ i) *\<^sub>R i else (interval_lowerbound (cbox u v) \ j) *\<^sub>R j) \ k)" - apply (rule prod.cong [OF refl]) using \i \ Basis\ - apply (subst sum_if_inner; simp add: algebra_simps)+ - done + by (intro prod.cong [OF refl]) (subst sum_if_inner; simp add: algebra_simps)+ also have "... = (content \ extend) K" - using \i \ Basis\ K box_ne_empty - apply (simp add: extend_def) - apply (subst content_cbox, auto) - done - finally show "(b \ i - a \ i) * content K / (interval_upperbound K \ i - interval_lowerbound K \ i) - = (content \ extend) K" . + using \i \ Basis\ K box_ne_empty \K \ \\ extend(1) + by (auto simp add: extend_def content_cbox_if) + finally show "(b \ i - a \ i) * content K / (interv_diff K i) = (content \ extend) K" . qed also have "... = sum content (extend ` \)" proof - have "\K1 \ \; K2 \ \; K1 \ K2; extend K1 = extend K2\ \ content (extend K1) = 0" for K1 K2 using int_extend_disjoint [of K1 K2] extend_def by (simp add: content_eq_0_interior) then show ?thesis by (simp add: comm_monoid_add_class.sum.reindex_nontrivial [OF \finite \\]) qed also have "... \ ?rhs" proof (rule subadditive_content_division) show "extend ` \ division_of \ (extend ` \)" - using int_extend_disjoint apply (auto simp: division_of_def \finite \\ extend) - using extend_def apply blast - done + using int_extend_disjoint by (auto simp: division_of_def \finite \\ extend extend_cbox) show "\ (extend ` \) \ cbox a b" using extend by fastforce qed finally show ?thesis . qed proposition sum_content_area_over_thin_division: assumes div: "\ division_of S" and S: "S \ cbox a b" and i: "i \ Basis" and "a \ i \ c" "c \ b \ i" and nonmt: "\K. K \ \ \ K \ {x. x \ i = c} \ {}" shows "(b \ i - a \ i) * (\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ 2 * content(cbox a b)" proof (cases "content(cbox a b) = 0") case True have "(\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) = 0" using S div by (force intro!: sum.neutral content_0_subset [OF True]) then show ?thesis by (auto simp: True) next case False then have "content(cbox a b) > 0" using zero_less_measure_iff by blast then have "a \ i < b \ i" if "i \ Basis" for i using content_pos_lt_eq that by blast have "finite \" using div by blast define Dlec where "Dlec \ {L \ (\L. L \ {x. x \ i \ c}) ` \. content L \ 0}" define Dgec where "Dgec \ {L \ (\L. L \ {x. x \ i \ c}) ` \. content L \ 0}" define a' where "a' \ (\j\Basis. (if j = i then c else a \ j) *\<^sub>R j)" define b' where "b' \ (\j\Basis. (if j = i then c else b \ j) *\<^sub>R j)" + define interv_diff where "interv_diff \ \K. \i::'a. interval_upperbound K \ i - interval_lowerbound K \ i" have Dlec_cbox: "\K. K \ Dlec \ \a b. K = cbox a b" using interval_split [OF i] div by (fastforce simp: Dlec_def division_of_def) then have lec_is_cbox: "\content (L \ {x. x \ i \ c}) \ 0; L \ \\ \ \a b. L \ {x. x \ i \ c} = cbox a b" for L using Dlec_def by blast have Dgec_cbox: "\K. K \ Dgec \ \a b. K = cbox a b" using interval_split [OF i] div by (fastforce simp: Dgec_def division_of_def) then have gec_is_cbox: "\content (L \ {x. x \ i \ c}) \ 0; L \ \\ \ \a b. L \ {x. x \ i \ c} = cbox a b" for L using Dgec_def by blast - have "(b' \ i - a \ i) * (\K\Dlec. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ content(cbox a b')" + + have zero_left: "\x y. \x \ \; y \ \; x \ y; x \ {x. x \ i \ c} = y \ {x. x \ i \ c}\ + \ content (y \ {x. x \ i \ c}) = 0" + by (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior) + have zero_right: "\x y. \x \ \; y \ \; x \ y; x \ {x. c \ x \ i} = y \ {x. c \ x \ i}\ + \ content (y \ {x. c \ x \ i}) = 0" + by (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior) + + have "(b' \ i - a \ i) * (\K\Dlec. content K / interv_diff K i) \ content(cbox a b')" + unfolding interv_diff_def proof (rule content_division_lemma1) show "Dlec division_of \Dlec" unfolding division_of_def proof (intro conjI ballI Dlec_cbox) show "\K1 K2. \K1 \ Dlec; K2 \ Dlec\ \ K1 \ K2 \ interior K1 \ interior K2 = {}" by (clarsimp simp: Dlec_def) (use div in auto) qed (use \finite \\ Dlec_def in auto) show "\Dlec \ cbox a b'" using Dlec_def div S by (auto simp: b'_def division_of_def mem_box) show "(\K\Dlec. K \ {x. x \ i = a \ i} \ {}) \ (\K\Dlec. K \ {x. x \ i = b' \ i} \ {})" using nonmt by (fastforce simp: Dlec_def b'_def i) qed (use i Dlec_def in auto) moreover - have "(\K\Dlec. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) = - (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K)" - apply (subst sum.reindex_nontrivial [OF \finite \\, symmetric], simp) - apply (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior) - unfolding Dlec_def using \finite \\ apply (auto simp: sum.mono_neutral_left) - done + have "(\K\Dlec. content K / (interv_diff K i)) = (\K\(\K. K \ {x. x \ i \ c}) ` \. content K / interv_diff K i)" + unfolding Dlec_def using \finite \\ by (auto simp: sum.mono_neutral_left) + moreover have "... = + (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K)" + by (simp add: zero_left sum.reindex_nontrivial [OF \finite \\]) moreover have "(b' \ i - a \ i) = (c - a \ i)" by (simp add: b'_def i) ultimately - have lec: "(c - a \ i) * (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K) + have lec: "(c - a \ i) * (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K) \ content(cbox a b')" by simp - have "(b \ i - a' \ i) * (\K\Dgec. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ content(cbox a' b)" + have "(b \ i - a' \ i) * (\K\Dgec. content K / (interv_diff K i)) \ content(cbox a' b)" + unfolding interv_diff_def proof (rule content_division_lemma1) show "Dgec division_of \Dgec" unfolding division_of_def proof (intro conjI ballI Dgec_cbox) show "\K1 K2. \K1 \ Dgec; K2 \ Dgec\ \ K1 \ K2 \ interior K1 \ interior K2 = {}" by (clarsimp simp: Dgec_def) (use div in auto) qed (use \finite \\ Dgec_def in auto) show "\Dgec \ cbox a' b" using Dgec_def div S by (auto simp: a'_def division_of_def mem_box) show "(\K\Dgec. K \ {x. x \ i = a' \ i} \ {}) \ (\K\Dgec. K \ {x. x \ i = b \ i} \ {})" using nonmt by (fastforce simp: Dgec_def a'_def i) qed (use i Dgec_def in auto) moreover - have "(\K\Dgec. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) = - (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K)" - apply (subst sum.reindex_nontrivial [OF \finite \\, symmetric], simp) - apply (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior) - unfolding Dgec_def using \finite \\ apply (auto simp: sum.mono_neutral_left) - done + have "(\K\Dgec. content K / (interv_diff K i)) = (\K\(\K. K \ {x. c \ x \ i}) ` \. + content K / interv_diff K i)" + unfolding Dgec_def using \finite \\ by (auto simp: sum.mono_neutral_left) + moreover have "\ = + (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K)" + by (simp add: zero_right sum.reindex_nontrivial [OF \finite \\]) moreover have "(b \ i - a' \ i) = (b \ i - c)" by (simp add: a'_def i) ultimately - have gec: "(b \ i - c) * (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K) + have gec: "(b \ i - c) * (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K) \ content(cbox a' b)" by simp + show ?thesis proof (cases "c = a \ i \ c = b \ i") case True then show ?thesis proof assume c: "c = a \ i" - then have "a' = a" - apply (simp add: i a'_def cong: if_cong) + moreover + have "(\j\Basis. (if j = i then a \ i else a \ j) *\<^sub>R j) = a" using euclidean_representation [of a] sum.cong [OF refl, of Basis "\i. (a \ i) *\<^sub>R i"] by presburger + ultimately have "a' = a" + by (simp add: i a'_def cong: if_cong) then have "content (cbox a' b) \ 2 * content (cbox a b)" by simp moreover - have eq: "(\K\\. content (K \ {x. a \ i \ x \ i}) / - (interval_upperbound (K \ {x. a \ i \ x \ i}) \ i - interval_lowerbound (K \ {x. a \ i \ x \ i}) \ i)) - = (\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i))" - (is "sum ?f _ = sum ?g _") + have eq: "(\K\\. content (K \ {x. a \ i \ x \ i}) / interv_diff (K \ {x. a \ i \ x \ i}) i) + = (\K\\. content K / interv_diff K i)" + (is "sum ?f _ = sum ?g _") proof (rule sum.cong [OF refl]) fix K assume "K \ \" then have "a \ i \ x \ i" if "x \ K" for x by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that) then have "K \ {x. a \ i \ x \ i} = K" by blast then show "?f K = ?g K" by simp qed ultimately show ?thesis - using gec c eq by auto + using gec c eq interv_diff_def by auto next assume c: "c = b \ i" - then have "b' = b" - apply (simp add: i b'_def cong: if_cong) + moreover have "(\j\Basis. (if j = i then b \ i else b \ j) *\<^sub>R j) = b" using euclidean_representation [of b] sum.cong [OF refl, of Basis "\i. (b \ i) *\<^sub>R i"] by presburger + ultimately have "b' = b" + by (simp add: i b'_def cong: if_cong) then have "content (cbox a b') \ 2 * content (cbox a b)" by simp moreover - have eq: "(\K\\. content (K \ {x. x \ i \ b \ i}) / - (interval_upperbound (K \ {x. x \ i \ b \ i}) \ i - interval_lowerbound (K \ {x. x \ i \ b \ i}) \ i)) - = (\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i))" + have eq: "(\K\\. content (K \ {x. x \ i \ b \ i}) / interv_diff (K \ {x. x \ i \ b \ i}) i) + = (\K\\. content K / interv_diff K i)" (is "sum ?f _ = sum ?g _") proof (rule sum.cong [OF refl]) fix K assume "K \ \" then have "x \ i \ b \ i" if "x \ K" for x by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that) then have "K \ {x. x \ i \ b \ i} = K" by blast then show "?f K = ?g K" by simp qed ultimately show ?thesis - using lec c eq by auto + using lec c eq interv_diff_def by auto qed next case False have prod_if: "(\k\Basis \ - {i}. f k) = (\k\Basis. f k) / f i" if "f i \ (0::real)" for f - using that mk_disjoint_insert [OF i] - apply (clarsimp simp add: field_split_simps) - by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton) + proof - + have "f i * prod f (Basis \ - {i}) = prod f Basis" + using that mk_disjoint_insert [OF i] + by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton) + then show ?thesis + by (metis nonzero_mult_div_cancel_left that) + qed have abc: "a \ i < c" "c < b \ i" using False assms by auto - then have "(\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K) + then have "(\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K) \ content(cbox a b') / (c - a \ i)" - "(\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K) + "(\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K) \ content(cbox a' b) / (b \ i - c)" using lec gec by (simp_all add: field_split_simps) moreover - have "(\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) - \ (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K) + - (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K)" + have "(\K\\. content K / (interv_diff K i)) + \ (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K) + + (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K)" (is "?lhs \ ?rhs") proof - have "?lhs \ - (\K\\. ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K + - ((\K. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ ((\K. K \ {x. x \ i \ c}))) K)" + (\K\\. ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K + + ((\K. content K / (interv_diff K i)) \ ((\K. K \ {x. x \ i \ c}))) K)" (is "sum ?f _ \ sum ?g _") proof (rule sum_mono) fix K assume "K \ \" then obtain u v where uv: "K = cbox u v" using div by blast obtain u' v' where uv': "cbox u v \ {x. x \ i \ c} = cbox u v'" "cbox u v \ {x. c \ x \ i} = cbox u' v" "\k. k \ Basis \ u' \ k = (if k = i then max (u \ i) c else u \ k)" "\k. k \ Basis \ v' \ k = (if k = i then min (v \ i) c else v \ k)" using i by (auto simp: interval_split) have *: "\content (cbox u v') = 0; content (cbox u' v) = 0\ \ content (cbox u v) = 0" "content (cbox u' v) \ 0 \ content (cbox u v) \ 0" "content (cbox u v') \ 0 \ content (cbox u v) \ 0" using i uv uv' by (auto simp: content_eq_0 le_max_iff_disj min_le_iff_disj split: if_split_asm intro: order_trans) + have uniq: "\j. \j \ Basis; \ u \ j \ v \ j\ \ j = i" + by (metis \K \ \\ box_ne_empty(1) div division_of_def uv) show "?f K \ ?g K" - using i uv uv' apply (clarsimp simp add: lemma0 * intro!: prod_nonneg) - by (metis content_eq_0 le_less_linear order.strict_implies_order) + using i uv uv' by (auto simp add: interv_diff_def lemma0 dest: uniq * intro!: prod_nonneg) qed also have "... = ?rhs" by (simp add: sum.distrib) finally show ?thesis . qed moreover have "content (cbox a b') / (c - a \ i) = content (cbox a b) / (b \ i - a \ i)" using i abc apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff) apply (auto simp: if_distrib if_distrib [of "\f. f x" for x] prod.If_cases [of Basis "\x. x = i", simplified] prod_if field_simps) done moreover have "content (cbox a' b) / (b \ i - c) = content (cbox a b) / (b \ i - a \ i)" using i abc apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff) apply (auto simp: if_distrib prod.If_cases [of Basis "\x. x = i", simplified] prod_if field_simps) done ultimately - have "(\K\\. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) - \ 2 * content (cbox a b) / (b \ i - a \ i)" + have "(\K\\. content K / (interv_diff K i)) \ 2 * content (cbox a b) / (b \ i - a \ i)" by linarith then show ?thesis - using abc by (simp add: field_split_simps) + using abc interv_diff_def by (simp add: field_split_simps) qed qed - - proposition bounded_equiintegral_over_thin_tagged_partial_division: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes F: "F equiintegrable_on cbox a b" and f: "f \ F" and "0 < \" and norm_f: "\h x. \h \ F; x \ cbox a b\ \ norm(h x) \ norm(f x)" obtains \ where "gauge \" "\c i S h. \c \ cbox a b; i \ Basis; S tagged_partial_division_of cbox a b; \ fine S; h \ F; \x K. (x,K) \ S \ (K \ {x. x \ i = c \ i} \ {})\ \ (\(x,K) \ S. norm (integral K h)) < \" proof (cases "content(cbox a b) = 0") case True show ?thesis proof show "gauge (\x. ball x 1)" by (simp add: gauge_trivial) show "(\(x,K) \ S. norm (integral K h)) < \" if "S tagged_partial_division_of cbox a b" "(\x. ball x 1) fine S" for S and h:: "'a \ 'b" proof - have "(\(x,K) \ S. norm (integral K h)) = 0" using that True content_0_subset by (fastforce simp: tagged_partial_division_of_def intro: sum.neutral) with \0 < \\ show ?thesis by simp qed qed next case False then have contab_gt0: "content(cbox a b) > 0" by (simp add: zero_less_measure_iff) then have a_less_b: "\i. i \ Basis \ a\i < b\i" by (auto simp: content_pos_lt_eq) obtain \0 where "gauge \0" and \0: "\S h. \S tagged_partial_division_of cbox a b; \0 fine S; h \ F\ \ (\(x,K) \ S. norm (content K *\<^sub>R h x - integral K h)) < \/2" proof - obtain \ where "gauge \" and \: "\f \. \f \ F; \ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - integral (cbox a b) f) < \/(5 * (Suc DIM('b)))" proof - have e5: "\/(5 * (Suc DIM('b))) > 0" using \\ > 0\ by auto then show ?thesis using F that by (auto simp: equiintegrable_on_def) qed show ?thesis proof show "gauge \" by (rule \gauge \\) show "(\(x,K) \ S. norm (content K *\<^sub>R h x - integral K h)) < \/2" if "S tagged_partial_division_of cbox a b" "\ fine S" "h \ F" for S h proof - have "(\(x,K) \ S. norm (content K *\<^sub>R h x - integral K h)) \ 2 * real DIM('b) * (\/(5 * Suc DIM('b)))" proof (rule Henstock_lemma_part2 [of h a b]) show "h integrable_on cbox a b" using that F equiintegrable_on_def by metis show "gauge \" by (rule \gauge \\) qed (use that \\ > 0\ \ in auto) also have "... < \/2" using \\ > 0\ by (simp add: divide_simps) finally show ?thesis . qed qed qed define \ where "\ \ \x. \0 x \ ball x ((\/8 / (norm(f x) + 1)) * (INF m\Basis. b \ m - a \ m) / content(cbox a b))" - have "gauge (\x. ball x + define interv_diff where "interv_diff \ \K. \i::'a. interval_upperbound K \ i - interval_lowerbound K \ i" + have "8 * content (cbox a b) + norm (f x) * (8 * content (cbox a b)) > 0" for x + by (metis add.right_neutral add_pos_pos contab_gt0 mult_pos_pos mult_zero_left norm_eq_zero zero_less_norm_iff zero_less_numeral) + then have "gauge (\x. ball x (\ * (INF m\Basis. b \ m - a \ m) / ((8 * norm (f x) + 8) * content (cbox a b))))" - using \0 < content (cbox a b)\ \0 < \\ a_less_b - apply (auto simp: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff) - apply (meson add_increasing measure_nonneg mult_nonneg_nonneg norm_ge_zero not_le zero_le_numeral) - done + using \0 < content (cbox a b)\ \0 < \\ a_less_b + by (auto simp add: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff) then have "gauge \" unfolding \_def using \gauge \0\ gauge_Int by auto moreover have "(\(x,K) \ S. norm (integral K h)) < \" if "c \ cbox a b" "i \ Basis" and S: "S tagged_partial_division_of cbox a b" and "\ fine S" "h \ F" and ne: "\x K. (x,K) \ S \ K \ {x. x \ i = c \ i} \ {}" for c i S h proof - have "cbox c b \ cbox a b" by (meson mem_box(2) order_refl subset_box(1) that(1)) have "finite S" - using S by blast + using S unfolding tagged_partial_division_of_def by blast have "\0 fine S" and fineS: "(\x. ball x (\ * (INF m\Basis. b \ m - a \ m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S" using \\ fine S\ by (auto simp: \_def fine_Int) then have "(\(x,K) \ S. norm (content K *\<^sub>R h x - integral K h)) < \/2" by (intro \0 that fineS) moreover have "(\(x,K) \ S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h)) \ \/2" proof - have "(\(x,K) \ S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h)) \ (\(x,K) \ S. norm (content K *\<^sub>R h x))" proof (clarify intro!: sum_mono) fix x K assume xK: "(x,K) \ S" have "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \ norm (integral K h - (integral K h - content K *\<^sub>R h x))" by (metis norm_minus_commute norm_triangle_ineq2) also have "... \ norm (content K *\<^sub>R h x)" by simp finally show "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \ norm (content K *\<^sub>R h x)" . qed - also have "... \ (\(x,K) \ S. \/4 * (b \ i - a \ i) / content (cbox a b) * - content K / (interval_upperbound K \ i - interval_lowerbound K \ i))" + also have "... \ (\(x,K) \ S. \/4 * (b \ i - a \ i) / content (cbox a b) * content K / interv_diff K i)" proof (clarify intro!: sum_mono) fix x K assume xK: "(x,K) \ S" then have x: "x \ cbox a b" using S unfolding tagged_partial_division_of_def by (meson subset_iff) - let ?\ = "interval_upperbound K \ i - interval_lowerbound K \ i" - show "norm (content K *\<^sub>R h x) \ \/4 * (b \ i - a \ i) / content (cbox a b) * content K / ?\" + show "norm (content K *\<^sub>R h x) \ \/4 * (b \ i - a \ i) / content (cbox a b) * content K / interv_diff K i" proof (cases "content K = 0") case True then show ?thesis by simp next case False then have Kgt0: "content K > 0" using zero_less_measure_iff by blast moreover obtain u v where uv: "K = cbox u v" - using S \(x,K) \ S\ by blast + using S \(x,K) \ S\ unfolding tagged_partial_division_of_def by blast then have u_less_v: "\i. i \ Basis \ u \ i < v \ i" using content_pos_lt_eq uv Kgt0 by blast then have dist_uv: "dist u v > 0" using that by auto - ultimately have "norm (h x) \ (\ * (b \ i - a \ i)) / (4 * content (cbox a b) * ?\)" + ultimately have "norm (h x) \ (\ * (b \ i - a \ i)) / (4 * content (cbox a b) * interv_diff K i)" proof - have "dist x u < \ * (INF m\Basis. b \ m - a \ m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" "dist x v < \ * (INF m\Basis. b \ m - a \ m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" using fineS u_less_v uv xK by (force simp: fine_def mem_box field_simps dest!: bspec)+ moreover have "\ * (INF m\Basis. b \ m - a \ m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2 \ \ * (b \ i - a \ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" - apply (intro mult_left_mono divide_right_mono) - using \i \ Basis\ \0 < \\ apply (auto simp: intro!: cInf_le_finite) - done + proof (intro mult_left_mono divide_right_mono) + show "(INF m\Basis. b \ m - a \ m) \ b \ i - a \ i" + using \i \ Basis\ by (auto intro!: cInf_le_finite) + qed (use \0 < \\ in auto) ultimately have "dist x u < \ * (b \ i - a \ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" "dist x v < \ * (b \ i - a \ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" by linarith+ then have duv: "dist u v < \ * (b \ i - a \ i) / (4 * (norm (f x) + 1) * content (cbox a b))" using dist_triangle_half_r by blast have uvi: "\v \ i - u \ i\ \ norm (v - u)" by (metis inner_commute inner_diff_right \i \ Basis\ Basis_le_norm) have "norm (h x) \ norm (f x)" using x that by (auto simp: norm_f) also have "... < (norm (f x) + 1)" by simp also have "... < \ * (b \ i - a \ i) / dist u v / (4 * content (cbox a b))" - using duv dist_uv contab_gt0 - apply (simp add: field_split_simps split: if_split_asm) - by (meson add_nonneg_nonneg linorder_not_le measure_nonneg mult_nonneg_nonneg norm_ge_zero zero_le_numeral) + proof - + have "0 < norm (f x) + 1" + by (simp add: add.commute add_pos_nonneg) + then show ?thesis + using duv dist_uv contab_gt0 + by (simp only: mult_ac divide_simps) auto + qed also have "... = \ * (b \ i - a \ i) / norm (v - u) / (4 * content (cbox a b))" by (simp add: dist_norm norm_minus_commute) also have "... \ \ * (b \ i - a \ i) / \v \ i - u \ i\ / (4 * content (cbox a b))" - apply (intro mult_right_mono divide_left_mono divide_right_mono uvi) - using \0 < \\ a_less_b [OF \i \ Basis\] u_less_v [OF \i \ Basis\] contab_gt0 - by (auto simp: less_eq_real_def zero_less_mult_iff that) - also have "... = \ * (b \ i - a \ i) - / (4 * content (cbox a b) * ?\)" - using uv False that(2) u_less_v by fastforce + proof (intro mult_right_mono divide_left_mono divide_right_mono uvi) + show "norm (v - u) * \v \ i - u \ i\ > 0" + using u_less_v [OF \i \ Basis\] + by (auto simp: less_eq_real_def zero_less_mult_iff that) + show "\ * (b \ i - a \ i) \ 0" + using a_less_b \0 < \\ \i \ Basis\ by force + qed auto + also have "... = \ * (b \ i - a \ i) / (4 * content (cbox a b) * interv_diff K i)" + using uv False that(2) u_less_v interv_diff_def by fastforce finally show ?thesis by simp qed - with Kgt0 have "norm (content K *\<^sub>R h x) \ content K * ((\/4 * (b \ i - a \ i) / content (cbox a b)) / ?\)" + with Kgt0 have "norm (content K *\<^sub>R h x) \ content K * ((\/4 * (b \ i - a \ i) / content (cbox a b)) / interv_diff K i)" using mult_left_mono by fastforce - also have "... = \/4 * (b \ i - a \ i) / content (cbox a b) * - content K / ?\" + also have "... = \/4 * (b \ i - a \ i) / content (cbox a b) * content K / interv_diff K i" by (simp add: field_split_simps) finally show ?thesis . qed qed - also have "... = (\K\snd ` S. \/4 * (b \ i - a \ i) / content (cbox a b) * content K - / (interval_upperbound K \ i - interval_lowerbound K \ i))" + also have "... = (\K\snd ` S. \/4 * (b \ i - a \ i) / content (cbox a b) * content K / interv_diff K i)" + unfolding interv_diff_def apply (rule sum.over_tagged_division_lemma [OF tagged_partial_division_of_Union_self [OF S]]) apply (simp add: box_eq_empty(1) content_eq_0) done - also have "... = \/2 * ((b \ i - a \ i) / (2 * content (cbox a b)) * (\K\snd ` S. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)))" - by (simp add: sum_distrib_left mult.assoc) + also have "... = \/2 * ((b \ i - a \ i) / (2 * content (cbox a b)) * (\K\snd ` S. content K / interv_diff K i))" + by (simp add: interv_diff_def sum_distrib_left mult.assoc) also have "... \ (\/2) * 1" proof (rule mult_left_mono) - have "(b \ i - a \ i) * (\K\snd ` S. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) - \ 2 * content (cbox a b)" + have "(b \ i - a \ i) * (\K\snd ` S. content K / interv_diff K i) \ 2 * content (cbox a b)" + unfolding interv_diff_def proof (rule sum_content_area_over_thin_division) show "snd ` S division_of \(snd ` S)" by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division) show "\(snd ` S) \ cbox a b" - using S by force + using S unfolding tagged_partial_division_of_def by force show "a \ i \ c \ i" "c \ i \ b \ i" using mem_box(2) that by blast+ qed (use that in auto) - then show "(b \ i - a \ i) / (2 * content (cbox a b)) * (\K\snd ` S. content K / (interval_upperbound K \ i - interval_lowerbound K \ i)) \ 1" + then show "(b \ i - a \ i) / (2 * content (cbox a b)) * (\K\snd ` S. content K / interv_diff K i) \ 1" by (simp add: contab_gt0) qed (use \0 < \\ in auto) finally show ?thesis by simp qed then have "(\(x,K) \ S. norm (integral K h)) - (\(x,K) \ S. norm (content K *\<^sub>R h x - integral K h)) \ \/2" by (simp add: Groups_Big.sum_subtractf [symmetric]) ultimately show "(\(x,K) \ S. norm (integral K h)) < \" by linarith qed ultimately show ?thesis using that by auto qed proposition equiintegrable_halfspace_restrictions_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes F: "F equiintegrable_on cbox a b" and f: "f \ F" and norm_f: "\h x. \h \ F; x \ cbox a b\ \ norm(h x) \ norm(f x)" shows "(\i \ Basis. \c. \h \ F. {(\x. if x \ i \ c then h x else 0)}) equiintegrable_on cbox a b" proof (cases "content(cbox a b) = 0") case True then show ?thesis by simp next case False then have "content(cbox a b) > 0" using zero_less_measure_iff by blast then have "a \ i < b \ i" if "i \ Basis" for i using content_pos_lt_eq that by blast have int_F: "f integrable_on cbox a b" if "f \ F" for f using F that by (simp add: equiintegrable_on_def) let ?CI = "\K h x. content K *\<^sub>R h x - integral K h" show ?thesis unfolding equiintegrable_on_def proof (intro conjI; clarify) show int_lec: "\i \ Basis; h \ F\ \ (\x. if x \ i \ c then h x else 0) integrable_on cbox a b" for i c h using integrable_restrict_Int [of "{x. x \ i \ c}" h] - apply (auto simp: interval_split Int_commute mem_box intro!: integrable_on_subcbox int_F) - by (metis (full_types, hide_lams) min.bounded_iff) + by (simp add: inf_commute int_F integrable_split(1)) show "\\. gauge \ \ (\f T. f \ (\i\Basis. \c. \h\F. {\x. if x \ i \ c then h x else 0}) \ T tagged_division_of cbox a b \ \ fine T \ norm ((\(x,K) \ T. content K *\<^sub>R f x) - integral (cbox a b) f) < \)" if "\ > 0" for \ proof - obtain \0 where "gauge \0" and \0: "\c i S h. \c \ cbox a b; i \ Basis; S tagged_partial_division_of cbox a b; \0 fine S; h \ F; \x K. (x,K) \ S \ (K \ {x. x \ i = c \ i} \ {})\ \ (\(x,K) \ S. norm (integral K h)) < \/12" - apply (rule bounded_equiintegral_over_thin_tagged_partial_division [OF F f, of \\/12\]) - using \\ > 0\ by (auto simp: norm_f) + proof (rule bounded_equiintegral_over_thin_tagged_partial_division [OF F f, of \\/12\]) + show "\h x. \h \ F; x \ cbox a b\ \ norm (h x) \ norm (f x)" + by (auto simp: norm_f) + qed (use \\ > 0\ in auto) obtain \1 where "gauge \1" and \1: "\h T. \h \ F; T tagged_division_of cbox a b; \1 fine T\ \ norm ((\(x,K) \ T. content K *\<^sub>R h x) - integral (cbox a b) h) < \/(7 * (Suc DIM('b)))" proof - have e5: "\/(7 * (Suc DIM('b))) > 0" using \\ > 0\ by auto then show ?thesis using F that by (auto simp: equiintegrable_on_def) qed have h_less3: "(\(x,K) \ T. norm (?CI K h x)) < \/3" if "T tagged_partial_division_of cbox a b" "\1 fine T" "h \ F" for T h proof - have "(\(x,K) \ T. norm (?CI K h x)) \ 2 * real DIM('b) * (\/(7 * Suc DIM('b)))" proof (rule Henstock_lemma_part2 [of h a b]) show "h integrable_on cbox a b" using that F equiintegrable_on_def by metis - show "gauge \1" - by (rule \gauge \1\) - qed (use that \\ > 0\ \1 in auto) + qed (use that \\ > 0\ \gauge \1\ \1 in auto) also have "... < \/3" using \\ > 0\ by (simp add: divide_simps) finally show ?thesis . qed have *: "norm ((\(x,K) \ T. content K *\<^sub>R f x) - integral (cbox a b) f) < \" if f: "f = (\x. if x \ i \ c then h x else 0)" and T: "T tagged_division_of cbox a b" and fine: "(\x. \0 x \ \1 x) fine T" and "i \ Basis" "h \ F" for f T i c h proof (cases "a \ i \ c \ c \ b \ i") case True have "finite T" using T by blast define T' where "T' \ {(x,K) \ T. K \ {x. x \ i \ c} \ {}}" then have "T' \ T" by auto then have "finite T'" using \finite T\ infinite_super by blast have T'_tagged: "T' tagged_partial_division_of cbox a b" by (meson T \T' \ T\ tagged_division_of_def tagged_partial_division_subset) have fine': "\0 fine T'" "\1 fine T'" using \T' \ T\ fine_Int fine_subset fine by blast+ have int_KK': "(\(x,K) \ T. integral K f) = (\(x,K) \ T'. integral K f)" - apply (rule sum.mono_neutral_right [OF \finite T\ \T' \ T\]) - using f \finite T\ \T' \ T\ - using integral_restrict_Int [of _ "{x. x \ i \ c}" h] - apply (auto simp: T'_def Int_commute) - done + proof (rule sum.mono_neutral_right [OF \finite T\ \T' \ T\]) + show "\i \ T - T'. (case i of (x, K) \ integral K f) = 0" + using f \finite T\ \T' \ T\ integral_restrict_Int [of _ "{x. x \ i \ c}" h] + by (auto simp: T'_def Int_commute) + qed have "(\(x,K) \ T. content K *\<^sub>R f x) = (\(x,K) \ T'. content K *\<^sub>R f x)" - apply (rule sum.mono_neutral_right [OF \finite T\ \T' \ T\]) - using T f \finite T\ \T' \ T\ apply (force simp: T'_def) - done + proof (rule sum.mono_neutral_right [OF \finite T\ \T' \ T\]) + show "\i \ T - T'. (case i of (x, K) \ content K *\<^sub>R f x) = 0" + using T f \finite T\ \T' \ T\ by (force simp: T'_def) + qed moreover have "norm ((\(x,K) \ T'. content K *\<^sub>R f x) - integral (cbox a b) f) < \" proof - have *: "norm y < \" if "norm x < \/3" "norm(x - y) \ 2 * \/3" for x y::'b proof - have "norm y \ norm x + norm(x - y)" by (metis norm_minus_commute norm_triangle_sub) also have "\ < \/3 + 2*\/3" using that by linarith also have "... = \" by simp finally show ?thesis . qed have "norm (\(x,K) \ T'. ?CI K h x) \ (\(x,K) \ T'. norm (?CI K h x))" by (simp add: norm_sum split_def) also have "... < \/3" by (intro h_less3 T'_tagged fine' that) finally have "norm (\(x,K) \ T'. ?CI K h x) < \/3" . moreover have "integral (cbox a b) f = (\(x,K) \ T. integral K f)" using int_lec that by (auto simp: integral_combine_tagged_division_topdown) moreover have "norm (\(x,K) \ T'. ?CI K h x - ?CI K f x) \ 2*\/3" proof - define T'' where "T'' \ {(x,K) \ T'. \ (K \ {x. x \ i \ c})}" then have "T'' \ T'" by auto then have "finite T''" using \finite T'\ infinite_super by blast have T''_tagged: "T'' tagged_partial_division_of cbox a b" using T'_tagged \T'' \ T'\ tagged_partial_division_subset by blast have fine'': "\0 fine T''" "\1 fine T''" using \T'' \ T'\ fine' by (blast intro: fine_subset)+ have "(\(x,K) \ T'. ?CI K h x - ?CI K f x) = (\(x,K) \ T''. ?CI K h x - ?CI K f x)" proof (clarify intro!: sum.mono_neutral_right [OF \finite T'\ \T'' \ T'\]) fix x K assume "(x,K) \ T'" "(x,K) \ T''" then have "x \ K" "x \ i \ c" "{x. x \ i \ c} \ K = K" - using T''_def T'_tagged by blast+ + using T''_def T'_tagged tagged_partial_division_of_def by blast+ then show "?CI K h x - ?CI K f x = 0" using integral_restrict_Int [of _ "{x. x \ i \ c}" h] by (auto simp: f) qed moreover have "norm (\(x,K) \ T''. ?CI K h x - ?CI K f x) \ 2*\/3" proof - define A where "A \ {(x,K) \ T''. x \ i \ c}" define B where "B \ {(x,K) \ T''. x \ i > c}" then have "A \ T''" "B \ T''" and disj: "A \ B = {}" and T''_eq: "T'' = A \ B" by (auto simp: A_def B_def) then have "finite A" "finite B" using \finite T''\ by (auto intro: finite_subset) have A_tagged: "A tagged_partial_division_of cbox a b" using T''_tagged \A \ T''\ tagged_partial_division_subset by blast have fineA: "\0 fine A" "\1 fine A" using \A \ T''\ fine'' by (blast intro: fine_subset)+ have B_tagged: "B tagged_partial_division_of cbox a b" using T''_tagged \B \ T''\ tagged_partial_division_subset by blast have fineB: "\0 fine B" "\1 fine B" using \B \ T''\ fine'' by (blast intro: fine_subset)+ have "norm (\(x,K) \ T''. ?CI K h x - ?CI K f x) \ (\(x,K) \ T''. norm (?CI K h x - ?CI K f x))" by (simp add: norm_sum split_def) also have "... = (\(x,K) \ A. norm (?CI K h x - ?CI K f x)) + (\(x,K) \ B. norm (?CI K h x - ?CI K f x))" by (simp add: sum.union_disjoint T''_eq disj \finite A\ \finite B\) also have "... = (\(x,K) \ A. norm (integral K h - integral K f)) + (\(x,K) \ B. norm (?CI K h x + integral K f))" by (auto simp: A_def B_def f norm_minus_commute intro!: sum.cong arg_cong2 [where f= "(+)"]) also have "... \ (\(x,K)\A. norm (integral K h)) + (\(x,K)\(\(x,K). (x,K \ {x. x \ i \ c})) ` A. norm (integral K h)) + ((\(x,K)\B. norm (?CI K h x)) + (\(x,K)\B. norm (integral K h)) + (\(x,K)\(\(x,K). (x,K \ {x. c \ x \ i})) ` B. norm (integral K h)))" proof (rule add_mono) show "(\(x,K)\A. norm (integral K h - integral K f)) \ (\(x,K)\A. norm (integral K h)) + (\(x,K)\(\(x,K). (x,K \ {x. x \ i \ c})) ` A. norm (integral K h))" proof (subst sum.reindex_nontrivial [OF \finite A\], clarsimp) fix x K L assume "(x,K) \ A" "(x,L) \ A" and int_ne0: "integral (L \ {x. x \ i \ c}) h \ 0" and eq: "K \ {x. x \ i \ c} = L \ {x. x \ i \ c}" have False if "K \ L" proof - obtain u v where uv: "L = cbox u v" - using T'_tagged \(x, L) \ A\ \A \ T''\ \T'' \ T'\ by blast - have "A tagged_division_of \(snd ` A)" - using A_tagged tagged_partial_division_of_Union_self by auto - then have "interior (K \ {x. x \ i \ c}) = {}" - apply (rule tagged_division_split_left_inj [OF _ \(x,K) \ A\ \(x,L) \ A\]) - using that eq \i \ Basis\ by auto + using T'_tagged \(x, L) \ A\ \A \ T''\ \T'' \ T'\ by (blast dest: tagged_partial_division_ofD) + have "interior (K \ {x. x \ i \ c}) = {}" + proof (rule tagged_division_split_left_inj [OF _ \(x,K) \ A\ \(x,L) \ A\]) + show "A tagged_division_of \(snd ` A)" + using A_tagged tagged_partial_division_of_Union_self by auto + show "K \ {x. x \ i \ c} = L \ {x. x \ i \ c}" + using eq \i \ Basis\ by auto + qed (use that in auto) then show False using interval_split [OF \i \ Basis\] int_ne0 content_eq_0_interior eq uv by fastforce qed then show "K = L" by blast next show "(\(x,K) \ A. norm (integral K h - integral K f)) \ (\(x,K) \ A. norm (integral K h)) + sum ((\(x,K). norm (integral K h)) \ (\(x,K). (x,K \ {x. x \ i \ c}))) A" using integral_restrict_Int [of _ "{x. x \ i \ c}" h] f by (auto simp: Int_commute A_def [symmetric] sum.distrib [symmetric] intro!: sum_mono norm_triangle_ineq4) qed next show "(\(x,K)\B. norm (?CI K h x + integral K f)) \ (\(x,K)\B. norm (?CI K h x)) + (\(x,K)\B. norm (integral K h)) + (\(x,K)\(\(x,K). (x,K \ {x. c \ x \ i})) ` B. norm (integral K h))" proof (subst sum.reindex_nontrivial [OF \finite B\], clarsimp) fix x K L assume "(x,K) \ B" "(x,L) \ B" and int_ne0: "integral (L \ {x. c \ x \ i}) h \ 0" and eq: "K \ {x. c \ x \ i} = L \ {x. c \ x \ i}" have False if "K \ L" proof - obtain u v where uv: "L = cbox u v" - using T'_tagged \(x, L) \ B\ \B \ T''\ \T'' \ T'\ by blast - have "B tagged_division_of \(snd ` B)" - using B_tagged tagged_partial_division_of_Union_self by auto - then have "interior (K \ {x. c \ x \ i}) = {}" - apply (rule tagged_division_split_right_inj [OF _ \(x,K) \ B\ \(x,L) \ B\]) - using that eq \i \ Basis\ by auto + using T'_tagged \(x, L) \ B\ \B \ T''\ \T'' \ T'\ by (blast dest: tagged_partial_division_ofD) + have "interior (K \ {x. c \ x \ i}) = {}" + proof (rule tagged_division_split_right_inj [OF _ \(x,K) \ B\ \(x,L) \ B\]) + show "B tagged_division_of \(snd ` B)" + using B_tagged tagged_partial_division_of_Union_self by auto + show "K \ {x. c \ x \ i} = L \ {x. c \ x \ i}" + using eq \i \ Basis\ by auto + qed (use that in auto) then show False using interval_split [OF \i \ Basis\] int_ne0 content_eq_0_interior eq uv by fastforce qed then show "K = L" by blast next show "(\(x,K) \ B. norm (?CI K h x + integral K f)) \ (\(x,K) \ B. norm (?CI K h x)) + (\(x,K) \ B. norm (integral K h)) + sum ((\(x,K). norm (integral K h)) \ (\(x,K). (x,K \ {x. c \ x \ i}))) B" proof (clarsimp simp: B_def [symmetric] sum.distrib [symmetric] intro!: sum_mono) fix x K assume "(x,K) \ B" have *: "i = i1 + i2 \ norm(c + i1) \ norm c + norm i + norm(i2)" for i::'b and c i1 i2 by (metis add.commute add.left_commute add_diff_cancel_right' dual_order.refl norm_add_rule_thm norm_triangle_ineq4) obtain u v where uv: "K = cbox u v" - using T'_tagged \(x,K) \ B\ \B \ T''\ \T'' \ T'\ by blast - have "h integrable_on cbox a b" - by (simp add: int_F \h \ F\) - then have huv: "h integrable_on cbox u v" - apply (rule integrable_on_subcbox) - using B_tagged \(x,K) \ B\ uv by blast + using T'_tagged \(x,K) \ B\ \B \ T''\ \T'' \ T'\ by (blast dest: tagged_partial_division_ofD) + have huv: "h integrable_on cbox u v" + proof (rule integrable_on_subcbox) + show "cbox u v \ cbox a b" + using B_tagged \(x,K) \ B\ uv by (blast dest: tagged_partial_division_ofD) + show "h integrable_on cbox a b" + by (simp add: int_F \h \ F\) + qed have "integral K h = integral K f + integral (K \ {x. c \ x \ i}) h" using integral_restrict_Int [of _ "{x. x \ i \ c}" h] f uv \i \ Basis\ by (simp add: Int_commute integral_split [OF huv \i \ Basis\]) then show "norm (?CI K h x + integral K f) \ norm (?CI K h x) + norm (integral K h) + norm (integral (K \ {x. c \ x \ i}) h)" by (rule *) qed qed qed also have "... \ 2*\/3" proof - have overlap: "K \ {x. x \ i = c} \ {}" if "(x,K) \ T''" for x K proof - obtain y y' where y: "y' \ K" "c < y' \ i" "y \ K" "y \ i \ c" using that T''_def T'_def \(x,K) \ T''\ by fastforce obtain u v where uv: "K = cbox u v" - using T''_tagged \(x,K) \ T''\ by blast + using T''_tagged \(x,K) \ T''\ by (blast dest: tagged_partial_division_ofD) then have "connected K" by (simp add: is_interval_connected) then have "(\z \ K. z \ i = c)" using y connected_ivt_component by fastforce then show ?thesis by fastforce qed have **: "\x < \/12; y < \/12; z \ \/2\ \ x + y + z \ 2 * \/3" for x y z by auto show ?thesis proof (rule **) have cb_ab: "(\j \ Basis. if j = i then c *\<^sub>R i else (a \ j) *\<^sub>R j) \ cbox a b" using \i \ Basis\ True \\i. i \ Basis \ a \ i < b \ i\ - apply (clarsimp simp add: mem_box) - apply (subst sum_if_inner | force)+ - done + by (force simp add: mem_box sum_if_inner [where f = "\j. c"]) show "(\(x,K) \ A. norm (integral K h)) < \/12" - apply (rule \0 [OF cb_ab \i \ Basis\ A_tagged fineA(1) \h \ F\]) using \i \ Basis\ \A \ T''\ overlap - apply (subst sum_if_inner | force)+ - done - have 1: "(\(x,K). (x,K \ {x. x \ i \ c})) ` A tagged_partial_division_of cbox a b" - using \finite A\ \i \ Basis\ - apply (auto simp: tagged_partial_division_of_def) - using A_tagged apply (auto simp: A_def) - using interval_split(1) by blast + by (force simp add: sum_if_inner [where f = "\j. c"] + intro!: \0 [OF cb_ab \i \ Basis\ A_tagged fineA(1) \h \ F\]) + let ?F = "\(x,K). (x, K \ {x. x \ i \ c})" + have 1: "?F ` A tagged_partial_division_of cbox a b" + unfolding tagged_partial_division_of_def + proof (intro conjI strip) + show "\x K. (x, K) \ ?F ` A \ \a b. K = cbox a b" + using A_tagged interval_split(1) [OF \i \ Basis\, of _ _ c] + by (force dest: tagged_partial_division_ofD(4)) + show "\x K. (x, K) \ ?F ` A \ x \ K" + using A_def A_tagged by (fastforce dest: tagged_partial_division_ofD) + qed (use A_tagged in \fastforce dest: tagged_partial_division_ofD\)+ have 2: "\0 fine (\(x,K). (x,K \ {x. x \ i \ c})) ` A" using fineA(1) fine_def by fastforce show "(\(x,K) \ (\(x,K). (x,K \ {x. x \ i \ c})) ` A. norm (integral K h)) < \/12" - apply (rule \0 [OF cb_ab \i \ Basis\ 1 2 \h \ F\]) - using \i \ Basis\ apply (subst sum_if_inner | force)+ - using overlap apply (auto simp: A_def) - done + using \i \ Basis\ \A \ T''\ overlap + by (force simp add: sum_if_inner [where f = "\j. c"] + intro!: \0 [OF cb_ab \i \ Basis\ 1 2 \h \ F\]) have *: "\x < \/3; y < \/12; z < \/12\ \ x + y + z \ \/2" for x y z by auto show "(\(x,K) \ B. norm (?CI K h x)) + (\(x,K) \ B. norm (integral K h)) + (\(x,K) \ (\(x,K). (x,K \ {x. c \ x \ i})) ` B. norm (integral K h)) \ \/2" proof (rule *) show "(\(x,K) \ B. norm (?CI K h x)) < \/3" by (intro h_less3 B_tagged fineB that) show "(\(x,K) \ B. norm (integral K h)) < \/12" - apply (rule \0 [OF cb_ab \i \ Basis\ B_tagged fineB(1) \h \ F\]) - using \i \ Basis\ \B \ T''\ overlap by (subst sum_if_inner | force)+ - have 1: "(\(x,K). (x,K \ {x. c \ x \ i})) ` B tagged_partial_division_of cbox a b" - using \finite B\ \i \ Basis\ - apply (auto simp: tagged_partial_division_of_def) - using B_tagged apply (auto simp: B_def) - using interval_split(2) by blast + using \i \ Basis\ \B \ T''\ overlap + by (force simp add: sum_if_inner [where f = "\j. c"] + intro!: \0 [OF cb_ab \i \ Basis\ B_tagged fineB(1) \h \ F\]) + let ?F = "\(x,K). (x, K \ {x. c \ x \ i})" + have 1: "?F ` B tagged_partial_division_of cbox a b" + unfolding tagged_partial_division_of_def + proof (intro conjI strip) + show "\x K. (x, K) \ ?F ` B \ \a b. K = cbox a b" + using B_tagged interval_split(2) [OF \i \ Basis\, of _ _ c] + by (force dest: tagged_partial_division_ofD(4)) + show "\x K. (x, K) \ ?F ` B \ x \ K" + using B_def B_tagged by (fastforce dest: tagged_partial_division_ofD) + qed (use B_tagged in \fastforce dest: tagged_partial_division_ofD\)+ have 2: "\0 fine (\(x,K). (x,K \ {x. c \ x \ i})) ` B" using fineB(1) fine_def by fastforce show "(\(x,K) \ (\(x,K). (x,K \ {x. c \ x \ i})) ` B. norm (integral K h)) < \/12" - apply (rule \0 [OF cb_ab \i \ Basis\ 1 2 \h \ F\]) - using \i \ Basis\ apply (subst sum_if_inner | force)+ - using overlap apply (auto simp: B_def) - done + using \i \ Basis\ \A \ T''\ overlap + by (force simp add: B_def sum_if_inner [where f = "\j. c"] + intro!: \0 [OF cb_ab \i \ Basis\ 1 2 \h \ F\]) qed qed qed finally show ?thesis . qed ultimately show ?thesis by metis qed ultimately show ?thesis by (simp add: sum_subtractf [symmetric] int_KK' *) qed ultimately show ?thesis by metis next case False then consider "c < a \ i" | "b \ i < c" by auto then show ?thesis proof cases case 1 then have f0: "f x = 0" if "x \ cbox a b" for x using that f \i \ Basis\ mem_box(2) by force then have int_f0: "integral (cbox a b) f = 0" by (simp add: integral_cong) have f0_tag: "f x = 0" if "(x,K) \ T" for x K - using T f0 that by (force simp: tagged_division_of_def) + using T f0 that by (meson tag_in_interval) then have "(\(x,K) \ T. content K *\<^sub>R f x) = 0" by (metis (mono_tags, lifting) real_vector.scale_eq_0_iff split_conv sum.neutral surj_pair) then show ?thesis using \0 < \\ by (simp add: int_f0) next case 2 then have fh: "f x = h x" if "x \ cbox a b" for x using that f \i \ Basis\ mem_box(2) by force then have int_f: "integral (cbox a b) f = integral (cbox a b) h" using integral_cong by blast have fh_tag: "f x = h x" if "(x,K) \ T" for x K - using T fh that by (force simp: tagged_division_of_def) - then have "(\(x,K) \ T. content K *\<^sub>R f x) = (\(x,K) \ T. content K *\<^sub>R h x)" + using T fh that by (meson tag_in_interval) + then have fh: "(\(x,K) \ T. content K *\<^sub>R f x) = (\(x,K) \ T. content K *\<^sub>R h x)" by (metis (mono_tags, lifting) split_cong sum.cong) - with \0 < \\ show ?thesis - apply (simp add: int_f) - apply (rule less_trans [OF \1]) - using that fine_Int apply (force simp: divide_simps)+ - done + show ?thesis + unfolding fh int_f + proof (rule less_trans [OF \1]) + show "\1 fine T" + by (meson fine fine_Int) + show "\ / (7 * Suc DIM('b)) < \" + using \0 < \\ by (force simp: divide_simps)+ + qed (use that in auto) qed qed have "gauge (\x. \0 x \ \1 x)" by (simp add: \gauge \0\ \gauge \1\ gauge_Int) then show ?thesis by (auto intro: *) qed qed qed corollary equiintegrable_halfspace_restrictions_ge: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes F: "F equiintegrable_on cbox a b" and f: "f \ F" and norm_f: "\h x. \h \ F; x \ cbox a b\ \ norm(h x) \ norm(f x)" shows "(\i \ Basis. \c. \h \ F. {(\x. if x \ i \ c then h x else 0)}) equiintegrable_on cbox a b" proof - have *: "(\i\Basis. \c. \h\(\f. f \ uminus) ` F. {\x. if x \ i \ c then h x else 0}) equiintegrable_on cbox (- b) (- a)" proof (rule equiintegrable_halfspace_restrictions_le) show "(\f. f \ uminus) ` F equiintegrable_on cbox (- b) (- a)" using F equiintegrable_reflect by blast show "f \ uminus \ (\f. f \ uminus) ` F" using f by auto show "\h x. \h \ (\f. f \ uminus) ` F; x \ cbox (- b) (- a)\ \ norm (h x) \ norm ((f \ uminus) x)" - using f apply (clarsimp simp:) - by (metis add.inverse_inverse image_eqI norm_f uminus_interval_vector) + using f unfolding comp_def image_iff + by (metis (no_types, lifting) equation_minus_iff imageE norm_f uminus_interval_vector) qed have eq: "(\f. f \ uminus) ` (\i\Basis. \c. \h\F. {\x. if x \ i \ c then (h \ uminus) x else 0}) = - (\i\Basis. \c. \h\F. {\x. if c \ x \ i then h x else 0})" - apply (auto simp: o_def cong: if_cong) - using minus_le_iff apply fastforce + (\i\Basis. \c. \h\F. {\x. if c \ x \ i then h x else 0})" (is "?lhs = ?rhs") + proof + show "?lhs \ ?rhs" + using minus_le_iff by fastforce + show "?rhs \ ?lhs" + apply clarsimp apply (rule_tac x="\x. if c \ (-x) \ i then h(-x) else 0" in image_eqI) - using le_minus_iff apply fastforce+ - done + using le_minus_iff by fastforce+ + qed show ?thesis using equiintegrable_reflect [OF *] by (auto simp: eq) qed corollary equiintegrable_halfspace_restrictions_lt: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes F: "F equiintegrable_on cbox a b" and f: "f \ F" and norm_f: "\h x. \h \ F; x \ cbox a b\ \ norm(h x) \ norm(f x)" shows "(\i \ Basis. \c. \h \ F. {(\x. if x \ i < c then h x else 0)}) equiintegrable_on cbox a b" (is "?G equiintegrable_on cbox a b") proof - have *: "(\i\Basis. \c. \h\F. {\x. if c \ x \ i then h x else 0}) equiintegrable_on cbox a b" using equiintegrable_halfspace_restrictions_ge [OF F f] norm_f by auto have "(\x. if x \ i < c then h x else 0) = (\x. h x - (if c \ x \ i then h x else 0))" if "i \ Basis" "h \ F" for i c h using that by force then show ?thesis by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]]) qed corollary equiintegrable_halfspace_restrictions_gt: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes F: "F equiintegrable_on cbox a b" and f: "f \ F" and norm_f: "\h x. \h \ F; x \ cbox a b\ \ norm(h x) \ norm(f x)" shows "(\i \ Basis. \c. \h \ F. {(\x. if x \ i > c then h x else 0)}) equiintegrable_on cbox a b" (is "?G equiintegrable_on cbox a b") proof - have *: "(\i\Basis. \c. \h\F. {\x. if c \ x \ i then h x else 0}) equiintegrable_on cbox a b" using equiintegrable_halfspace_restrictions_le [OF F f] norm_f by auto have "(\x. if x \ i > c then h x else 0) = (\x. h x - (if c \ x \ i then h x else 0))" if "i \ Basis" "h \ F" for i c h using that by force then show ?thesis by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]]) qed proposition equiintegrable_closed_interval_restrictions: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f integrable_on cbox a b" shows "(\c d. {(\x. if x \ cbox c d then f x else 0)}) equiintegrable_on cbox a b" proof - let ?g = "\B c d x. if \i\B. c \ i \ x \ i \ x \ i \ d \ i then f x else 0" have *: "insert f (\c d. {?g B c d}) equiintegrable_on cbox a b" if "B \ Basis" for B proof - have "finite B" using finite_Basis finite_subset \B \ Basis\ by blast then show ?thesis using \B \ Basis\ proof (induction B) case empty with f show ?case by auto next case (insert i B) - then have "i \ Basis" + then have "i \ Basis" "B \ Basis" by auto have *: "norm (h x) \ norm (f x)" if "h \ insert f (\c d. {?g B c d})" "x \ cbox a b" for h x using that by auto - have "(\i\Basis. + define F where "F \ (\i\Basis. \\. \h\insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). {\x. if x \ i \ \ then h x else 0}). - {\x. if \ \ x \ i then h x else 0}) + {\x. if \ \ x \ i then h x else 0})" + show ?case + proof (rule equiintegrable_on_subset) + have "F equiintegrable_on cbox a b" + unfolding F_def + proof (rule equiintegrable_halfspace_restrictions_ge) + show "insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). + {\x. if x \ i \ \ then h x else 0}) equiintegrable_on cbox a b" + by (intro * f equiintegrable_on_insert equiintegrable_halfspace_restrictions_le [OF insert.IH insertI1] \B \ Basis\) + show "norm(h x) \ norm(f x)" + if "h \ insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). {\x. if x \ i \ \ then h x else 0})" + "x \ cbox a b" for h x + using that by auto + qed auto + then show "insert f F equiintegrable_on cbox a b" - proof (rule equiintegrable_halfspace_restrictions_ge [where f=f]) - show "insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). - {\x. if x \ i \ \ then h x else 0}) equiintegrable_on cbox a b" - apply (intro * f equiintegrable_on_insert equiintegrable_halfspace_restrictions_le [OF insert.IH insertI1]) - using insert.prems apply auto + by (blast intro: f equiintegrable_on_insert) + show "insert f (\c d. {\x. if \j\insert i B. c \ j \ x \ j \ x \ j \ d \ j then f x else 0}) + \ insert f F" + using \i \ Basis\ + apply clarify + apply (simp add: F_def) + apply (drule_tac x=i in bspec, assumption) + apply (drule_tac x="c \ i" in spec, clarify) + apply (drule_tac x=i in bspec, assumption) + apply (drule_tac x="d \ i" in spec) + apply (clarsimp simp: fun_eq_iff) + apply (drule_tac x=c in spec) + apply (drule_tac x=d in spec) + apply (simp split: if_split_asm) done - show"norm(h x) \ norm(f x)" - if "h \ insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). {\x. if x \ i \ \ then h x else 0})" - "x \ cbox a b" for h x - using that by auto - qed auto - then have "insert f (\i\Basis. - \\. \h\insert f (\i\Basis. \\. \h\insert f (\c d. {?g B c d}). {\x. if x \ i \ \ then h x else 0}). - {\x. if \ \ x \ i then h x else 0}) - equiintegrable_on cbox a b" - by (blast intro: f equiintegrable_on_insert) - then show ?case - apply (rule equiintegrable_on_subset, clarify) - using \i \ Basis\ apply simp - apply (drule_tac x=i in bspec, assumption) - apply (drule_tac x="c \ i" in spec, clarify) - apply (drule_tac x=i in bspec, assumption) - apply (drule_tac x="d \ i" in spec) - apply (clarsimp simp add: fun_eq_iff) - apply (drule_tac x=c in spec) - apply (drule_tac x=d in spec) - apply (simp add: split: if_split_asm) - done + qed qed qed show ?thesis by (rule equiintegrable_on_subset [OF * [OF subset_refl]]) (auto simp: mem_box) qed - subsection\Continuity of the indefinite integral\ proposition indefinite_integral_continuous: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes int_f: "f integrable_on cbox a b" and c: "c \ cbox a b" and d: "d \ cbox a b" "0 < \" obtains \ where "0 < \" "\c' d'. \c' \ cbox a b; d' \ cbox a b; norm(c' - c) \ \; norm(d' - d) \ \\ \ norm(integral(cbox c' d') f - integral(cbox c d) f) < \" proof - { assume "\c' d'. c' \ cbox a b \ d' \ cbox a b \ norm(c' - c) \ \ \ norm(d' - d) \ \ \ norm(integral(cbox c' d') f - integral(cbox c d) f) \ \" (is "\c' d'. ?\ c' d' \") if "0 < \" for \ then have "\c' d'. ?\ c' d' (1 / Suc n)" for n by simp then obtain u v where "\n. ?\ (u n) (v n) (1 / Suc n)" by metis then have u: "u n \ cbox a b" and norm_u: "norm(u n - c) \ 1 / Suc n" and v: "v n \ cbox a b" and norm_v: "norm(v n - d) \ 1 / Suc n" and \: "\ \ norm (integral (cbox (u n) (v n)) f - integral (cbox c d) f)" for n by blast+ then have False proof - have uvn: "cbox (u n) (v n) \ cbox a b" for n by (meson u v mem_box(2) subset_box(1)) define S where "S \ \i \ Basis. {x. x \ i = c \ i} \ {x. x \ i = d \ i}" have "negligible S" unfolding S_def by force then have int_f': "(\x. if x \ S then 0 else f x) integrable_on cbox a b" by (force intro: integrable_spike assms) have get_n: "\n. \m\n. x \ cbox (u m) (v m) \ x \ cbox c d" if x: "x \ S" for x proof - define \ where "\ \ Min ((\i. min \x \ i - c \ i\ \x \ i - d \ i\) ` Basis)" have "\ > 0" using \x \ S\ by (auto simp: S_def \_def) then obtain n where "n \ 0" and n: "1 / (real n) < \" by (metis inverse_eq_divide real_arch_inverse) have emin: "\ \ min \x \ i - c \ i\ \x \ i - d \ i\" if "i \ Basis" for i - unfolding \_def - apply (rule Min.coboundedI) - using that by force+ + unfolding \_def + by (meson Min.coboundedI euclidean_space_class.finite_Basis finite_imageI image_iff that) have "1 / real (Suc n) < \" using n \n \ 0\ \\ > 0\ by (simp add: field_simps) have "x \ cbox (u m) (v m) \ x \ cbox c d" if "m \ n" for m proof - have *: "\\u - c\ \ n; \v - d\ \ n; N < \x - c\; N < \x - d\; n \ N\ \ u \ x \ x \ v \ c \ x \ x \ d" for N n u v c d and x::real by linarith have "(u m \ i \ x \ i \ x \ i \ v m \ i) = (c \ i \ x \ i \ x \ i \ d \ i)" if "i \ Basis" for i proof (rule *) show "\u m \ i - c \ i\ \ 1 / Suc m" using norm_u [of m] by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that) show "\v m \ i - d \ i\ \ 1 / real (Suc m)" using norm_v [of m] by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that) show "1/n < \x \ i - c \ i\" "1/n < \x \ i - d \ i\" using n \n \ 0\ emin [OF \i \ Basis\] by (simp_all add: inverse_eq_divide) show "1 / real (Suc m) \ 1 / real n" using \n \ 0\ \m \ n\ by (simp add: field_split_simps) qed then show ?thesis by (simp add: mem_box) qed then show ?thesis by blast qed have 1: "range (\n x. if x \ cbox (u n) (v n) then if x \ S then 0 else f x else 0) equiintegrable_on cbox a b" by (blast intro: equiintegrable_on_subset [OF equiintegrable_closed_interval_restrictions [OF int_f']]) have 2: "(\n. if x \ cbox (u n) (v n) then if x \ S then 0 else f x else 0) \ (if x \ cbox c d then if x \ S then 0 else f x else 0)" for x by (fastforce simp: dest: get_n intro: tendsto_eventually eventually_sequentiallyI) have [simp]: "cbox c d \ cbox a b = cbox c d" using c d by (force simp: mem_box) have [simp]: "cbox (u n) (v n) \ cbox a b = cbox (u n) (v n)" for n using u v by (fastforce simp: mem_box intro: order.trans) have "\y A. y \ A - S \ f y = (\x. if x \ S then 0 else f x) y" by simp then have "\A. integral A (\x. if x \ S then 0 else f (x)) = integral A (\x. f (x))" by (blast intro: integral_spike [OF \negligible S\]) moreover obtain N where "dist (integral (cbox (u N) (v N)) (\x. if x \ S then 0 else f x)) (integral (cbox c d) (\x. if x \ S then 0 else f x)) < \" using equiintegrable_limit [OF 1 2] \0 < \\ by (force simp: integral_restrict_Int lim_sequentially) ultimately have "dist (integral (cbox (u N) (v N)) f) (integral (cbox c d) f) < \" by simp then show False by (metis dist_norm not_le \) qed } then show ?thesis by (meson not_le that) qed corollary indefinite_integral_uniformly_continuous: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on cbox a b" shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\y. integral (cbox (fst y) (snd y)) f)" proof - show ?thesis proof (rule compact_uniformly_continuous, clarsimp simp add: continuous_on_iff) fix c d and \::real assume c: "c \ cbox a b" and d: "d \ cbox a b" and "0 < \" obtain \ where "0 < \" and \: "\c' d'. \c' \ cbox a b; d' \ cbox a b; norm(c' - c) \ \; norm(d' - d) \ \\ \ norm(integral(cbox c' d') f - integral(cbox c d) f) < \" using indefinite_integral_continuous \0 < \\ assms c d by blast show "\\ > 0. \x' \ cbox (a, a) (b, b). dist x' (c, d) < \ \ dist (integral (cbox (fst x') (snd x')) f) (integral (cbox c d) f) < \" using \0 < \\ by (force simp: dist_norm intro: \ order_trans [OF norm_fst_le] order_trans [OF norm_snd_le] less_imp_le) qed auto qed corollary bounded_integrals_over_subintervals: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on cbox a b" shows "bounded {integral (cbox c d) f |c d. cbox c d \ cbox a b}" proof - have "bounded ((\y. integral (cbox (fst y) (snd y)) f) ` cbox (a, a) (b, b))" (is "bounded ?I") by (blast intro: bounded_cbox bounded_uniformly_continuous_image indefinite_integral_uniformly_continuous [OF assms]) then obtain B where "B > 0" and B: "\x. x \ ?I \ norm x \ B" by (auto simp: bounded_pos) have "norm x \ B" if "x = integral (cbox c d) f" "cbox c d \ cbox a b" for x c d proof (cases "cbox c d = {}") case True with \0 < B\ that show ?thesis by auto next case False - show ?thesis - apply (rule B) - using that \B > 0\ False apply (clarsimp simp: image_def) - by (metis cbox_Pair_iff interval_subset_is_interval is_interval_cbox prod.sel) + then have "\x \ cbox (a,a) (b,b). integral (cbox c d) f = integral (cbox (fst x) (snd x)) f" + using that by (metis cbox_Pair_iff interval_subset_is_interval is_interval_cbox prod.sel) + then show ?thesis + using B that(1) by blast qed then show ?thesis by (blast intro: boundedI) qed text\An existence theorem for "improper" integrals. Hake's theorem implies that if the integrals over subintervals have a limit, the integral exists. We only need to assume that the integrals are bounded, and we get absolute integrability, but we also need a (rather weak) bound assumption on the function.\ theorem absolutely_integrable_improper: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes int_f: "\c d. cbox c d \ box a b \ f integrable_on cbox c d" and bo: "bounded {integral (cbox c d) f |c d. cbox c d \ box a b}" and absi: "\i. i \ Basis \ \g. g absolutely_integrable_on cbox a b \ ((\x \ cbox a b. f x \ i \ g x) \ (\x \ cbox a b. f x \ i \ g x))" shows "f absolutely_integrable_on cbox a b" proof (cases "content(cbox a b) = 0") case True then show ?thesis by auto next case False then have pos: "content(cbox a b) > 0" using zero_less_measure_iff by blast show ?thesis unfolding absolutely_integrable_componentwise_iff [where f = f] proof fix j::'N assume "j \ Basis" then obtain g where absint_g: "g absolutely_integrable_on cbox a b" and g: "(\x \ cbox a b. f x \ j \ g x) \ (\x \ cbox a b. f x \ j \ g x)" using absi by blast have int_gab: "g integrable_on cbox a b" using absint_g set_lebesgue_integral_eq_integral(1) by blast - have 1: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \ box a b" for n - apply (rule subset_box_imp) - using pos apply (auto simp: content_pos_lt_eq algebra_simps) - done - have 2: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \ - cbox (a + (b - a) /\<^sub>R real (Suc n + 1)) (b - (b - a) /\<^sub>R real (Suc n + 1))" for n - apply (rule subset_box_imp) - using pos apply (simp add: content_pos_lt_eq algebra_simps) - apply (simp add: divide_simps) - apply (auto simp: field_simps) - done - have getN: "\N::nat. \k. k \ N \ x \ cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)" + define \ where "\ \ \k. a + (b - a) /\<^sub>R real k" + define \ where "\ \ \k. b - (b - a) /\<^sub>R real k" + define I where "I \ \k. cbox (\ k) (\ k)" + have ISuc_box: "I (Suc n) \ box a b" for n + using pos unfolding I_def + by (intro subset_box_imp) (auto simp: \_def \_def content_pos_lt_eq algebra_simps) + have ISucSuc: "I (Suc n) \ I (Suc (Suc n))" for n + proof - + have "\i. i \ Basis + \ a \ i / Suc n + b \ i / (real n + 2) \ b \ i / Suc n + a \ i / (real n + 2)" + using pos + by (simp add: content_pos_lt_eq divide_simps) (auto simp: algebra_simps) + then show ?thesis + unfolding I_def + by (intro subset_box_imp) (auto simp: algebra_simps inverse_eq_divide \_def \_def) + qed + have getN: "\N::nat. \k. k \ N \ x \ I k" if x: "x \ box a b" for x proof - - let ?\ = "(\i \ Basis. {((x - a) \ i) / ((b - a) \ i), (b - x) \ i / ((b - a) \ i)})" - obtain N where N: "real N > 1 / Inf ?\" + define \ where "\ \ (\i \ Basis. {((x - a) \ i) / ((b - a) \ i), (b - x) \ i / ((b - a) \ i)})" + obtain N where N: "real N > 1 / Inf \" using reals_Archimedean2 by blast - moreover have \: "Inf ?\ > 0" - using that by (auto simp: finite_less_Inf_iff mem_box algebra_simps divide_simps) + moreover have \: "Inf \ > 0" + using that by (auto simp: \_def finite_less_Inf_iff mem_box algebra_simps divide_simps) ultimately have "N > 0" using of_nat_0_less_iff by fastforce show ?thesis proof (intro exI impI allI) fix k assume "N \ k" with \0 < N\ have "k > 0" by linarith have xa_gt: "(x - a) \ i > ((b - a) \ i) / (real k)" if "i \ Basis" for i proof - - have *: "Inf ?\ \ ((x - a) \ i) / ((b - a) \ i)" - using that by (force intro: cInf_le_finite) - have "1 / Inf ?\ \ ((b - a) \ i) / ((x - a) \ i)" + have *: "Inf \ \ ((x - a) \ i) / ((b - a) \ i)" + unfolding \_def using that by (force intro: cInf_le_finite) + have "1 / Inf \ \ ((b - a) \ i) / ((x - a) \ i)" using le_imp_inverse_le [OF * \] by (simp add: field_simps) with N have "k > ((b - a) \ i) / ((x - a) \ i)" using \N \ k\ by linarith with x that show ?thesis by (auto simp: mem_box algebra_simps field_split_simps) qed have bx_gt: "(b - x) \ i > ((b - a) \ i) / k" if "i \ Basis" for i proof - - have *: "Inf ?\ \ ((b - x) \ i) / ((b - a) \ i)" - using that by (force intro: cInf_le_finite) - have "1 / Inf ?\ \ ((b - a) \ i) / ((b - x) \ i)" + have *: "Inf \ \ ((b - x) \ i) / ((b - a) \ i)" + using that unfolding \_def by (force intro: cInf_le_finite) + have "1 / Inf \ \ ((b - a) \ i) / ((b - x) \ i)" using le_imp_inverse_le [OF * \] by (simp add: field_simps) with N have "k > ((b - a) \ i) / ((b - x) \ i)" using \N \ k\ by linarith with x that show ?thesis by (auto simp: mem_box algebra_simps field_split_simps) qed - show "x \ cbox (a + (b - a) /\<^sub>R k) (b - (b - a) /\<^sub>R k)" - using that \ \k > 0\ - by (auto simp: mem_box algebra_simps divide_inverse dest: xa_gt bx_gt) + show "x \ I k" + using that \ \k > 0\ unfolding I_def + by (auto simp: \_def \_def mem_box algebra_simps divide_inverse dest: xa_gt bx_gt) qed qed - obtain Bf where "Bf > 0" and Bf: "\c d. cbox c d \ box a b \ norm (integral (cbox c d) f) \ Bf" - using bo unfolding bounded_pos by blast - obtain Bg where "Bg > 0" and Bg:"\c d. cbox c d \ cbox a b \ \integral (cbox c d) g\ \ Bg" - using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_pos real_norm_def by blast + obtain Bf where Bf: "\c d. cbox c d \ box a b \ norm (integral (cbox c d) f) \ Bf" + using bo unfolding bounded_iff by blast + obtain Bg where Bg:"\c d. cbox c d \ cbox a b \ \integral (cbox c d) g\ \ Bg" + using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_iff real_norm_def by blast show "(\x. f x \ j) absolutely_integrable_on cbox a b" using g proof \ \A lot of duplication in the two proofs\ assume fg [rule_format]: "\x\cbox a b. f x \ j \ g x" have "(\x. (f x \ j)) = (\x. g x - (g x - (f x \ j)))" by simp moreover have "(\x. g x - (g x - (f x \ j))) integrable_on cbox a b" proof (rule Henstock_Kurzweil_Integration.integrable_diff [OF int_gab]) - let ?\ = "\k x. if x \ cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k)) - then g x - f x \ j else 0" + define \ where "\ \ \k x. if x \ I (Suc k) then g x - f x \ j else 0" have "(\x. g x - f x \ j) integrable_on box a b" - proof (rule monotone_convergence_increasing [of ?\, THEN conjunct1]) - have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \ box a b - = cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k - using box_subset_cbox "1" by fastforce - show "?\ k integrable_on box a b" for k - apply (simp add: integrable_restrict_Int integral_restrict_Int *) - apply (rule integrable_diff [OF integrable_on_subcbox [OF int_gab]]) - using "*" box_subset_cbox apply blast - by (metis "1" int_f integrable_component of_nat_Suc) - have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) - \ cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k - using False content_eq_0 - apply (simp add: subset_box algebra_simps) - apply (simp add: divide_simps) - apply (fastforce simp: field_simps) - done - show "?\ k x \ ?\ (Suc k) x" if "x \ box a b" for k x - using cb12 box_subset_cbox that by (force simp: intro!: fg) - show "(\k. ?\ k x) \ g x - f x \ j" if x: "x \ box a b" for x + proof (rule monotone_convergence_increasing [of \, THEN conjunct1]) + have *: "I (Suc k) \ box a b = I (Suc k)" for k + using box_subset_cbox ISuc_box by fastforce + show "\ k integrable_on box a b" for k + proof - + have "I (Suc k) \ cbox a b" + using "*" box_subset_cbox by blast + moreover have "(\m. f m \ j) integrable_on I (Suc k)" + by (metis ISuc_box I_def int_f integrable_component) + ultimately have "(\m. g m - f m \ j) integrable_on I (Suc k)" + by (metis Henstock_Kurzweil_Integration.integrable_diff I_def int_gab integrable_on_subcbox) + then show ?thesis + by (simp add: "*" \_def integrable_restrict_Int) + qed + show "\ k x \ \ (Suc k) x" if "x \ box a b" for k x + using ISucSuc box_subset_cbox that by (force simp: \_def intro!: fg) + show "(\k. \ k x) \ g x - f x \ j" if x: "x \ box a b" for x proof (rule tendsto_eventually) - obtain N::nat where N: "\k. k \ N \ x \ cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)" + obtain N::nat where N: "\k. k \ N \ x \ I k" using getN [OF x] by blast - show "\\<^sub>F k in sequentially. ?\ k x = g x - f x \ j" + show "\\<^sub>F k in sequentially. \ k x = g x - f x \ j" proof fix k::nat assume "N \ k" - have "x \ cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))" + have "x \ I (Suc k)" by (metis \N \ k\ le_Suc_eq N) - then show "?\ k x = g x - f x \ j" - by simp + then show "\ k x = g x - f x \ j" + by (simp add: \_def) qed qed - have "\integral (box a b) - (\x. if x \ cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) - then g x - f x \ j else 0)\ \ Bg + Bf" for k + have "\integral (box a b) (\x. if x \ I (Suc k) then g x - f x \ j else 0)\ \ Bg + Bf" for k proof - - let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" - have I_int [simp]: "?I \ box a b = ?I" - using 1 by (simp add: Int_absorb2) - have int_fI: "f integrable_on ?I" - by (metis I_int inf_le2 int_f) - then have "(\x. f x \ j) integrable_on ?I" - by (simp add: integrable_component) - moreover have "g integrable_on ?I" - by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox) + have ABK_def [simp]: "I (Suc k) \ box a b = I (Suc k)" + using ISuc_box by (simp add: Int_absorb2) + have int_fI: "f integrable_on I (Suc k)" + using ISuc_box I_def int_f by auto moreover - have "\integral ?I (\x. f x \ j)\ \ norm (integral ?I f)" + have "\integral (I (Suc k)) (\x. f x \ j)\ \ norm (integral (I (Suc k)) f)" by (simp add: Basis_le_norm int_fI \j \ Basis\) - with 1 I_int have "\integral ?I (\x. f x \ j)\ \ Bf" - by (blast intro: order_trans [OF _ Bf]) - ultimately show ?thesis - apply (simp add: integral_restrict_Int integral_diff) - using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) + with ISuc_box ABK_def have "\integral (I (Suc k)) (\x. f x \ j)\ \ Bf" + by (metis Bf I_def \j \ Basis\ int_fI integral_component_eq norm_bound_Basis_le) + ultimately + have "\integral (I (Suc k)) g - integral (I (Suc k)) (\x. f x \ j)\ \ Bg + Bf" + using "*" box_subset_cbox unfolding I_def + by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) + moreover have "g integrable_on I (Suc k)" + by (metis ISuc_box I_def int_gab integrable_on_open_interval integrable_on_subcbox) + moreover have "(\x. f x \ j) integrable_on I (Suc k)" + using int_fI by (simp add: integrable_component) + ultimately show ?thesis + by (simp add: integral_restrict_Int integral_diff) qed - then show "bounded (range (\k. integral (box a b) (?\ k)))" - apply (simp add: bounded_pos) - apply (rule_tac x="Bg+Bf" in exI) - using \0 < Bf\ \0 < Bg\ apply auto - done + then show "bounded (range (\k. integral (box a b) (\ k)))" + by (auto simp add: bounded_iff \_def) qed then show "(\x. g x - f x \ j) integrable_on cbox a b" by (simp add: integrable_on_open_interval) qed ultimately have "(\x. f x \ j) integrable_on cbox a b" by auto then show ?thesis - apply (rule absolutely_integrable_component_ubound [OF _ absint_g]) - by (simp add: fg) + using absolutely_integrable_component_ubound [OF _ absint_g] fg by force next assume gf [rule_format]: "\x\cbox a b. g x \ f x \ j" have "(\x. (f x \ j)) = (\x. ((f x \ j) - g x) + g x)" by simp moreover have "(\x. (f x \ j - g x) + g x) integrable_on cbox a b" proof (rule Henstock_Kurzweil_Integration.integrable_add [OF _ int_gab]) - let ?\ = "\k x. if x \ cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k)) - then f x \ j - g x else 0" + let ?\ = "\k x. if x \ I(Suc k) then f x \ j - g x else 0" have "(\x. f x \ j - g x) integrable_on box a b" proof (rule monotone_convergence_increasing [of ?\, THEN conjunct1]) - have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \ box a b - = cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k - using box_subset_cbox "1" by fastforce + have *: "I (Suc k) \ box a b = I (Suc k)" for k + using box_subset_cbox ISuc_box by fastforce show "?\ k integrable_on box a b" for k - apply (simp add: integrable_restrict_Int integral_restrict_Int *) - apply (rule integrable_diff) - apply (metis "1" int_f integrable_component of_nat_Suc) - apply (rule integrable_on_subcbox [OF int_gab]) - using "*" box_subset_cbox apply blast - done - have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) - \ cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k - using False content_eq_0 - apply (simp add: subset_box algebra_simps) - apply (simp add: divide_simps) - apply (fastforce simp: field_simps) - done + proof (simp add: integrable_restrict_Int integral_restrict_Int *) + show "(\x. f x \ j - g x) integrable_on I (Suc k)" + by (metis ISuc_box Henstock_Kurzweil_Integration.integrable_diff I_def int_f int_gab integrable_component integrable_on_open_interval integrable_on_subcbox) + qed show "?\ k x \ ?\ (Suc k) x" if "x \ box a b" for k x - using cb12 box_subset_cbox that by (force simp: intro!: gf) + using ISucSuc box_subset_cbox that by (force simp: I_def intro!: gf) show "(\k. ?\ k x) \ f x \ j - g x" if x: "x \ box a b" for x proof (rule tendsto_eventually) - obtain N::nat where N: "\k. k \ N \ x \ cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)" + obtain N::nat where N: "\k. k \ N \ x \ I k" using getN [OF x] by blast - show "\\<^sub>F k in sequentially. ?\ k x = f x \ j - g x" - proof - fix k::nat assume "N \ k" - have "x \ cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))" - by (metis \N \ k\ le_Suc_eq N) - then show "?\ k x = f x \ j - g x" - by simp - qed + then show "\\<^sub>F k in sequentially. ?\ k x = f x \ j - g x" + by (metis (no_types, lifting) eventually_at_top_linorderI le_Suc_eq) qed have "\integral (box a b) - (\x. if x \ cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) - then f x \ j - g x else 0)\ \ Bf + Bg" for k + (\x. if x \ I (Suc k) then f x \ j - g x else 0)\ \ Bf + Bg" for k proof - - let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" - have I_int [simp]: "?I \ box a b = ?I" - using 1 by (simp add: Int_absorb2) - have int_fI: "f integrable_on ?I" - by (simp add: inf.orderI int_f) - then have "(\x. f x \ j) integrable_on ?I" + define ABK where "ABK \ cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" + have ABK_eq [simp]: "ABK \ box a b = ABK" + using "*" I_def \_def \_def ABK_def by auto + have int_fI: "f integrable_on ABK" + unfolding ABK_def + using ISuc_box I_def \_def \_def int_f by force + then have "(\x. f x \ j) integrable_on ABK" by (simp add: integrable_component) - moreover have "g integrable_on ?I" - by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox) + moreover have "g integrable_on ABK" + by (metis ABK_def ABK_eq IntE box_subset_cbox int_gab integrable_on_subcbox subset_eq) moreover - have "\integral ?I (\x. f x \ j)\ \ norm (integral ?I f)" + have "\integral ABK (\x. f x \ j)\ \ norm (integral ABK f)" by (simp add: Basis_le_norm int_fI \j \ Basis\) - with 1 I_int have "\integral ?I (\x. f x \ j)\ \ Bf" - by (blast intro: order_trans [OF _ Bf]) + then have "\integral ABK (\x. f x \ j)\ \ Bf" + by (metis ABK_eq ABK_def Bf IntE dual_order.trans subset_eq) ultimately show ?thesis - apply (simp add: integral_restrict_Int integral_diff) - using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) + using "*" box_subset_cbox + apply (simp add: integral_restrict_Int integral_diff ABK_def I_def \_def \_def) + by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) qed then show "bounded (range (\k. integral (box a b) (?\ k)))" - apply (simp add: bounded_pos) - apply (rule_tac x="Bf+Bg" in exI) - using \0 < Bf\ \0 < Bg\ by auto + by (auto simp add: bounded_iff) qed then show "(\x. f x \ j - g x) integrable_on cbox a b" by (simp add: integrable_on_open_interval) qed ultimately have "(\x. f x \ j) integrable_on cbox a b" by auto then show ?thesis - apply (rule absolutely_integrable_component_lbound [OF absint_g]) - by (simp add: gf) + using absint_g absolutely_integrable_absolutely_integrable_lbound gf by blast qed qed qed -subsection\Second Mean Value Theorem\ - - - subsection\Second mean value theorem and corollaries\ lemma level_approx: fixes f :: "real \ real" and n::nat assumes f: "\x. x \ S \ 0 \ f x \ f x \ 1" and "x \ S" "n \ 0" shows "\f x - (\k = Suc 0..n. if k / n \ f x then inverse n else 0)\ < inverse n" (is "?lhs < _") proof - have "n * f x \ 0" using assms by auto then obtain m::nat where m: "floor(n * f x) = int m" using nonneg_int_cases zero_le_floor by blast then have kn: "real k / real n \ f x \ k \ m" for k using \n \ 0\ by (simp add: field_split_simps) linarith then have "Suc n / real n \ f x \ Suc n \ m" by blast have "real n * f x \ real n" by (simp add: \x \ S\ f mult_left_le) then have "m \ n" using m by linarith have "?lhs = \f x - (\k \ {Suc 0..n} \ {..m}. inverse n)\" by (subst sum.inter_restrict) (auto simp: kn) also have "\ < inverse n" using \m \ n\ \n \ 0\ m by (simp add: min_absorb2 field_split_simps) linarith finally show ?thesis . qed lemma SMVT_lemma2: fixes f :: "real \ real" assumes f: "f integrable_on {a..b}" and g: "\x y. x \ y \ g x \ g y" shows "(\y::real. {\x. if g x \ y then f x else 0}) equiintegrable_on {a..b}" proof - have ffab: "{f} equiintegrable_on {a..b}" by (metis equiintegrable_on_sing f interval_cbox) then have ff: "{f} equiintegrable_on (cbox a b)" by simp have ge: "(\c. {\x. if x \ c then f x else 0}) equiintegrable_on {a..b}" using equiintegrable_halfspace_restrictions_ge [OF ff] by auto have gt: "(\c. {\x. if x > c then f x else 0}) equiintegrable_on {a..b}" using equiintegrable_halfspace_restrictions_gt [OF ff] by auto have 0: "{(\x. 0)} equiintegrable_on {a..b}" by (metis box_real(2) equiintegrable_on_sing integrable_0) have \: "(\x. if g x \ y then f x else 0) \ {(\x. 0), f} \ (\z. {\x. if z < x then f x else 0}) \ (\z. {\x. if z \ x then f x else 0})" for y proof (cases "(\x. g x \ y) \ (\x. \ (g x \ y))") let ?\ = "Inf {x. g x \ y}" case False have lower: "?\ \ x" if "g x \ y" for x - apply (rule cInf_lower) - using False - apply (auto simp: that bdd_below_def) - by (meson dual_order.trans g linear) + proof (rule cInf_lower) + show "x \ {x. y \ g x}" + using False by (auto simp: that) + show "bdd_below {x. y \ g x}" + by (metis False bdd_belowI dual_order.trans g linear mem_Collect_eq) + qed have greatest: "?\ \ z" if "(\x. g x \ y \ z \ x)" for z by (metis False cInf_greatest empty_iff mem_Collect_eq that) show ?thesis proof (cases "g ?\ \ y") case True then obtain \ where \: "\x. g x \ y \ x \ \" by (metis g lower order.trans) \ \in fact y is @{term ?\}\ then show ?thesis by (force simp: \) next case False - have "(y \ g x) = (?\ < x)" for x + have "(y \ g x) \ (?\ < x)" for x proof show "?\ < x" if "y \ g x" using that False less_eq_real_def lower by blast show "y \ g x" if "?\ < x" by (metis g greatest le_less_trans that less_le_trans linear not_less) qed then obtain \ where \: "\x. g x \ y \ x > \" .. then show ?thesis by (force simp: \) qed qed auto show ?thesis - apply (rule equiintegrable_on_subset [OF equiintegrable_on_Un [OF gt equiintegrable_on_Un [OF ge equiintegrable_on_Un [OF ffab 0]]]]) - using \ apply (simp add: UN_subset_iff) - done + using \ by (simp add: UN_subset_iff equiintegrable_on_subset [OF equiintegrable_on_Un [OF gt equiintegrable_on_Un [OF ge equiintegrable_on_Un [OF ffab 0]]]]) qed lemma SMVT_lemma4: fixes f :: "real \ real" assumes f: "f integrable_on {a..b}" and "a \ b" and g: "\x y. x \ y \ g x \ g y" and 01: "\x. \a \ x; x \ b\ \ 0 \ g x \ g x \ 1" obtains c where "a \ c" "c \ b" "((\x. g x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}" proof - have "connected ((\x. integral {x..b} f) ` {a..b})" by (simp add: f indefinite_integral_continuous_1' connected_continuous_image) moreover have "compact ((\x. integral {x..b} f) ` {a..b})" by (simp add: compact_continuous_image f indefinite_integral_continuous_1') ultimately obtain m M where int_fab: "(\x. integral {x..b} f) ` {a..b} = {m..M}" using connected_compact_interval_1 by meson have "\c. c \ {a..b} \ integral {c..b} f = integral {a..b} (\x. (\k = 1..n. if g x \ real k / real n then inverse n *\<^sub>R f x else 0))" for n proof (cases "n=0") case True then show ?thesis using \a \ b\ by auto next case False have "(\c::real. {\x. if g x \ c then f x else 0}) equiintegrable_on {a..b}" using SMVT_lemma2 [OF f g] . then have int: "(\x. if g x \ c then f x else 0) integrable_on {a..b}" for c by (simp add: equiintegrable_on_def) have int': "(\x. if g x \ c then u * f x else 0) integrable_on {a..b}" for c u proof - have "(\x. if g x \ c then u * f x else 0) = (\x. u * (if g x \ c then f x else 0))" by (force simp: if_distrib) then show ?thesis using integrable_on_cmult_left [OF int] by simp qed have "\d. d \ {a..b} \ integral {a..b} (\x. if g x \ y then f x else 0) = integral {d..b} f" for y proof - let ?X = "{x. g x \ y}" - have *: "\a. ?X = {x. a \ x} \ ?X = {x. a < x}" + have *: "\a. ?X = {a..} \ ?X = {a<..}" if 1: "?X \ {}" and 2: "?X \ UNIV" proof - let ?\ = "Inf{x. g x \ y}" have lower: "?\ \ x" if "g x \ y" for x - apply (rule cInf_lower) - using 1 2 apply (auto simp: that bdd_below_def) - by (meson dual_order.trans g linear) - have greatest: "?\ \ z" if "(\x. g x \ y \ z \ x)" for z + proof (rule cInf_lower) + show "x \ {x. y \ g x}" + using 1 2 by (auto simp: that) + show "bdd_below {x. y \ g x}" + unfolding bdd_below_def + by (metis "2" UNIV_eq_I dual_order.trans g less_eq_real_def mem_Collect_eq not_le) + qed + have greatest: "?\ \ z" if "\x. g x \ y \ z \ x" for z by (metis cInf_greatest mem_Collect_eq that 1) show ?thesis proof (cases "g ?\ \ y") case True then obtain \ where \: "\x. g x \ y \ x \ \" by (metis g lower order.trans) \ \in fact y is @{term ?\}\ then show ?thesis by (force simp: \) next case False have "(y \ g x) = (?\ < x)" for x proof show "?\ < x" if "y \ g x" using that False less_eq_real_def lower by blast show "y \ g x" if "?\ < x" by (metis g greatest le_less_trans that less_le_trans linear not_less) qed then obtain \ where \: "\x. g x \ y \ x > \" .. then show ?thesis by (force simp: \) qed qed - then consider "?X = {}" | "?X = UNIV" | d where "?X = {x. d \ x} \?X = {x. d < x}" + then consider "?X = {}" | "?X = UNIV" | (intv) d where "?X = {d..} \ ?X = {d<..}" by metis then have "\d. d \ {a..b} \ integral {a..b} (\x. if x \ ?X then f x else 0) = integral {d..b} f" proof cases - case 1 - then show ?thesis - using \a \ b\ by auto - next - case 2 - then show ?thesis - using \a \ b\ by auto - next - case (3 d) + case (intv d) show ?thesis proof (cases "d < a") case True - with 3 show ?thesis - apply (rule_tac x=a in exI) - apply (auto simp: \a \ b\ iff del: mem_Collect_eq intro!: Henstock_Kurzweil_Integration.integral_cong, simp_all) - done + with intv have "integral {a..b} (\x. if y \ g x then f x else 0) = integral {a..b} f" + by (intro Henstock_Kurzweil_Integration.integral_cong) force + then show ?thesis + by (rule_tac x=a in exI) (simp add: \a \ b\) next case False show ?thesis proof (cases "b < d") case True have "integral {a..b} (\x. if x \ {x. y \ g x} then f x else 0) = integral {a..b} (\x. 0)" - by (rule Henstock_Kurzweil_Integration.integral_cong) (use 3 True in fastforce) + by (rule Henstock_Kurzweil_Integration.integral_cong) (use intv True in fastforce) then show ?thesis using \a \ b\ by auto next case False - with \\ d < a\ have eq: "Collect ((\) d) \ {a..b} = {d..b}" "Collect ((<) d) \ {a..b} = {d<..b}" + with \\ d < a\ have eq: "{d..} \ {a..b} = {d..b}" "{d<..} \ {a..b} = {d<..b}" by force+ moreover have "integral {d<..b} f = integral {d..b} f" by (rule integral_spike_set [OF empty_imp_negligible negligible_subset [OF negligible_sing [of d]]]) auto - ultimately - show ?thesis - using \\ d < a\ False 3 - apply (rule_tac x=d in exI) - apply (auto simp: \a \ b\ iff del: mem_Collect_eq; subst Henstock_Kurzweil_Integration.integral_restrict_Int) - apply (simp_all add: \a \ b\ not_less eq) - done + ultimately + have "integral {a..b} (\x. if x \ {x. y \ g x} then f x else 0) = integral {d..b} f" + unfolding integral_restrict_Int using intv by presburger + moreover have "d \ {a..b}" + using \\ d < a\ \a \ b\ False by auto + ultimately show ?thesis + by auto qed qed - qed + qed (use \a \ b\ in auto) then show ?thesis by auto qed then have "\k. \d. d \ {a..b} \ integral {a..b} (\x. if real k / real n \ g x then f x else 0) = integral {d..b} f" by meson then obtain d where dab: "\k. d k \ {a..b}" and deq: "\k::nat. integral {a..b} (\x. if k/n \ g x then f x else 0) = integral {d k..b} f" by metis have "(\k = 1..n. integral {a..b} (\x. if real k / real n \ g x then f x else 0)) /\<^sub>R n \ {m..M}" unfolding scaleR_right.sum proof (intro conjI allI impI convex [THEN iffD1, rule_format]) show "integral {a..b} (\xa. if real k / real n \ g xa then f xa else 0) \ {m..M}" for k by (metis (no_types, lifting) deq image_eqI int_fab dab) qed (use \n \ 0\ in auto) then have "\c. c \ {a..b} \ integral {c..b} f = inverse n *\<^sub>R (\k = 1..n. integral {a..b} (\x. if g x \ real k / real n then f x else 0))" by (metis (no_types, lifting) int_fab imageE) then show ?thesis by (simp add: sum_distrib_left if_distrib integral_sum int' flip: integral_mult_right cong: if_cong) qed then obtain c where cab: "\n. c n \ {a..b}" and c: "\n. integral {c n..b} f = integral {a..b} (\x. (\k = 1..n. if g x \ real k / real n then f x /\<^sub>R n else 0))" by metis obtain d and \ :: "nat\nat" where "d \ {a..b}" and \: "strict_mono \" and d: "(c \ \) \ d" and non0: "\n. \ n \ Suc 0" proof - have "compact{a..b}" by auto with cab obtain d and s0 where "d \ {a..b}" and s0: "strict_mono s0" and tends: "(c \ s0) \ d" unfolding compact_def using that by blast show thesis proof show "d \ {a..b}" by fact show "strict_mono (s0 \ Suc)" using s0 by (auto simp: strict_mono_def) show "(c \ (s0 \ Suc)) \ d" by (metis tends LIMSEQ_subseq_LIMSEQ Suc_less_eq comp_assoc strict_mono_def) show "\n. (s0 \ Suc) n \ Suc 0" by (metis comp_apply le0 not_less_eq_eq old.nat.exhaust s0 seq_suble) qed qed define \ where "\ \ \n x. \k = Suc 0..\ n. if k/(\ n) \ g x then f x /\<^sub>R (\ n) else 0" define \ where "\ \ \n x. \k = Suc 0..\ n. if k/(\ n) \ g x then inverse (\ n) else 0" have **: "(\x. g x *\<^sub>R f x) integrable_on cbox a b \ (\n. integral (cbox a b) (\ n)) \ integral (cbox a b) (\x. g x *\<^sub>R f x)" proof (rule equiintegrable_limit) have \: "((\n. \x. (\k = Suc 0..n. if k / n \ g x then inverse n *\<^sub>R f x else 0)) ` {Suc 0..}) equiintegrable_on {a..b}" proof - have *: "(\c::real. {\x. if g x \ c then f x else 0}) equiintegrable_on {a..b}" using SMVT_lemma2 [OF f g] . show ?thesis apply (rule equiintegrable_on_subset [OF equiintegrable_sum_real [OF *]], clarify) apply (rule_tac a="{Suc 0..n}" in UN_I, force) apply (rule_tac a="\k. inverse n" in UN_I, auto) apply (rule_tac x="\k x. if real k / real n \ g x then f x else 0" in bexI) apply (force intro: sum.cong)+ done qed show "range \ equiintegrable_on cbox a b" unfolding \_def by (auto simp: non0 intro: equiintegrable_on_subset [OF \]) show "(\n. \ n x) \ g x *\<^sub>R f x" if x: "x \ cbox a b" for x proof - have eq: "\ n x = \ n x *\<^sub>R f x" for n by (auto simp: \_def \_def sum_distrib_right if_distrib intro: sum.cong) show ?thesis unfolding eq proof (rule tendsto_scaleR [OF _ tendsto_const]) show "(\n. \ n x) \ g x" unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e :: real assume "e > 0" then obtain N where "N \ 0" "0 < inverse (real N)" and N: "inverse (real N) < e" using real_arch_inverse by metis moreover have "\\ n x - g x\ < inverse (real N)" if "n\N" for n proof - have "\g x - \ n x\ < inverse (real (\ n))" unfolding \_def proof (rule level_approx [of "{a..b}" g]) show "\ n \ 0" by (metis Suc_n_not_le_n non0) qed (use x 01 non0 in auto) also have "\ \ inverse N" using seq_suble [OF \] \N \ 0\ non0 that by (auto intro: order_trans simp: field_split_simps) finally show ?thesis by linarith qed ultimately show "\N. \n\N. \\ n x - g x\ < e" using less_trans by blast qed qed qed qed show thesis proof show "a \ d" "d \ b" using \d \ {a..b}\ atLeastAtMost_iff by blast+ show "((\x. g x *\<^sub>R f x) has_integral integral {d..b} f) {a..b}" unfolding has_integral_iff proof show "(\x. g x *\<^sub>R f x) integrable_on {a..b}" using ** by simp show "integral {a..b} (\x. g x *\<^sub>R f x) = integral {d..b} f" proof (rule tendsto_unique) show "(\n. integral {c(\ n)..b} f) \ integral {a..b} (\x. g x *\<^sub>R f x)" using ** by (simp add: c \_def) - show "(\n. integral {c(\ n)..b} f) \ integral {d..b} f" - using indefinite_integral_continuous_1' [OF f] - using d unfolding o_def - apply (simp add: continuous_on_eq_continuous_within) - apply (drule_tac x=d in bspec) - using \d \ {a..b}\ apply blast - apply (simp add: continuous_within_sequentially o_def) - using cab by auto + have "continuous (at d within {a..b}) (\x. integral {x..b} f)" + using indefinite_integral_continuous_1' [OF f] \d \ {a..b}\ + by (simp add: continuous_on_eq_continuous_within) + then show "(\n. integral {c(\ n)..b} f) \ integral {d..b} f" + using d cab unfolding o_def + by (simp add: continuous_within_sequentially o_def) qed auto qed qed qed theorem second_mean_value_theorem_full: fixes f :: "real \ real" assumes f: "f integrable_on {a..b}" and "a \ b" and g: "\x y. \a \ x; x \ y; y \ b\ \ g x \ g y" obtains c where "c \ {a..b}" and "((\x. g x * f x) has_integral (g a * integral {a..c} f + g b * integral {c..b} f)) {a..b}" proof - have gab: "g a \ g b" using \a \ b\ g by blast then consider "g a < g b" | "g a = g b" by linarith then show thesis proof cases case 1 define h where "h \ \x. if x < a then 0 else if b < x then 1 else (g x - g a) / (g b - g a)" obtain c where "a \ c" "c \ b" and c: "((\x. h x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}" proof (rule SMVT_lemma4 [OF f \a \ b\, of h]) show "h x \ h y" "0 \ h x \ h x \ 1" if "x \ y" for x y using that gab by (auto simp: divide_simps g h_def) qed show ?thesis proof show "c \ {a..b}" using \a \ c\ \c \ b\ by auto have I: "((\x. g x * f x - g a * f x) has_integral (g b - g a) * integral {c..b} f) {a..b}" proof (subst has_integral_cong) show "g x * f x - g a * f x = (g b - g a) * h x *\<^sub>R f x" if "x \ {a..b}" for x using 1 that by (simp add: h_def field_split_simps) show "((\x. (g b - g a) * h x *\<^sub>R f x) has_integral (g b - g a) * integral {c..b} f) {a..b}" using has_integral_mult_right [OF c, of "g b - g a"] . qed have II: "((\x. g a * f x) has_integral g a * integral {a..b} f) {a..b}" using has_integral_mult_right [where c = "g a", OF integrable_integral [OF f]] . have "((\x. g x * f x) has_integral (g b - g a) * integral {c..b} f + g a * integral {a..b} f) {a..b}" using has_integral_add [OF I II] by simp then show "((\x. g x * f x) has_integral g a * integral {a..c} f + g b * integral {c..b} f) {a..b}" by (simp add: algebra_simps flip: integral_combine [OF \a \ c\ \c \ b\ f]) qed next case 2 show ?thesis proof show "a \ {a..b}" by (simp add: \a \ b\) have "((\x. g x * f x) has_integral g a * integral {a..b} f) {a..b}" proof (rule has_integral_eq) show "((\x. g a * f x) has_integral g a * integral {a..b} f) {a..b}" using f has_integral_mult_right by blast show "g a * f x = g x * f x" if "x \ {a..b}" for x by (metis atLeastAtMost_iff g less_eq_real_def not_le that 2) qed then show "((\x. g x * f x) has_integral g a * integral {a..a} f + g b * integral {a..b} f) {a..b}" by (simp add: 2) qed qed qed corollary second_mean_value_theorem: fixes f :: "real \ real" assumes f: "f integrable_on {a..b}" and "a \ b" and g: "\x y. \a \ x; x \ y; y \ b\ \ g x \ g y" obtains c where "c \ {a..b}" - "integral {a..b} (\x. g x * f x) - = g a * integral {a..c} f + g b * integral {c..b} f" + "integral {a..b} (\x. g x * f x) = g a * integral {a..c} f + g b * integral {c..b} f" using second_mean_value_theorem_full [where g=g, OF assms] by (metis (full_types) integral_unique) end diff --git a/src/HOL/Analysis/Tagged_Division.thy b/src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy +++ b/src/HOL/Analysis/Tagged_Division.thy @@ -1,2602 +1,2602 @@ (* Title: HOL/Analysis/Tagged_Division.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP *) section \Tagged Divisions for Henstock-Kurzweil Integration\ theory Tagged_Division imports Topology_Euclidean_Space begin lemma sum_Sigma_product: assumes "finite S" and "\i. i \ S \ finite (T i)" shows "(\i\S. sum (x i) (T i)) = (\(i, j)\Sigma S T. x i j)" using assms proof induct case empty then show ?case by simp next case (insert a S) show ?case by (simp add: insert.hyps(1) insert.prems sum.Sigma) qed lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one subsection\<^marker>\tag unimportant\ \Sundries\ text\A transitive relation is well-founded if all initial segments are finite.\ lemma wf_finite_segments: assumes "irrefl r" and "trans r" and "\x. finite {y. (y, x) \ r}" shows "wf (r)" apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms) using acyclic_def assms irrefl_def trans_Restr by fastforce text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: fixes u::"'a::linordered_field" assumes "u \ v" "0 \ r" "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms using divide_le_eq_1 by fastforce then have "(r/s) * (v - u) \ 1 * (v - u)" by (meson diff_ge_0_iff_ge mult_right_mono \u \ v\) then show ?thesis by (simp add: field_simps) qed subsection \Some useful lemmas about intervals\ lemma interior_subset_union_intervals: assumes "i = cbox a b" and "j = cbox c d" and "interior j \ {}" and "i \ j \ S" and "interior i \ interior j = {}" shows "interior i \ interior S" proof - have "box a b \ cbox c d = {}" using Int_interval_mixed_eq_empty[of c d a b] assms unfolding interior_cbox by auto moreover have "box a b \ cbox c d \ S" apply (rule order_trans,rule box_subset_cbox) using assms by auto ultimately show ?thesis unfolding assms interior_cbox by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI) qed lemma interior_Union_subset_cbox: assumes "finite f" assumes f: "\s. s \ f \ \a b. s = cbox a b" "\s. s \ f \ interior s \ t" and t: "closed t" shows "interior (\f) \ t" proof - have [simp]: "s \ f \ closed s" for s using f by auto define E where "E = {s\f. interior s = {}}" then have "finite E" "E \ {s\f. interior s = {}}" using \finite f\ by auto then have "interior (\f) = interior (\(f - E))" proof (induction E rule: finite_subset_induct') case (insert s f') have "interior (\(f - insert s f') \ s) = interior (\(f - insert s f'))" using insert.hyps \finite f\ by (intro interior_closed_Un_empty_interior) auto also have "\(f - insert s f') \ s = \(f - f')" using insert.hyps by auto finally show ?case by (simp add: insert.IH) qed simp also have "\ \ \(f - E)" by (rule interior_subset) also have "\ \ t" proof (rule Union_least) fix s assume "s \ f - E" with f[of s] obtain a b where s: "s \ f" "s = cbox a b" "box a b \ {}" by (fastforce simp: E_def) have "closure (interior s) \ closure t" by (intro closure_mono f \s \ f\) with s \closed t\ show "s \ t" by simp qed finally show ?thesis . qed lemma Int_interior_Union_intervals: "\finite \; open S; \T. T\\ \ \a b. T = cbox a b; \T. T\\ \ S \ (interior T) = {}\ \ S \ interior (\\) = {}" using interior_Union_subset_cbox[of \ "UNIV - S"] by auto lemma interval_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "cbox a b \ {x. x\k \ c} = cbox a (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)" "cbox a b \ {x. x\k \ c} = cbox (\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) b" using assms by (rule_tac set_eqI; auto simp: mem_box)+ lemma interval_not_empty: "\i\Basis. a\i \ b\i \ cbox a b \ {}" by (simp add: box_ne_empty) subsection \Bounds on intervals where they exist\ definition\<^marker>\tag important\ interval_upperbound :: "('a::euclidean_space) set \ 'a" where "interval_upperbound s = (\i\Basis. (SUP x\s. x\i) *\<^sub>R i)" definition\<^marker>\tag important\ interval_lowerbound :: "('a::euclidean_space) set \ 'a" where "interval_lowerbound s = (\i\Basis. (INF x\s. x\i) *\<^sub>R i)" lemma interval_upperbound[simp]: "\i\Basis. a\i \ b\i \ interval_upperbound (cbox a b) = (b::'a::euclidean_space)" unfolding interval_upperbound_def euclidean_representation_sum cbox_def by (safe intro!: cSup_eq) auto lemma interval_lowerbound[simp]: "\i\Basis. a\i \ b\i \ interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" unfolding interval_lowerbound_def euclidean_representation_sum cbox_def by (safe intro!: cInf_eq) auto lemmas interval_bounds = interval_upperbound interval_lowerbound lemma fixes X::"real set" shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" by (auto simp: interval_upperbound_def interval_lowerbound_def) lemma interval_bounds'[simp]: assumes "cbox a b \ {}" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" using assms unfolding box_ne_empty by auto lemma interval_upperbound_Times: assumes "A \ {}" and "B \ {}" shows "interval_upperbound (A \ B) = (interval_upperbound A, interval_upperbound B)" proof- from assms have fst_image_times': "A = fst ` (A \ B)" by simp have "(\i\Basis. (SUP x\A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (SUP x\A. x \ i) *\<^sub>R i)" by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0) moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp have "(\i\Basis. (SUP x\A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (SUP x\B. x \ i) *\<^sub>R i)" by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0) ultimately show ?thesis unfolding interval_upperbound_def by (subst sum_Basis_prod_eq) (auto simp add: sum_prod) qed lemma interval_lowerbound_Times: assumes "A \ {}" and "B \ {}" shows "interval_lowerbound (A \ B) = (interval_lowerbound A, interval_lowerbound B)" proof- from assms have fst_image_times': "A = fst ` (A \ B)" by simp have "(\i\Basis. (INF x\A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (INF x\A. x \ i) *\<^sub>R i)" by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0) moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp have "(\i\Basis. (INF x\A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (INF x\B. x \ i) *\<^sub>R i)" by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0) ultimately show ?thesis unfolding interval_lowerbound_def by (subst sum_Basis_prod_eq) (auto simp add: sum_prod) qed subsection \The notion of a gauge --- simply an open set containing the point\ definition\<^marker>\tag important\ "gauge \ \ (\x. x \ \ x \ open (\ x))" lemma gaugeI: assumes "\x. x \ \ x" and "\x. open (\ x)" shows "gauge \" using assms unfolding gauge_def by auto lemma gaugeD[dest]: assumes "gauge \" shows "x \ \ x" and "open (\ x)" using assms unfolding gauge_def by auto lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" unfolding gauge_def by auto lemma gauge_ball[intro]: "0 < e \ gauge (\x. ball x e)" unfolding gauge_def by auto lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" by (rule gauge_ball) auto lemma gauge_Int[intro]: "gauge \1 \ gauge \2 \ gauge (\x. \1 x \ \2 x)" unfolding gauge_def by auto lemma gauge_reflect: fixes \ :: "'a::euclidean_space \ 'a set" shows "gauge \ \ gauge (\x. uminus ` \ (- x))" using equation_minus_iff by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus) lemma gauge_Inter: assumes "finite S" and "\u. u\S \ gauge (f u)" shows "gauge (\x. \{f u x | u. u \ S})" proof - have *: "\x. {f u x |u. u \ S} = (\u. f u x) ` S" by auto show ?thesis unfolding gauge_def unfolding * using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto qed lemma gauge_existence_lemma: "(\x. \d :: real. p x \ 0 < d \ q d x) \ (\x. \d>0. p x \ q d x)" by (metis zero_less_one) subsection \Attempt a systematic general set of "offset" results for components\ (*FIXME Restructure, the subsection consists only of 1 lemma *) lemma gauge_modify: assumes "(\S. open S \ open {x. f(x) \ S})" "gauge d" shows "gauge (\x. {y. f y \ d (f x)})" using assms unfolding gauge_def by force subsection \Divisions\ definition\<^marker>\tag important\ division_of (infixl "division'_of" 40) where "s division_of i \ finite s \ (\K\s. K \ i \ K \ {} \ (\a b. K = cbox a b)) \ (\K1\s. \K2\s. K1 \ K2 \ interior(K1) \ interior(K2) = {}) \ (\s = i)" lemma division_ofD[dest]: assumes "s division_of i" shows "finite s" and "\K. K \ s \ K \ i" and "\K. K \ s \ K \ {}" and "\K. K \ s \ \a b. K = cbox a b" and "\K1 K2. K1 \ s \ K2 \ s \ K1 \ K2 \ interior(K1) \ interior(K2) = {}" and "\s = i" using assms unfolding division_of_def by auto lemma division_ofI: assumes "finite s" and "\K. K \ s \ K \ i" and "\K. K \ s \ K \ {}" and "\K. K \ s \ \a b. K = cbox a b" and "\K1 K2. K1 \ s \ K2 \ s \ K1 \ K2 \ interior K1 \ interior K2 = {}" and "\s = i" shows "s division_of i" using assms unfolding division_of_def by auto lemma division_of_finite: "s division_of i \ finite s" by auto lemma division_of_self[intro]: "cbox a b \ {} \ {cbox a b} division_of (cbox a b)" unfolding division_of_def by auto lemma division_of_trivial[simp]: "s division_of {} \ s = {}" unfolding division_of_def by auto lemma division_of_sing[simp]: "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" (is "?l = ?r") proof assume ?r moreover { fix K assume "s = {{a}}" "K\s" then have "\x y. K = cbox x y" apply (rule_tac x=a in exI)+ apply force done } ultimately show ?l unfolding division_of_def cbox_sing by auto next assume ?l have "x = {a}" if "x \ s" for x by (metis \s division_of cbox a a\ cbox_sing division_ofD(2) division_ofD(3) subset_singletonD that) moreover have "s \ {}" using \s division_of cbox a a\ by auto ultimately show ?r unfolding cbox_sing by auto qed lemma elementary_empty: obtains p where "p division_of {}" unfolding division_of_trivial by auto lemma elementary_interval: obtains p where "p division_of (cbox a b)" by (metis division_of_trivial division_of_self) lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" unfolding division_of_def by auto lemma forall_in_division: "d division_of i \ (\x\d. P x) \ (\a b. cbox a b \ d \ P (cbox a b))" unfolding division_of_def by fastforce lemma cbox_division_memE: assumes \: "K \ \" "\ division_of S" obtains a b where "K = cbox a b" "K \ {}" "K \ S" using assms unfolding division_of_def by metis lemma division_of_subset: assumes "p division_of (\p)" and "q \ p" shows "q division_of (\q)" proof (rule division_ofI) note * = division_ofD[OF assms(1)] show "finite q" using "*"(1) assms(2) infinite_super by auto { fix k assume "k \ q" then have kp: "k \ p" using assms(2) by auto show "k \ \q" using \k \ q\ by auto show "\a b. k = cbox a b" using *(4)[OF kp] by auto show "k \ {}" using *(3)[OF kp] by auto } fix k1 k2 assume "k1 \ q" "k2 \ q" "k1 \ k2" then have **: "k1 \ p" "k2 \ p" "k1 \ k2" using assms(2) by auto show "interior k1 \ interior k2 = {}" using *(5)[OF **] by auto qed auto lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" unfolding division_of_def by auto lemma division_inter: fixes s1 s2 :: "'a::euclidean_space set" assumes "p1 division_of s1" and "p2 division_of s2" shows "{k1 \ k2 | k1 k2. k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" (is "?A' division_of _") proof - let ?A = "{s. s \ (\(k1,k2). k1 \ k2) ` (p1 \ p2) \ s \ {}}" have *: "?A' = ?A" by auto show ?thesis unfolding * proof (rule division_ofI) have "?A \ (\(x, y). x \ y) ` (p1 \ p2)" by auto moreover have "finite (p1 \ p2)" using assms unfolding division_of_def by auto ultimately show "finite ?A" by auto have *: "\s. \{x\s. x \ {}} = \s" by auto show "\?A = s1 \ s2" unfolding * using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto { fix k assume "k \ ?A" then obtain k1 k2 where k: "k = k1 \ k2" "k1 \ p1" "k2 \ p2" "k \ {}" by auto then show "k \ {}" by auto show "k \ s1 \ s2" using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] unfolding k by auto obtain a1 b1 where k1: "k1 = cbox a1 b1" using division_ofD(4)[OF assms(1) k(2)] by blast obtain a2 b2 where k2: "k2 = cbox a2 b2" using division_ofD(4)[OF assms(2) k(3)] by blast show "\a b. k = cbox a b" unfolding k k1 k2 unfolding Int_interval by auto } fix k1 k2 assume "k1 \ ?A" then obtain x1 y1 where k1: "k1 = x1 \ y1" "x1 \ p1" "y1 \ p2" "k1 \ {}" by auto assume "k2 \ ?A" then obtain x2 y2 where k2: "k2 = x2 \ y2" "x2 \ p1" "y2 \ p2" "k2 \ {}" by auto assume "k1 \ k2" then have th: "x1 \ x2 \ y1 \ y2" unfolding k1 k2 by auto have *: "interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {} \ interior (x1 \ y1) \ interior x1 \ interior (x1 \ y1) \ interior y1 \ interior (x2 \ y2) \ interior x2 \ interior (x2 \ y2) \ interior y2 \ interior (x1 \ y1) \ interior (x2 \ y2) = {}" by auto show "interior k1 \ interior k2 = {}" unfolding k1 k2 apply (rule *) using assms division_ofD(5) k1 k2(2) k2(3) th apply auto done qed qed lemma division_inter_1: assumes "d division_of i" and "cbox a (b::'a::euclidean_space) \ i" shows "{cbox a b \ k | k. k \ d \ cbox a b \ k \ {}} division_of (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis unfolding True and division_of_trivial by auto next case False have *: "cbox a b \ i = cbox a b" using assms(2) by auto show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed lemma elementary_Int: fixes s t :: "'a::euclidean_space set" assumes "p1 division_of s" and "p2 division_of t" shows "\p. p division_of (s \ t)" using assms division_inter by blast lemma elementary_Inter: assumes "finite f" and "f \ {}" and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" shows "\p. p division_of (\f)" using assms proof (induct f rule: finite_induct) case (insert x f) show ?case proof (cases "f = {}") case True then show ?thesis unfolding True using insert by auto next case False obtain p where "p division_of \f" using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. moreover obtain px where "px division_of x" using insert(5)[rule_format,OF insertI1] .. ultimately show ?thesis by (simp add: elementary_Int Inter_insert) qed qed auto lemma division_disjoint_union: assumes "p1 division_of s1" and "p2 division_of s2" and "interior s1 \ interior s2 = {}" shows "(p1 \ p2) division_of (s1 \ s2)" proof (rule division_ofI) note d1 = division_ofD[OF assms(1)] note d2 = division_ofD[OF assms(2)] show "finite (p1 \ p2)" using d1(1) d2(1) by auto show "\(p1 \ p2) = s1 \ s2" using d1(6) d2(6) by auto { fix k1 k2 assume as: "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" moreover let ?g="interior k1 \ interior k2 = {}" { assume as: "k1\p1" "k2\p2" have ?g using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]] using assms(3) by blast } moreover { assume as: "k1\p2" "k2\p1" have ?g using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]] using assms(3) by blast } ultimately show ?g using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto } fix k assume k: "k \ p1 \ p2" show "k \ s1 \ s2" using k d1(2) d2(2) by auto show "k \ {}" using k d1(3) d2(3) by auto show "\a b. k = cbox a b" using k d1(4) d2(4) by auto qed lemma partial_division_extend_1: fixes a b c d :: "'a::euclidean_space" assumes incl: "cbox c d \ cbox a b" and nonempty: "cbox c d \ {}" obtains p where "p division_of (cbox a b)" "cbox c d \ p" proof let ?B = "\f::'a\'a \ 'a. cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" define p where "p = ?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" show "cbox c d \ p" unfolding p_def by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) have ord: "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" if "i \ Basis" for i using incl nonempty that unfolding box_eq_empty subset_box by (auto simp: not_le) show "p division_of (cbox a b)" proof (rule division_ofI) show "finite p" unfolding p_def by (auto intro!: finite_PiE) { fix k assume "k \ p" then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" by (auto simp: p_def) then show "\a b. k = cbox a b" by auto have "k \ cbox a b \ k \ {}" proof (simp add: k box_eq_empty subset_box not_less, safe) fix i :: 'a assume i: "i \ Basis" with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" by (auto simp: PiE_iff) with i ord[of i] show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" by auto qed then show "k \ {}" "k \ cbox a b" by auto { fix l assume "l \ p" then obtain g where g: "g \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" by (auto simp: p_def) assume "l \ k" have "\i\Basis. f i \ g i" proof (rule ccontr) assume "\ ?thesis" with f g have "f = g" by (auto simp: PiE_iff extensional_def fun_eq_iff) with \l \ k\ show False by (simp add: l k) qed then obtain i where *: "i \ Basis" "f i \ g i" .. then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" using f g by (auto simp: PiE_iff) with * ord[of i] show "interior l \ interior k = {}" by (auto simp add: l k disjoint_interval intro!: bexI[of _ i]) } note \k \ cbox a b\ } moreover { fix x assume x: "x \ cbox a b" have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" proof fix i :: 'a assume "i \ Basis" with x ord[of i] have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ (d \ i \ x \ i \ x \ i \ b \ i)" by (auto simp: cbox_def) then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" by auto qed then obtain f where f: "\i\Basis. x \ i \ {fst (f i) \ i..snd (f i) \ i} \ f i \ {(a, c), (c, d), (d, b)}" unfolding bchoice_iff .. moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" by auto moreover from f have "x \ ?B (restrict f Basis)" by (auto simp: mem_box) ultimately have "\k\p. x \ k" unfolding p_def by blast } ultimately show "\p = cbox a b" by auto qed qed proposition partial_division_extend_interval: assumes "p division_of (\p)" "(\p) \ cbox a b" obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" proof (cases "p = {}") case True obtain q where "q division_of (cbox a b)" by (rule elementary_interval) then show ?thesis using True that by blast next case False note p = division_ofD[OF assms(1)] have div_cbox: "\k\p. \q. q division_of cbox a b \ k \ q" proof fix k assume kp: "k \ p" obtain c d where k: "k = cbox c d" using p(4)[OF kp] by blast have *: "cbox c d \ cbox a b" "cbox c d \ {}" using p(2,3)[OF kp, unfolded k] using assms(2) by (blast intro: order.trans)+ obtain q where "q division_of cbox a b" "cbox c d \ q" by (rule partial_division_extend_1[OF *]) then show "\q. q division_of cbox a b \ k \ q" unfolding k by auto qed obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" using bchoice[OF div_cbox] by blast have "q x division_of \(q x)" if x: "x \ p" for x apply (rule division_ofI) using division_ofD[OF q(1)[OF x]] by auto then have di: "\x. x \ p \ \d. d division_of \(q x - {x})" by (meson Diff_subset division_of_subset) have "\d. d division_of \((\i. \(q i - {i})) ` p)" apply (rule elementary_Inter [OF finite_imageI[OF p(1)]]) apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]]) done then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. have "d \ p division_of cbox a b" proof - have te: "\S f T. S \ {} \ \i\S. f i \ i = T \ T = \(f ` S) \ \S" by auto have cbox_eq: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" proof (rule te[OF False], clarify) fix i assume i: "i \ p" show "\(q i - {i}) \ i = cbox a b" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed { fix K assume K: "K \ p" note qk=division_ofD[OF q(1)[OF K]] have *: "\u T S. T \ S = {} \ u \ S \ u \ T = {}" by auto have "interior (\i\p. \(q i - {i})) \ interior K = {}" proof (rule *[OF Int_interior_Union_intervals]) show "\T. T\q K - {K} \ interior K \ interior T = {}" using qk(5) using q(2)[OF K] by auto show "interior (\i\p. \(q i - {i})) \ interior (\(q K - {K}))" apply (rule interior_mono)+ using K by auto qed (use qk in auto)} note [simp] = this show "d \ p division_of (cbox a b)" unfolding cbox_eq apply (rule division_disjoint_union[OF d assms(1)]) apply (rule Int_interior_Union_intervals) apply (rule p open_interior ballI)+ apply simp_all done qed then show ?thesis by (meson Un_upper2 that) qed lemma elementary_bounded[dest]: fixes S :: "'a::euclidean_space set" shows "p division_of S \ bounded S" unfolding division_of_def by (metis bounded_Union bounded_cbox) lemma elementary_subset_cbox: "p division_of S \ \a b. S \ cbox a (b::'a::euclidean_space)" by (meson bounded_subset_cbox_symmetric elementary_bounded) proposition division_union_intervals_exists: fixes a b :: "'a::euclidean_space" assumes "cbox a b \ {}" obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" proof (cases "cbox c d = {}") case True with assms that show ?thesis by force next case False show ?thesis proof (cases "cbox a b \ cbox c d = {}") case True then show ?thesis by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty) next case False obtain u v where uv: "cbox a b \ cbox c d = cbox u v" unfolding Int_interval by auto have uv_sub: "cbox u v \ cbox c d" using uv by auto obtain p where pd: "p division_of cbox c d" and "cbox u v \ p" by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) note p = this division_ofD[OF pd] have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" apply (rule arg_cong[of _ _ interior]) using p(8) uv by auto also have "\ = {}" unfolding interior_Int apply (rule Int_interior_Union_intervals) using p(6) p(7)[OF p(2)] \finite p\ apply auto done finally have [simp]: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" by simp have cbe: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" using p(8) unfolding uv[symmetric] by auto have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" proof - have "{cbox a b} division_of cbox a b" by (simp add: assms division_of_self) then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" by (metis (no_types) Diff_subset \interior (cbox a b) \ interior (\(p - {cbox u v})) = {}\ division_disjoint_union division_of_subset insert_is_Un p(1) p(8)) qed with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe) qed qed lemma division_of_Union: assumes "finite f" and "\p. p \ f \ p division_of (\p)" and "\k1 k2. k1 \ \f \ k2 \ \f \ k1 \ k2 \ interior k1 \ interior k2 = {}" shows "\f division_of \(\f)" using assms by (auto intro!: division_ofI) lemma elementary_union_interval: fixes a b :: "'a::euclidean_space" assumes "p division_of \p" obtains q where "q division_of (cbox a b \ \p)" proof (cases "p = {} \ cbox a b = {}") case True obtain p where "p division_of (cbox a b)" by (rule elementary_interval) then show thesis using True assms that by auto next case False then have "p \ {}" "cbox a b \ {}" by auto note pdiv = division_ofD[OF assms] show ?thesis proof (cases "interior (cbox a b) = {}") case True show ?thesis apply (rule that [of "insert (cbox a b) p", OF division_ofI]) using pdiv(1-4) True False apply auto apply (metis IntI equals0D pdiv(5)) by (metis disjoint_iff_not_equal pdiv(5)) next case False have "\K\p. \q. (insert (cbox a b) q) division_of (cbox a b \ K)" proof fix K assume kp: "K \ p" from pdiv(4)[OF kp] obtain c d where "K = cbox c d" by blast then show "\q. (insert (cbox a b) q) division_of (cbox a b \ K)" by (meson \cbox a b \ {}\ division_union_intervals_exists) qed from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. note q = division_ofD[OF this[rule_format]] let ?D = "\{insert (cbox a b) (q K) | K. K \ p}" show thesis proof (rule that[OF division_ofI]) have *: "{insert (cbox a b) (q K) |K. K \ p} = (\K. insert (cbox a b) (q K)) ` p" by auto show "finite ?D" using "*" pdiv(1) q(1) by auto have "\?D = (\x\p. \(insert (cbox a b) (q x)))" by auto also have "... = \{cbox a b \ t |t. t \ p}" using q(6) by auto also have "... = cbox a b \ \p" using \p \ {}\ by auto finally show "\?D = cbox a b \ \p" . show "K \ cbox a b \ \p" "K \ {}" if "K \ ?D" for K using q that by blast+ show "\a b. K = cbox a b" if "K \ ?D" for K using q(4) that by auto next fix K' K assume K: "K \ ?D" and K': "K' \ ?D" "K \ K'" obtain x where x: "K \ insert (cbox a b) (q x)" "x\p" using K by auto obtain x' where x': "K'\insert (cbox a b) (q x')" "x'\p" using K' by auto show "interior K \ interior K' = {}" proof (cases "x = x' \ K = cbox a b \ K' = cbox a b") case True then show ?thesis using True K' q(5) x' x by auto next case False then have as': "K \ cbox a b" "K' \ cbox a b" by auto obtain c d where K: "K = cbox c d" using q(4) x by blast have "interior K \ interior (cbox a b) = {}" using as' K'(2) q(5) x by blast then have "interior K \ interior x" by (metis \interior (cbox a b) \ {}\ K q(2) x interior_subset_union_intervals) moreover obtain c d where c_d: "K' = cbox c d" using q(4)[OF x'(2,1)] by blast have "interior K' \ interior (cbox a b) = {}" using as'(2) q(5) x' by blast then have "interior K' \ interior x'" by (metis \interior (cbox a b) \ {}\ c_d interior_subset_union_intervals q(2) x'(1) x'(2)) moreover have "interior x \ interior x' = {}" by (meson False assms division_ofD(5) x'(2) x(2)) ultimately show ?thesis using \interior K \ interior x\ \interior K' \ interior x'\ by auto qed qed qed qed lemma elementary_unions_intervals: assumes fin: "finite f" and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" obtains p where "p division_of (\f)" proof - have "\p. p division_of (\f)" proof (induct_tac f rule:finite_subset_induct) show "\p. p division_of \{}" using elementary_empty by auto next fix x F assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" from this(3) obtain p where p: "p division_of \F" .. from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast have *: "\F = \p" using division_ofD[OF p] by auto show "\p. p division_of \(insert x F)" using elementary_union_interval[OF p[unfolded *], of a b] unfolding Union_insert x * by metis qed (insert assms, auto) then show ?thesis using that by auto qed lemma elementary_union: fixes S T :: "'a::euclidean_space set" assumes "ps division_of S" "pt division_of T" obtains p where "p division_of (S \ T)" proof - have *: "S \ T = \ps \ \pt" using assms unfolding division_of_def by auto show ?thesis apply (rule elementary_unions_intervals[of "ps \ pt"]) using assms apply auto by (simp add: * that) qed lemma partial_division_extend: fixes T :: "'a::euclidean_space set" assumes "p division_of S" and "q division_of T" and "S \ T" obtains r where "p \ r" and "r division_of T" proof - note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] obtain a b where ab: "T \ cbox a b" using elementary_subset_cbox[OF assms(2)] by auto obtain r1 where "p \ r1" "r1 division_of (cbox a b)" using assms by (metis ab dual_order.trans partial_division_extend_interval divp(6)) note r1 = this division_ofD[OF this(2)] obtain p' where "p' division_of \(r1 - p)" apply (rule elementary_unions_intervals[of "r1 - p"]) using r1(3,6) apply auto done then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" by (metis assms(2) divq(6) elementary_Int) { fix x assume x: "x \ T" "x \ S" then obtain R where r: "R \ r1" "x \ R" unfolding r1 using ab by (meson division_contains r1(2) subsetCE) moreover have "R \ p" proof assume "R \ p" then have "x \ S" using divp(2) r by auto then show False using x by auto qed ultimately have "x\\(r1 - p)" by auto } then have Teq: "T = \p \ (\(r1 - p) \ \q)" unfolding divp divq using assms(3) by auto have "interior S \ interior (\(r1-p)) = {}" proof (rule Int_interior_Union_intervals) have *: "\S. (\x. x \ S \ False) \ S = {}" by auto show "interior S \ interior m = {}" if "m \ r1 - p" for m proof - have "interior m \ interior (\p) = {}" proof (rule Int_interior_Union_intervals) show "\T. T\p \ interior m \ interior T = {}" by (metis DiffD1 DiffD2 that r1(1) r1(7) rev_subsetD) qed (use divp in auto) then show "interior S \ interior m = {}" unfolding divp by auto qed qed (use r1 in auto) then have "interior S \ interior (\(r1-p) \ (\q)) = {}" using interior_subset by auto then have div: "p \ r2 division_of \p \ \(r1 - p) \ \q" by (simp add: assms(1) division_disjoint_union divp(6) r2) show ?thesis apply (rule that[of "p \ r2"]) apply (auto simp: div Teq) done qed lemma division_split: fixes a :: "'a::euclidean_space" assumes "p division_of (cbox a b)" and k: "k\Basis" shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" (is "?p1 division_of ?I1") and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of (cbox a b \ {x. x\k \ c})" (is "?p2 division_of ?I2") proof (rule_tac[!] division_ofI) note p = division_ofD[OF assms(1)] show "finite ?p1" "finite ?p2" using p(1) by auto show "\?p1 = ?I1" "\?p2 = ?I2" unfolding p(6)[symmetric] by auto { fix K assume "K \ ?p1" then obtain l where l: "K = l \ {x. x \ k \ c}" "l \ p" "l \ {x. x \ k \ c} \ {}" by blast obtain u v where uv: "l = cbox u v" using assms(1) l(2) by blast show "K \ ?I1" using l p(2) uv by force show "K \ {}" by (simp add: l) show "\a b. K = cbox a b" apply (simp add: l uv p(2-3)[OF l(2)]) apply (subst interval_split[OF k]) apply (auto intro: order.trans) done fix K' assume "K' \ ?p1" then obtain l' where l': "K' = l' \ {x. x \ k \ c}" "l' \ p" "l' \ {x. x \ k \ c} \ {}" by blast assume "K \ K'" then show "interior K \ interior K' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } { fix K assume "K \ ?p2" then obtain l where l: "K = l \ {x. c \ x \ k}" "l \ p" "l \ {x. c \ x \ k} \ {}" by blast obtain u v where uv: "l = cbox u v" using l(2) p(4) by blast show "K \ ?I2" using l p(2) uv by force show "K \ {}" by (simp add: l) show "\a b. K = cbox a b" apply (simp add: l uv p(2-3)[OF l(2)]) apply (subst interval_split[OF k]) apply (auto intro: order.trans) done fix K' assume "K' \ ?p2" then obtain l' where l': "K' = l' \ {x. c \ x \ k}" "l' \ p" "l' \ {x. c \ x \ k} \ {}" by blast assume "K \ K'" then show "interior K \ interior K' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } qed subsection \Tagged (partial) divisions\ definition\<^marker>\tag important\ tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where "s tagged_partial_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ (\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {})" -lemma tagged_partial_division_ofD[dest]: +lemma tagged_partial_division_ofD: assumes "s tagged_partial_division_of i" shows "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1,K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" using assms unfolding tagged_partial_division_of_def by blast+ definition\<^marker>\tag important\ tagged_division_of (infixr "tagged'_division'_of" 40) where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{K. \x. (x,K) \ s} = i)" lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" unfolding tagged_division_of_def tagged_partial_division_of_def by auto lemma tagged_division_of: "s tagged_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ (\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}) \ (\{K. \x. (x,K) \ s} = i)" unfolding tagged_division_of_def tagged_partial_division_of_def by auto lemma tagged_division_ofI: assumes "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1,K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" and "(\{K. \x. (x,K) \ s} = i)" shows "s tagged_division_of i" unfolding tagged_division_of using assms by fastforce lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) assumes "s tagged_division_of i" shows "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" and "(\{K. \x. (x,K) \ s} = i)" using assms unfolding tagged_division_of by blast+ lemma division_of_tagged_division: assumes "s tagged_division_of i" shows "(snd ` s) division_of i" proof (rule division_ofI) note assm = tagged_division_ofD[OF assms] show "\(snd ` s) = i" "finite (snd ` s)" using assm by auto fix k assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto then show "k \ i" "k \ {}" "\a b. k = cbox a b" using assm by fastforce+ fix k' assume k': "k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk': "(xk', k') \ s" by auto then show "interior k \ interior k' = {}" using assm(5) k'(2) xk by blast qed lemma partial_division_of_tagged_division: assumes "s tagged_partial_division_of i" shows "(snd ` s) division_of \(snd ` s)" proof (rule division_ofI) note assm = tagged_partial_division_ofD[OF assms] show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" using assm by auto fix k assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" using assm by auto fix k' assume k': "k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk': "(xk', k') \ s" by auto then show "interior k \ interior k' = {}" using assm(5) k'(2) xk by auto qed lemma tagged_partial_division_subset: assumes "s tagged_partial_division_of i" and "t \ s" shows "t tagged_partial_division_of i" using assms finite_subset[OF assms(2)] unfolding tagged_partial_division_of_def by blast lemma tag_in_interval: "p tagged_division_of i \ (x, k) \ p \ x \ i" by auto lemma tagged_division_of_empty: "{} tagged_division_of {}" unfolding tagged_division_of by auto lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \ p = {}" unfolding tagged_partial_division_of_def by auto lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \ p = {}" unfolding tagged_division_of by auto lemma tagged_division_of_self: "x \ cbox a b \ {(x,cbox a b)} tagged_division_of (cbox a b)" by (rule tagged_division_ofI) auto lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" unfolding box_real[symmetric] by (rule tagged_division_of_self) lemma tagged_division_Un: assumes "p1 tagged_division_of s1" and "p2 tagged_division_of s2" and "interior s1 \ interior s2 = {}" shows "(p1 \ p2) tagged_division_of (s1 \ s2)" proof (rule tagged_division_ofI) note p1 = tagged_division_ofD[OF assms(1)] note p2 = tagged_division_ofD[OF assms(2)] show "finite (p1 \ p2)" using p1(1) p2(1) by auto show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" using p1(6) p2(6) by blast fix x k assume xk: "(x, k) \ p1 \ p2" show "x \ k" "\a b. k = cbox a b" using xk p1(2,4) p2(2,4) by auto show "k \ s1 \ s2" using xk p1(3) p2(3) by blast fix x' k' assume xk': "(x', k') \ p1 \ p2" "(x, k) \ (x', k')" have *: "\a b. a \ s1 \ b \ s2 \ interior a \ interior b = {}" using assms(3) interior_mono by blast show "interior k \ interior k' = {}" apply (cases "(x, k) \ p1") apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2)) by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2)) qed lemma tagged_division_Union: assumes "finite I" and tag: "\i. i\I \ pfn i tagged_division_of i" and disj: "\i1 i2. \i1 \ I; i2 \ I; i1 \ i2\ \ interior(i1) \ interior(i2) = {}" shows "\(pfn ` I) tagged_division_of (\I)" proof (rule tagged_division_ofI) note assm = tagged_division_ofD[OF tag] show "finite (\(pfn ` I))" using assms by auto have "\{k. \x. (x, k) \ \(pfn ` I)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` I)" by blast also have "\ = \I" using assm(6) by auto finally show "\{k. \x. (x, k) \ \(pfn ` I)} = \I" . fix x k assume xk: "(x, k) \ \(pfn ` I)" then obtain i where i: "i \ I" "(x, k) \ pfn i" by auto show "x \ k" "\a b. k = cbox a b" "k \ \I" using assm(2-4)[OF i] using i(1) by auto fix x' k' assume xk': "(x', k') \ \(pfn ` I)" "(x, k) \ (x', k')" then obtain i' where i': "i' \ I" "(x', k') \ pfn i'" by auto have *: "\a b. i \ i' \ a \ i \ b \ i' \ interior a \ interior b = {}" using i(1) i'(1) disj interior_mono by blast show "interior k \ interior k' = {}" proof (cases "i = i'") case True then show ?thesis using assm(5) i' i xk'(2) by blast next case False then show ?thesis using "*" assm(3) i' i by auto qed qed lemma tagged_partial_division_of_Union_self: assumes "p tagged_partial_division_of s" shows "p tagged_division_of (\(snd ` p))" apply (rule tagged_division_ofI) using tagged_partial_division_ofD[OF assms] apply auto done lemma tagged_division_of_union_self: assumes "p tagged_division_of s" shows "p tagged_division_of (\(snd ` p))" apply (rule tagged_division_ofI) using tagged_division_ofD[OF assms] apply auto done lemma tagged_division_Un_interval: fixes a :: "'a::euclidean_space" assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of (cbox a b)" proof - have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" by auto show ?thesis apply (subst *) apply (rule tagged_division_Un[OF assms(1-2)]) unfolding interval_split[OF k] interior_cbox using k apply (auto simp add: box_def elim!: ballE[where x=k]) done qed lemma tagged_division_Un_interval_real: fixes a :: real assumes "p1 tagged_division_of ({a .. b} \ {x. x\k \ (c::real)})" and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of {a .. b}" using assms unfolding box_real[symmetric] by (rule tagged_division_Un_interval) lemma tagged_division_split_left_inj: assumes d: "d tagged_division_of i" and tags: "(x1, K1) \ d" "(x2, K2) \ d" and "K1 \ K2" and eq: "K1 \ {x. x \ k \ c} = K2 \ {x. x \ k \ c}" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) then show ?thesis using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) qed lemma tagged_division_split_right_inj: assumes d: "d tagged_division_of i" and tags: "(x1, K1) \ d" "(x2, K2) \ d" and "K1 \ K2" and eq: "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) then show ?thesis using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) qed lemma (in comm_monoid_set) over_tagged_division_lemma: assumes "p tagged_division_of i" and "\u v. box u v = {} \ d (cbox u v) = \<^bold>1" shows "F (\(_, k). d k) p = F d (snd ` p)" proof - have *: "(\(_ ,k). d k) = d \ snd" by (simp add: fun_eq_iff) note assm = tagged_division_ofD[OF assms(1)] show ?thesis unfolding * proof (rule reindex_nontrivial[symmetric]) show "finite p" using assm by auto fix x y assume "x\p" "y\p" "x\y" "snd x = snd y" obtain a b where ab: "snd x = cbox a b" using assm(4)[of "fst x" "snd x"] \x\p\ by auto have "(fst x, snd y) \ p" "(fst x, snd y) \ y" using \x \ p\ \x \ y\ \snd x = snd y\ [symmetric] by auto with \x\p\ \y\p\ have "interior (snd x) \ interior (snd y) = {}" by (intro assm(5)[of "fst x" _ "fst y"]) auto then have "box a b = {}" unfolding \snd x = snd y\[symmetric] ab by auto then have "d (cbox a b) = \<^bold>1" using assm(2)[of "fst x" "snd x"] \x\p\ ab[symmetric] by (intro assms(2)) auto then show "d (snd x) = \<^bold>1" unfolding ab by auto qed qed subsection \Functions closed on boxes: morphisms from boxes to monoids\ text \This auxiliary structure is used to sum up over the elements of a division. Main theorem is \operative_division\. Instances for the monoid are \<^typ>\'a option\, \<^typ>\real\, and \<^typ>\bool\.\ paragraph\<^marker>\tag important\ \Using additivity of lifted function to encode definedness.\ text \%whitespace\ definition\<^marker>\tag important\ lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" where "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\b. Some (f a b)))" lemma lift_option_simps[simp]: "lift_option f (Some a) (Some b) = Some (f a b)" "lift_option f None b' = None" "lift_option f a' None = None" by (auto simp: lift_option_def) lemma\<^marker>\tag important\ comm_monoid_lift_option: assumes "comm_monoid f z" shows "comm_monoid (lift_option f) (Some z)" proof - from assms interpret comm_monoid f z . show ?thesis by standard (auto simp: lift_option_def ac_simps split: bind_split) qed lemma comm_monoid_and: "comm_monoid HOL.conj True" by standard auto lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True" by (rule comm_monoid_set.intro) (fact comm_monoid_and) paragraph \Misc\ lemma interval_real_split: "{a .. b::real} \ {x. x \ c} = {a .. min b c}" "{a .. b} \ {x. c \ x} = {max a c .. b}" apply (metis Int_atLeastAtMostL1 atMost_def) apply (metis Int_atLeastAtMostL2 atLeast_def) done lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" by (meson zero_less_one) paragraph \Division points\ text \%whitespace\ definition\<^marker>\tag important\ "division_points (k::('a::euclidean_space) set) d = {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" lemma division_points_finite: fixes i :: "'a::euclidean_space set" assumes "d division_of i" shows "finite (division_points i d)" proof - note assm = division_ofD[OF assms] let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" have *: "division_points i d = \(?M ` Basis)" unfolding division_points_def by auto show ?thesis unfolding * using assm by auto qed lemma division_points_subset: fixes a :: "'a::euclidean_space" assumes "d division_of (cbox a b)" and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and k: "k \ Basis" shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is ?t1) and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ \(l \ {x. x\k \ c} = {})} \ division_points (cbox a b) d" (is ?t2) proof - note assm = division_ofD[OF assms(1)] have *: "\i\Basis. a\i \ b\i" "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" "min (b \ k) c = c" "max (a \ k) c = c" using assms using less_imp_le by auto have "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" if "a \ x < y" "y < (if x = k then c else b \ x)" "interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" "x \ Basis" for i l x y proof - obtain u v where l: "l = cbox u v" using \l \ d\ assms(1) by blast have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" using that(6) unfolding l interval_split[OF k] box_ne_empty that . have **: "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto show ?thesis apply (rule bexI[OF _ \l \ d\]) using that(1-3,5) \x \ Basis\ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that apply (auto split: if_split_asm) done qed moreover have "\x y. \y < (if x = k then c else b \ x)\ \ y < b \ x" using \c < b\k\ by (auto split: if_split_asm) ultimately show ?t1 unfolding division_points_def interval_split[OF k, of a b] unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * by force have "\x y i l. (if x = k then c else a \ x) < y \ a \ x < y" using \a\k < c\ by (auto split: if_split_asm) moreover have "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" if "(if x = k then c else a \ x) < y" "y < b \ x" "interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" "x \ Basis" for x y i l proof - obtain u v where l: "l = cbox u v" using \l \ d\ assm(4) by blast have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" using that(6) unfolding l interval_split[OF k] box_ne_empty that . have **: "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto show "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" apply (rule bexI[OF _ \l \ d\]) using that(1-3,5) \x \ Basis\ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that apply (auto split: if_split_asm) done qed ultimately show ?t2 unfolding division_points_def interval_split[OF k, of a b] unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * by force qed lemma division_points_psubset: fixes a :: "'a::euclidean_space" assumes d: "d division_of (cbox a b)" and altb: "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and "l \ d" and disj: "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k: "k \ Basis" shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is "?D1 \ ?D") and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is "?D2 \ ?D") proof - have ab: "\i\Basis. a\i \ b\i" using altb by (auto intro!:less_imp_le) obtain u v where l: "l = cbox u v" using d \l \ d\ by blast have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l) by (metis assms(5) box_ne_empty(1) cbox_division_memE d l subset_box(1)) have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" unfolding l interval_split[OF k] interval_bounds[OF uv(1)] using uv[rule_format, of k] ab k by auto have "\x. x \ ?D - ?D1" using assms(3-) unfolding division_points_def interval_bounds[OF ab] by (force simp add: *) moreover have "?D1 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D1 \ ?D" by blast have *: "interval_lowerbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" "interval_lowerbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" unfolding l interval_split[OF k] interval_bounds[OF uv(1)] using uv[rule_format, of k] ab k by auto have "\x. x \ ?D - ?D2" using assms(3-) unfolding division_points_def interval_bounds[OF ab] by (force simp add: *) moreover have "?D2 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D2 \ ?D" by blast qed lemma division_split_left_inj: fixes S :: "'a::euclidean_space set" assumes div: "\ division_of S" and eq: "K1 \ {x::'a. x\k \ c} = K2 \ {x. x\k \ c}" and "K1 \ \" "K2 \ \" "K1 \ K2" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior K2 \ interior {a. a \ k \ c} = interior K1 \ interior {a. a \ k \ c}" by (metis (no_types) eq interior_Int) moreover have "\A. interior A \ interior K2 = {} \ A = K2 \ A \ \" by (meson div \K2 \ \\ division_of_def) ultimately show ?thesis using \K1 \ \\ \K1 \ K2\ by auto qed lemma division_split_right_inj: fixes S :: "'a::euclidean_space set" assumes div: "\ division_of S" and eq: "K1 \ {x::'a. x\k \ c} = K2 \ {x. x\k \ c}" and "K1 \ \" "K2 \ \" "K1 \ K2" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior K2 \ interior {a. a \ k \ c} = interior K1 \ interior {a. a \ k \ c}" by (metis (no_types) eq interior_Int) moreover have "\A. interior A \ interior K2 = {} \ A = K2 \ A \ \" by (meson div \K2 \ \\ division_of_def) ultimately show ?thesis using \K1 \ \\ \K1 \ K2\ by auto qed lemma interval_doublesplit: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "cbox a b \ {x . \x\k - c\ \ (e::real)} = cbox (\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)" proof - have *: "\x c e::real. \x - c\ \ e \ x \ c - e \ x \ c + e" by auto have **: "\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" by blast show ?thesis unfolding * ** interval_split[OF assms] by (rule refl) qed lemma division_doublesplit: fixes a :: "'a::euclidean_space" assumes "p division_of (cbox a b)" and k: "k \ Basis" shows "(\l. l \ {x. \x\k - c\ \ e}) ` {l\p. l \ {x. \x\k - c\ \ e} \ {}} division_of (cbox a b \ {x. \x\k - c\ \ e})" proof - have *: "\x c. \x - c\ \ e \ x \ c - e \ x \ c + e" by auto have **: "\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" by auto note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] note division_split(2)[OF this, where c="c-e" and k=k,OF k] then show ?thesis apply (rule **) subgoal apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image [symmetric] cong: image_cong_simp) apply (rule equalityI) apply blast apply clarsimp apply (rule_tac x="xa \ {x. c + e \ x \ k}" in exI) apply auto done by (simp add: interval_split k interval_doublesplit) qed paragraph \Operative\ locale operative = comm_monoid_set + fixes g :: "'b::euclidean_space set \ 'a" assumes box_empty_imp: "\a b. box a b = {} \ g (cbox a b) = \<^bold>1" and Basis_imp: "\a b c k. k \ Basis \ g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c})" begin lemma empty [simp]: "g {} = \<^bold>1" proof - have *: "cbox One (-One) = ({}::'b set)" by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv) moreover have "box One (-One) = ({}::'b set)" using box_subset_cbox[of One "-One"] by (auto simp: *) ultimately show ?thesis using box_empty_imp [of One "-One"] by simp qed lemma division: "F g d = g (cbox a b)" if "d division_of (cbox a b)" proof - define C where [abs_def]: "C = card (division_points (cbox a b) d)" then show ?thesis using that proof (induction C arbitrary: a b d rule: less_induct) case (less a b d) show ?case proof cases assume "box a b = {}" { fix k assume "k\d" then obtain a' b' where k: "k = cbox a' b'" using division_ofD(4)[OF less.prems] by blast with \k\d\ division_ofD(2)[OF less.prems] have "cbox a' b' \ cbox a b" by auto then have "box a' b' \ box a b" unfolding subset_box by auto then have "g k = \<^bold>1" using box_empty_imp [of a' b'] k by (simp add: \box a b = {}\) } then show "box a b = {} \ F g d = g (cbox a b)" by (auto intro!: neutral simp: box_empty_imp) next assume "box a b \ {}" then have ab: "\i\Basis. a\i < b\i" and ab': "\i\Basis. a\i \ b\i" by (auto simp: box_ne_empty) show "F g d = g (cbox a b)" proof (cases "division_points (cbox a b) d = {}") case True { fix u v and j :: 'b assume j: "j \ Basis" and as: "cbox u v \ d" then have "cbox u v \ {}" using less.prems by blast then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" using j unfolding box_ne_empty by auto have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q (cbox u v)" using as j by auto have "(j, u\j) \ division_points (cbox a b) d" "(j, v\j) \ division_points (cbox a b) d" using True by auto note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] moreover have "a\j \ u\j" "v\j \ b\j" using division_ofD(2,2,3)[OF \d division_of cbox a b\ as] apply (metis j subset_box(1) uv(1)) by (metis \cbox u v \ cbox a b\ j subset_box(1) uv(1)) ultimately have "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } then have d': "\i\d. \u v. i = cbox u v \ (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" unfolding forall_in_division[OF less.prems] by blast have "(1/2) *\<^sub>R (a+b) \ cbox a b" unfolding mem_box using ab by (auto simp: inner_simps) note this[unfolded division_ofD(6)[OF \d division_of cbox a b\,symmetric] Union_iff] then obtain i where i: "i \ d" "(1 / 2) *\<^sub>R (a + b) \ i" .. obtain u v where uv: "i = cbox u v" "\j\Basis. u \ j = a \ j \ v \ j = a \ j \ u \ j = b \ j \ v \ j = b \ j \ u \ j = a \ j \ v \ j = b \ j" using d' i(1) by auto have "cbox a b \ d" proof - have "u = a" "v = b" unfolding euclidean_eq_iff[where 'a='b] proof safe fix j :: 'b assume j: "j \ Basis" note i(2)[unfolded uv mem_box,rule_format,of j] then show "u \ j = a \ j" and "v \ j = b \ j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps) qed then have "i = cbox a b" using uv by auto then show ?thesis using i by auto qed then have deq: "d = insert (cbox a b) (d - {cbox a b})" by auto have "F g (d - {cbox a b}) = \<^bold>1" proof (intro neutral ballI) fix x assume x: "x \ d - {cbox a b}" then have "x\d" by auto note d'[rule_format,OF this] then obtain u v where uv: "x = cbox u v" "\j\Basis. u \ j = a \ j \ v \ j = a \ j \ u \ j = b \ j \ v \ j = b \ j \ u \ j = a \ j \ v \ j = b \ j" by blast have "u \ a \ v \ b" using x[unfolded uv] by auto then obtain j where "u\j \ a\j \ v\j \ b\j" and j: "j \ Basis" unfolding euclidean_eq_iff[where 'a='b] by auto then have "u\j = v\j" using uv(2)[rule_format,OF j] by auto then have "box u v = {}" using j unfolding box_eq_empty by (auto intro!: bexI[of _ j]) then show "g x = \<^bold>1" unfolding uv(1) by (rule box_empty_imp) qed then show "F g d = g (cbox a b)" using division_ofD[OF less.prems] apply (subst deq) apply (subst insert) apply auto done next case False then have "\x. x \ division_points (cbox a b) d" by auto then obtain k c where "k \ Basis" "interval_lowerbound (cbox a b) \ k < c" "c < interval_upperbound (cbox a b) \ k" "\i\d. interval_lowerbound i \ k = c \ interval_upperbound i \ k = c" unfolding division_points_def by auto then obtain j where "a \ k < c" "c < b \ k" and "j \ d" and j: "interval_lowerbound j \ k = c \ interval_upperbound j \ k = c" by (metis division_of_trivial empty_iff interval_bounds' less.prems) let ?lec = "{x. x\k \ c}" let ?gec = "{x. x\k \ c}" define d1 where "d1 = {l \ ?lec | l. l \ d \ l \ ?lec \ {}}" define d2 where "d2 = {l \ ?gec | l. l \ d \ l \ ?gec \ {}}" define cb where "cb = (\i\Basis. (if i = k then c else b\i) *\<^sub>R i)" define ca where "ca = (\i\Basis. (if i = k then c else a\i) *\<^sub>R i)" have "division_points (cbox a b \ ?lec) {l \ ?lec |l. l \ d \ l \ ?lec \ {}} \ division_points (cbox a b) d" by (rule division_points_psubset[OF \d division_of cbox a b\ ab \a \ k < c\ \c < b \ k\ \j \ d\ j \k \ Basis\]) with division_points_finite[OF \d division_of cbox a b\] have "card (division_points (cbox a b \ ?lec) {l \ ?lec |l. l \ d \ l \ ?lec \ {}}) < card (division_points (cbox a b) d)" by (rule psubset_card_mono) moreover have "division_points (cbox a b \ {x. c \ x \ k}) {l \ {x. c \ x \ k} |l. l \ d \ l \ {x. c \ x \ k} \ {}} \ division_points (cbox a b) d" by (rule division_points_psubset[OF \d division_of cbox a b\ ab \a \ k < c\ \c < b \ k\ \j \ d\ j \k \ Basis\]) with division_points_finite[OF \d division_of cbox a b\] have "card (division_points (cbox a b \ ?gec) {l \ ?gec |l. l \ d \ l \ ?gec \ {}}) < card (division_points (cbox a b) d)" by (rule psubset_card_mono) ultimately have *: "F g d1 = g (cbox a b \ ?lec)" "F g d2 = g (cbox a b \ ?gec)" unfolding interval_split[OF \k \ Basis\] apply (rule_tac[!] less.hyps) using division_split[OF \d division_of cbox a b\, where k=k and c=c] \k \ Basis\ by (simp_all add: interval_split d1_def d2_def division_points_finite[OF \d division_of cbox a b\]) have fxk_le: "g (l \ ?lec) = \<^bold>1" if "l \ d" "y \ d" "l \ ?lec = y \ ?lec" "l \ y" for l y proof - obtain u v where leq: "l = cbox u v" using \l \ d\ less.prems by auto have "interior (cbox u v \ ?lec) = {}" using that division_split_left_inj leq less.prems by blast then show ?thesis unfolding leq interval_split [OF \k \ Basis\] by (auto intro: box_empty_imp) qed have fxk_ge: "g (l \ {x. x \ k \ c}) = \<^bold>1" if "l \ d" "y \ d" "l \ ?gec = y \ ?gec" "l \ y" for l y proof - obtain u v where leq: "l = cbox u v" using \l \ d\ less.prems by auto have "interior (cbox u v \ ?gec) = {}" using that division_split_right_inj leq less.prems by blast then show ?thesis unfolding leq interval_split[OF \k \ Basis\] by (auto intro: box_empty_imp) qed have d1_alt: "d1 = (\l. l \ ?lec) ` {l \ d. l \ ?lec \ {}}" using d1_def by auto have d2_alt: "d2 = (\l. l \ ?gec) ` {l \ d. l \ ?gec \ {}}" using d2_def by auto have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev") unfolding * using \k \ Basis\ by (auto dest: Basis_imp) also have "F g d1 = F (\l. g (l \ ?lec)) d" unfolding d1_alt using division_of_finite[OF less.prems] fxk_le by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left) also have "F g d2 = F (\l. g (l \ ?gec)) d" unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left) also have *: "\x\d. g x = g (x \ ?lec) \<^bold>* g (x \ ?gec)" unfolding forall_in_division[OF \d division_of cbox a b\] using \k \ Basis\ by (auto dest: Basis_imp) have "F (\l. g (l \ ?lec)) d \<^bold>* F (\l. g (l \ ?gec)) d = F g d" using * by (simp add: distrib) finally show ?thesis by auto qed qed qed qed proposition tagged_division: assumes "d tagged_division_of (cbox a b)" shows "F (\(_, l). g l) d = g (cbox a b)" proof - have "F (\(_, k). g k) d = F g (snd ` d)" using assms box_empty_imp by (rule over_tagged_division_lemma) then show ?thesis unfolding assms [THEN division_of_tagged_division, THEN division] . qed end locale operative_real = comm_monoid_set + fixes g :: "real set \ 'a" assumes neutral: "b \ a \ g {a..b} = \<^bold>1" assumes coalesce_less: "a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" begin sublocale operative where g = g rewrites "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" and "\x::real. x \ Basis \ x = 1" proof - show "operative f z g" proof show "g (cbox a b) = \<^bold>1" if "box a b = {}" for a b using that by (simp add: neutral) show "g (cbox a b) = g (cbox a b \ {x. x \ k \ c}) \<^bold>* g (cbox a b \ {x. c \ x \ k})" if "k \ Basis" for a b c k proof - from that have [simp]: "k = 1" by simp from neutral [of 0 1] neutral [of a a for a] coalesce_less have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" by auto have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" by (auto simp: min_def max_def le_less) then show "g (cbox a b) = g (cbox a b \ {x. x \ k \ c}) \<^bold>* g (cbox a b \ {x. c \ x \ k})" by (simp add: atMost_def [symmetric] atLeast_def [symmetric]) qed qed show "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" and "\x::real. x \ Basis \ x = 1" by (simp_all add: fun_eq_iff) qed lemma coalesce_less_eq: "g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \ c" "c \ b" proof (cases "c = a \ c = b") case False with that have "a < c" "c < b" by auto then show ?thesis by (rule coalesce_less) next case True with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis by safe simp_all qed end lemma operative_realI: "operative_real f z g" if "operative f z g" proof - interpret operative f z g using that . show ?thesis proof show "g {a..b} = z" if "b \ a" for a b using that box_empty_imp by simp show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c using that using Basis_imp [of 1 a b c] by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def) qed qed subsection \Special case of additivity we need for the FTC\ (*fix add explanation here *) lemma additive_tagged_division_1: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and "p tagged_division_of {a..b}" shows "sum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" proof - let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" interpret operative_real plus 0 ?f rewrites "comm_monoid_set.F (+) 0 = sum" by standard[1] (auto simp add: sum_def) have p_td: "p tagged_division_of cbox a b" using assms(2) box_real(2) by presburger have **: "cbox a b \ {}" using assms(1) by auto then have "f b - f a = (\(x, l)\p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))" proof - have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a" using assms by auto then show ?thesis using p_td assms by (simp add: tagged_division) qed then show ?thesis using assms by (auto intro!: sum.cong) qed subsection \Fine-ness of a partition w.r.t. a gauge\ definition\<^marker>\tag important\ fine (infixr "fine" 46) where "d fine s \ (\(x,k) \ s. k \ d x)" lemma fineI: assumes "\x k. (x, k) \ s \ k \ d x" shows "d fine s" using assms unfolding fine_def by auto lemma fineD[dest]: assumes "d fine s" shows "\x k. (x,k) \ s \ k \ d x" using assms unfolding fine_def by auto lemma fine_Int: "(\x. d1 x \ d2 x) fine p \ d1 fine p \ d2 fine p" unfolding fine_def by auto lemma fine_Inter: "(\x. \{f d x | d. d \ s}) fine p \ (\d\s. (f d) fine p)" unfolding fine_def by blast lemma fine_Un: "d fine p1 \ d fine p2 \ d fine (p1 \ p2)" unfolding fine_def by blast lemma fine_Union: "(\p. p \ ps \ d fine p) \ d fine (\ps)" unfolding fine_def by auto lemma fine_subset: "p \ q \ d fine q \ d fine p" unfolding fine_def by blast subsection \Some basic combining lemmas\ lemma tagged_division_Union_exists: assumes "finite I" and "\i\I. \p. p tagged_division_of i \ d fine p" and "\i1\I. \i2\I. i1 \ i2 \ interior i1 \ interior i2 = {}" and "\I = i" obtains p where "p tagged_division_of i" and "d fine p" proof - obtain pfn where pfn: "\x. x \ I \ pfn x tagged_division_of x" "\x. x \ I \ d fine pfn x" using bchoice[OF assms(2)] by auto show thesis apply (rule_tac p="\(pfn ` I)" in that) using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force by (metis (mono_tags, lifting) fine_Union imageE pfn(2)) qed (* FIXME structure here, do not have one lemma in each subsection *) subsection\<^marker>\tag unimportant\ \The set we're concerned with must be closed\ lemma division_of_closed: fixes i :: "'n::euclidean_space set" shows "s division_of i \ closed i" unfolding division_of_def by fastforce (* FIXME structure here, do not have one lemma in each subsection *) subsection \General bisection principle for intervals; might be useful elsewhere\ (* FIXME move? *) lemma interval_bisection_step: fixes type :: "'a::euclidean_space" assumes emp: "P {}" and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" and non: "\ P (cbox a (b::'a))" obtains c d where "\ P (cbox c d)" and "\i. i \ Basis \ a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" proof - have "cbox a b \ {}" using emp non by metis then have ab: "\i. i\Basis \ a \ i \ b \ i" by (force simp: mem_box) have UN_cases: "\finite \; \S. S\\ \ P S; \S. S\\ \ \a b. S = cbox a b; \S T. S\\ \ T\\ \ S \ T \ interior S \ interior T = {}\ \ P (\\)" for \ proof (induct \ rule: finite_induct) case empty show ?case using emp by auto next case (insert x f) then show ?case unfolding Union_insert by (metis Int_interior_Union_intervals Un insert_iff open_interior) qed let ?ab = "\i. (a\i + b\i) / 2" let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = ?ab i) \ (c\i = ?ab i) \ (d\i = b\i)}" have "P (\?A)" if "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i \ P (cbox c d)" proof (rule UN_cases) let ?B = "(\S. cbox (\i\Basis. (if i \ S then a\i else ?ab i) *\<^sub>R i::'a) (\i\Basis. (if i \ S then ?ab i else b\i) *\<^sub>R i)) ` {s. s \ Basis}" have "?A \ ?B" proof fix x assume "x \ ?A" then obtain c d where x: "x = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast have "c = (\i\Basis. (if c \ i = a \ i then a \ i else ?ab i) *\<^sub>R i)" "d = (\i\Basis. (if c \ i = a \ i then ?ab i else b \ i) *\<^sub>R i)" using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+ then show "x \ ?B" unfolding x by (rule_tac x="{i. i\Basis \ c\i = a\i}" in image_eqI) auto qed then show "finite ?A" by (rule finite_subset) auto next fix S assume "S \ ?A" then obtain c d where s: "S = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast show "P S" unfolding s using ab s(2) by (fastforce intro!: that) show "\a b. S = cbox a b" unfolding s by auto fix T assume "T \ ?A" then obtain e f where t: "T = cbox e f" "\i. i \ Basis \ e \ i = a \ i \ f \ i = ?ab i \ e \ i = ?ab i \ f \ i = b \ i" by blast assume "S \ T" then have "\ (c = e \ d = f)" unfolding s t by auto then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" unfolding euclidean_eq_iff[where 'a='a] by auto then have i: "c\i \ e\i" "d\i \ f\i" using s(2) t(2) apply fastforce using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto show "interior S \ interior T = {}" unfolding s t interior_cbox proof (rule *) fix x assume "x \ box c d" "x \ box e f" then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" unfolding mem_box using i' by force+ show False using s(2)[OF i'] t(2)[OF i'] and i x by auto qed qed also have "\?A = cbox a b" proof (rule set_eqI,rule) fix x assume "x \ \?A" then obtain c d where x: "x \ cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast then show "x\cbox a b" unfolding mem_box by force next fix x assume x: "x \ cbox a b" then have "\i\Basis. \c d. (c = a\i \ d = ?ab i \ c = ?ab i \ d = b\i) \ c\x\i \ x\i \ d" (is "\i\Basis. \c d. ?P i c d") unfolding mem_box by (metis linear) then obtain \ \ where "\i\Basis. (\ \ i = a \ i \ \ \ i = ?ab i \ \ \ i = ?ab i \ \ \ i = b \ i) \ \ \ i \ x \ i \ x \ i \ \ \ i" by (auto simp: choice_Basis_iff) then show "x\\?A" by (force simp add: mem_box) qed finally show thesis by (metis (no_types, lifting) assms(3) that) qed lemma interval_bisection: fixes type :: "'a::euclidean_space" assumes "P {}" and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" and "\ P (cbox a (b::'a))" obtains x where "x \ cbox a b" and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" proof - have "\x. \y. \ P (cbox (fst x) (snd x)) \ (\ P (cbox (fst y) (snd y)) \ (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" (is "\x. ?P x") proof show "?P x" for x proof (cases "P (cbox (fst x) (snd x))") case True then show ?thesis by auto next case False obtain c d where "\ P (cbox c d)" "\i. i \ Basis \ fst x \ i \ c \ i \ c \ i \ d \ i \ d \ i \ snd x \ i \ 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" by (blast intro: interval_bisection_step[of P, OF assms(1-2) False]) then show ?thesis by (rule_tac x="(c,d)" in exI) auto qed qed then obtain f where f: "\x. \ P (cbox (fst x) (snd x)) \ \ P (cbox (fst (f x)) (snd (f x))) \ (\i\Basis. fst x \ i \ fst (f x) \ i \ fst (f x) \ i \ snd (f x) \ i \ snd (f x) \ i \ snd x \ i \ 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" by metis define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") proof - show "A 0 = a" "B 0 = b" unfolding ab_def by auto note S = ab_def funpow.simps o_def id_apply show "?P n" for n proof (induct n) case 0 then show ?case unfolding S using \\ P (cbox a b)\ f by auto next case (Suc n) show ?case unfolding S apply (rule f[rule_format]) using Suc unfolding S apply auto done qed qed then have AB: "A(n)\i \ A(Suc n)\i" "A(Suc n)\i \ B(Suc n)\i" "B(Suc n)\i \ B(n)\i" "2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i" if "i\Basis" for i n using that by blast+ have notPAB: "\ P (cbox (A(Suc n)) (B(Suc n)))" for n using ABRAW by blast have interv: "\n. \x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" if e: "0 < e" for e proof - obtain n where n: "(\i\Basis. b \ i - a \ i) / e < 2 ^ n" using real_arch_pow[of 2 "(sum (\i. b\i - a\i) Basis) / e"] by auto show ?thesis proof (rule exI [where x=n], clarify) fix x y assume xy: "x\cbox (A n) (B n)" "y\cbox (A n) (B n)" have "dist x y \ sum (\i. \(x - y)\i\) Basis" unfolding dist_norm by(rule norm_le_l1) also have "\ \ sum (\i. B n\i - A n\i) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" show "\(x - y) \ i\ \ B n \ i - A n \ i" using xy[unfolded mem_box,THEN bspec, OF i] by (auto simp: inner_diff_left) qed also have "\ \ sum (\i. b\i - a\i) Basis / 2^n" unfolding sum_divide_distrib proof (rule sum_mono) show "B n \ i - A n \ i \ (b \ i - a \ i) / 2 ^ n" if i: "i \ Basis" for i proof (induct n) case 0 then show ?case unfolding AB by auto next case (Suc n) have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" using AB(3) that AB(4)[of i n] using i by auto also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" using Suc by (auto simp add: field_simps) finally show ?case . qed qed also have "\ < e" using n using e by (auto simp add: field_simps) finally show "dist x y < e" . qed qed { fix n m :: nat assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" proof (induction rule: inc_induct) case (step i) show ?case using AB by (intro order_trans[OF step.IH] subset_box_imp) auto qed simp } note ABsubset = this have "\n. cbox (A n) (B n) \ {}" by (meson AB dual_order.trans interval_not_empty) then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" using decreasing_closed_nest [OF closed_cbox] ABsubset interv by blast show thesis proof (rule that[rule_format, of x0]) show "x0\cbox a b" using \A 0 = a\ \B 0 = b\ x0 by blast fix e :: real assume "e > 0" from interv[OF this] obtain n where n: "\x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" .. have "\ P (cbox (A n) (B n))" proof (cases "0 < n") case True then show ?thesis by (metis Suc_pred' notPAB) next case False then show ?thesis using \A 0 = a\ \B 0 = b\ \\ P (cbox a b)\ by blast qed moreover have "cbox (A n) (B n) \ ball x0 e" using n using x0[of n] by auto moreover have "cbox (A n) (B n) \ cbox a b" using ABsubset \A 0 = a\ \B 0 = b\ by blast ultimately show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" apply (rule_tac x="A n" in exI) apply (rule_tac x="B n" in exI) apply (auto simp: x0) done qed qed subsection \Cousin's lemma\ lemma fine_division_exists: (*FIXME rename(?) *) fixes a b :: "'a::euclidean_space" assumes "gauge g" obtains p where "p tagged_division_of (cbox a b)" "g fine p" proof (cases "\p. p tagged_division_of (cbox a b) \ g fine p") case True then show ?thesis using that by auto next case False assume "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" obtain x where x: "x \ (cbox a b)" "\e. 0 < e \ \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ (cbox a b) \ \ (\p. p tagged_division_of cbox c d \ g fine p)" apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p", OF _ _ False]) apply (simp add: fine_def) apply (metis tagged_division_Un fine_Un) apply auto done obtain e where e: "e > 0" "ball x e \ g x" using gaugeD[OF assms, of x] unfolding open_contains_ball by auto from x(2)[OF e(1)] obtain c d where c_d: "x \ cbox c d" "cbox c d \ ball x e" "cbox c d \ cbox a b" "\ (\p. p tagged_division_of cbox c d \ g fine p)" by blast have "g fine {(x, cbox c d)}" unfolding fine_def using e using c_d(2) by auto then show ?thesis using tagged_division_of_self[OF c_d(1)] using c_d by auto qed lemma fine_division_exists_real: fixes a b :: real assumes "gauge g" obtains p where "p tagged_division_of {a .. b}" "g fine p" by (metis assms box_real(2) fine_division_exists) subsection \A technical lemma about "refinement" of division\ lemma tagged_division_finer: fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" assumes ptag: "p tagged_division_of (cbox a b)" and "gauge d" obtains q where "q tagged_division_of (cbox a b)" and "d fine q" and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" proof - have p: "finite p" "p tagged_partial_division_of (cbox a b)" - using ptag unfolding tagged_division_of_def by auto + using ptag tagged_division_of_def by blast+ have "(\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" if "finite p" "p tagged_partial_division_of (cbox a b)" "gauge d" for p using that proof (induct p) case empty show ?case by (force simp add: fine_def) next case (insert xk p) obtain x k where xk: "xk = (x, k)" using surj_pair[of xk] by metis obtain q1 where q1: "q1 tagged_division_of \{k. \x. (x, k) \ p}" and "d fine q1" and q1I: "\x k. \(x, k)\p; k \ d x\ \ (x, k) \ q1" using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI] by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2)) have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" unfolding xk by auto note p = tagged_partial_division_ofD[OF insert(4)] obtain u v where uv: "k = cbox u v" using p(4) xk by blast have "finite {k. \x. (x, k) \ p}" apply (rule finite_subset[of _ "snd ` p"]) using image_iff apply fastforce using insert.hyps(1) by blast then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" proof (rule Int_interior_Union_intervals) show "open (interior (cbox u v))" by auto show "\T. T \ {k. \x. (x, k) \ p} \ \a b. T = cbox a b" using p(4) by auto show "\T. T \ {k. \x. (x, k) \ p} \ interior (cbox u v) \ interior T = {}" by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk) qed show ?case proof (cases "cbox u v \ d x") case True have "{(x, cbox u v)} tagged_division_of cbox u v" by (simp add: p(2) uv xk tagged_division_of_self) then have "{(x, cbox u v)} \ q1 tagged_division_of \{k. \x. (x, k) \ insert xk p}" unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un) with True show ?thesis apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) using \d fine q1\ fine_def q1I uv xk apply fastforce done next case False obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2" using fine_division_exists[OF assms(2)] by blast show ?thesis apply (rule_tac x="q2 \ q1" in exI) apply (intro conjI) unfolding * uv apply (rule tagged_division_Un q2 q1 int fine_Un)+ apply (auto intro: q1 q2 fine_Un \d fine q1\ simp add: False q1I uv xk) done qed qed with p obtain q where q: "q tagged_division_of \{k. \x. (x, k) \ p}" "d fine q" "\(x, k)\p. k \ d x \ (x, k) \ q" by (meson \gauge d\) with ptag that show ?thesis by auto qed subsubsection\<^marker>\tag important\ \Covering lemma\ text\ Some technical lemmas used in the approximation results that follow. Proof of the covering lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's "Introduction to Gauge Integrals". \ proposition covering_lemma: assumes "S \ cbox a b" "box a b \ {}" "gauge g" obtains \ where "countable \" "\\ \ cbox a b" "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\K. K \ \ \ \x \ S \ K. K \ g x" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" "S \ \\" proof - have aibi: "\i. i \ Basis \ a \ i \ b \ i" and normab: "0 < norm(b - a)" using \box a b \ {}\ box_eq_empty box_sing by fastforce+ let ?K0 = "\(n, f::'a\nat). cbox (\i \ Basis. (a \ i + (f i / 2^n) * (b \ i - a \ i)) *\<^sub>R i) (\i \ Basis. (a \ i + ((f i + 1) / 2^n) * (b \ i - a \ i)) *\<^sub>R i)" let ?D0 = "?K0 ` (SIGMA n:UNIV. Pi\<^sub>E Basis (\i::'a. lessThan (2^n)))" obtain \0 where count: "countable \0" and sub: "\\0 \ cbox a b" and int: "\K. K \ \0 \ (interior K \ {}) \ (\c d. K = cbox c d)" and intdj: "\A B. \A \ \0; B \ \0\ \ A \ B \ B \ A \ interior A \ interior B = {}" and SK: "\x. x \ S \ \K \ \0. x \ K \ K \ g x" and cbox: "\u v. cbox u v \ \0 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" and fin: "\K. K \ \0 \ finite {L \ \0. K \ L}" proof show "countable ?D0" by (simp add: countable_PiE) next show "\?D0 \ cbox a b" apply (simp add: UN_subset_iff) apply (intro conjI allI ballI subset_box_imp) apply (simp add: field_simps) apply (auto intro: mult_right_mono aibi) apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem intro: mult_right_mono) done next show "\K. K \ ?D0 \ interior K \ {} \ (\c d. K = cbox c d)" using \box a b \ {}\ by (clarsimp simp: box_eq_empty) (fastforce simp add: field_split_simps dest: PiE_mem) next have realff: "(real w) * 2^m < (real v) * 2^n \ w * 2^m < v * 2^n" for m n v w using of_nat_less_iff less_imp_of_nat_less by fastforce have *: "\v w. ?K0(m,v) \ ?K0(n,w) \ ?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" for m n \ \The symmetry argument requires a single HOL formula\ proof (rule linorder_wlog [where a=m and b=n], intro allI impI) fix v w m and n::nat assume "m \ n" \ \WLOG we can assume \<^term>\m \ n\, when the first disjunct becomes impossible\ have "?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" apply (simp add: subset_box disjoint_interval) apply (rule ccontr) apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le) apply (drule_tac x=i in bspec, assumption) using \m\n\ realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"] apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff) apply (simp_all add: power_add) apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) done then show "?K0(m,v) \ ?K0(n,w) \ ?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" by meson qed auto show "\A B. \A \ ?D0; B \ ?D0\ \ A \ B \ B \ A \ interior A \ interior B = {}" apply (erule imageE SigmaE)+ using * by simp next show "\K \ ?D0. x \ K \ K \ g x" if "x \ S" for x proof (simp only: bex_simps split_paired_Bex_Sigma) show "\n. \f \ Basis \\<^sub>E {..<2 ^ n}. x \ ?K0(n,f) \ ?K0(n,f) \ g x" proof - obtain e where "0 < e" and e: "\y. (\i. i \ Basis \ \x \ i - y \ i\ \ e) \ y \ g x" proof - have "x \ g x" "open (g x)" using \gauge g\ by (auto simp: gauge_def) then obtain \ where "0 < \" and \: "ball x \ \ g x" using openE by blast have "norm (x - y) < \" if "(\i. i \ Basis \ \x \ i - y \ i\ \ \ / (2 * real DIM('a)))" for y proof - have "norm (x - y) \ (\i\Basis. \x \ i - y \ i\)" by (metis (no_types, lifting) inner_diff_left norm_le_l1 sum.cong) also have "... \ DIM('a) * (\ / (2 * real DIM('a)))" by (meson sum_bounded_above that) also have "... = \ / 2" by (simp add: field_split_simps) also have "... < \" by (simp add: \0 < \\) finally show ?thesis . qed then show ?thesis by (rule_tac e = "\ / 2 / DIM('a)" in that) (simp_all add: \0 < \\ dist_norm subsetD [OF \]) qed have xab: "x \ cbox a b" using \x \ S\ \S \ cbox a b\ by blast obtain n where n: "norm (b - a) / 2^n < e" using real_arch_pow_inv [of "e / norm(b - a)" "1/2"] normab \0 < e\ by (auto simp: field_split_simps) then have "norm (b - a) < e * 2^n" by (auto simp: field_split_simps) then have bai: "b \ i - a \ i < e * 2 ^ n" if "i \ Basis" for i proof - have "b \ i - a \ i \ norm (b - a)" by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that) also have "... < e * 2 ^ n" using \norm (b - a) < e * 2 ^ n\ by blast finally show ?thesis . qed have D: "(a + n \ x \ x \ a + m) \ (a + n \ y \ y \ a + m) \ abs(x - y) \ m - n" for a m n x and y::real by auto have "\i\Basis. \k<2 ^ n. (a \ i + real k * (b \ i - a \ i) / 2 ^ n \ x \ i \ x \ i \ a \ i + (real k + 1) * (b \ i - a \ i) / 2 ^ n)" proof fix i::'a assume "i \ Basis" consider "x \ i = b \ i" | "x \ i < b \ i" using \i \ Basis\ mem_box(2) xab by force then show "\k<2 ^ n. (a \ i + real k * (b \ i - a \ i) / 2 ^ n \ x \ i \ x \ i \ a \ i + (real k + 1) * (b \ i - a \ i) / 2 ^ n)" proof cases case 1 then show ?thesis by (rule_tac x = "2^n - 1" in exI) (auto simp: algebra_simps field_split_simps of_nat_diff \i \ Basis\ aibi) next case 2 then have abi_less: "a \ i < b \ i" using \i \ Basis\ xab by (auto simp: mem_box) let ?k = "nat \2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)\" show ?thesis proof (intro exI conjI) show "?k < 2 ^ n" using aibi xab \i \ Basis\ by (force simp: nat_less_iff floor_less_iff field_split_simps 2 mem_box) next have "a \ i + real ?k * (b \ i - a \ i) / 2 ^ n \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) using aibi [OF \i \ Basis\] xab 2 apply (simp_all add: \i \ Basis\ mem_box field_split_simps) done also have "... = x \ i" using abi_less by (simp add: field_split_simps) finally show "a \ i + real ?k * (b \ i - a \ i) / 2 ^ n \ x \ i" . next have "x \ i \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" using abi_less by (simp add: field_split_simps) also have "... \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) using aibi [OF \i \ Basis\] xab apply (auto simp: \i \ Basis\ mem_box divide_simps) done finally show "x \ i \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" . qed qed qed then have "\f\Basis \\<^sub>E {..<2 ^ n}. x \ ?K0(n,f)" apply (simp add: mem_box Bex_def) apply (clarify dest!: bchoice) apply (rule_tac x="restrict f Basis" in exI, simp) done moreover have "\f. x \ ?K0(n,f) \ ?K0(n,f) \ g x" apply (clarsimp simp add: mem_box) apply (rule e) apply (drule bspec D, assumption)+ apply (erule order_trans) apply (simp add: divide_simps) using bai apply (force simp add: algebra_simps) done ultimately show ?thesis by auto qed qed next show "\u v. cbox u v \ ?D0 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" by (force simp: eq_cbox box_eq_empty field_simps dest!: aibi) next obtain j::'a where "j \ Basis" using nonempty_Basis by blast have "finite {L \ ?D0. ?K0(n,f) \ L}" if "f \ Basis \\<^sub>E {..<2 ^ n}" for n f proof (rule finite_subset) let ?B = "(\(n, f::'a\nat). cbox (\i\Basis. (a \ i + (f i) / 2^n * (b \ i - a \ i)) *\<^sub>R i) (\i\Basis. (a \ i + ((f i) + 1) / 2^n * (b \ i - a \ i)) *\<^sub>R i)) ` (SIGMA m:atMost n. Pi\<^sub>E Basis (\i::'a. lessThan (2^m)))" have "?K0(m,g) \ ?B" if "g \ Basis \\<^sub>E {..<2 ^ m}" "?K0(n,f) \ ?K0(m,g)" for m g proof - have dd: "w / m \ v / n \ (v+1) / n \ (w+1) / m \ inverse n \ inverse m" for w m v n::real by (auto simp: field_split_simps) have bjaj: "b \ j - a \ j > 0" using \j \ Basis\ \box a b \ {}\ box_eq_empty(1) by fastforce have "((g j) / 2 ^ m) * (b \ j - a \ j) \ ((f j) / 2 ^ n) * (b \ j - a \ j) \ (((f j) + 1) / 2 ^ n) * (b \ j - a \ j) \ (((g j) + 1) / 2 ^ m) * (b \ j - a \ j)" using that \j \ Basis\ by (simp add: subset_box field_split_simps aibi) then have "((g j) / 2 ^ m) \ ((f j) / 2 ^ n) \ ((real(f j) + 1) / 2 ^ n) \ ((real(g j) + 1) / 2 ^ m)" by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2) then have "inverse (2^n) \ (inverse (2^m) :: real)" by (rule dd) then have "m \ n" by auto show ?thesis by (rule imageI) (simp add: \m \ n\ that) qed then show "{L \ ?D0. ?K0(n,f) \ L} \ ?B" by auto show "finite ?B" by (intro finite_imageI finite_SigmaI finite_atMost finite_lessThan finite_PiE finite_Basis) qed then show "finite {L \ ?D0. K \ L}" if "K \ ?D0" for K using that by auto qed let ?D1 = "{K \ \0. \x \ S \ K. K \ g x}" obtain \ where count: "countable \" and sub: "\\ \ cbox a b" "S \ \\" and int: "\K. K \ \ \ (interior K \ {}) \ (\c d. K = cbox c d)" and intdj: "\A B. \A \ \; B \ \\ \ A \ B \ B \ A \ interior A \ interior B = {}" and SK: "\K. K \ \ \ \x. x \ S \ K \ K \ g x" and cbox: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" and fin: "\K. K \ \ \ finite {L. L \ \ \ K \ L}" proof show "countable ?D1" using count countable_subset by (simp add: count countable_subset) show "\?D1 \ cbox a b" using sub by blast show "S \ \?D1" using SK by (force simp:) show "\K. K \ ?D1 \ (interior K \ {}) \ (\c d. K = cbox c d)" using int by blast show "\A B. \A \ ?D1; B \ ?D1\ \ A \ B \ B \ A \ interior A \ interior B = {}" using intdj by blast show "\K. K \ ?D1 \ \x. x \ S \ K \ K \ g x" by auto show "\u v. cbox u v \ ?D1 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" using cbox by blast show "\K. K \ ?D1 \ finite {L. L \ ?D1 \ K \ L}" using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset) qed let ?\ = "{K \ \. \K'. K' \ \ \ K \ K' \ \(K \ K')}" show ?thesis proof (rule that) show "countable ?\" by (blast intro: countable_subset [OF _ count]) show "\?\ \ cbox a b" using sub by blast show "S \ \?\" proof clarsimp fix x assume "x \ S" then obtain X where "x \ X" "X \ \" using \S \ \\\ by blast let ?R = "{(K,L). K \ \ \ L \ \ \ L \ K}" have irrR: "irrefl ?R" by (force simp: irrefl_def) have traR: "trans ?R" by (force simp: trans_def) have finR: "\x. finite {y. (y, x) \ ?R}" by simp (metis (mono_tags, lifting) fin \X \ \\ finite_subset mem_Collect_eq psubset_imp_subset subsetI) have "{X \ \. x \ X} \ {}" using \X \ \\ \x \ X\ by blast then obtain Y where "Y \ {X \ \. x \ X}" "\Y'. (Y', Y) \ ?R \ Y' \ {X \ \. x \ X}" by (rule wfE_min' [OF wf_finite_segments [OF irrR traR finR]]) blast then show "\Y. Y \ \ \ (\K'. K' \ \ \ Y \ K' \ \ Y \ K') \ x \ Y" by blast qed show "\K. K \ ?\ \ interior K \ {} \ (\c d. K = cbox c d)" by (blast intro: dest: int) show "pairwise (\A B. interior A \ interior B = {}) ?\" using intdj by (simp add: pairwise_def) metis show "\K. K \ ?\ \ \x \ S \ K. K \ g x" using SK by force show "\u v. cbox u v \ ?\ \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" using cbox by force qed qed subsection \Division filter\ text \Divisions over all gauges towards finer divisions.\ definition\<^marker>\tag important\ division_filter :: "'a::euclidean_space set \ ('a \ 'a set) set filter" where "division_filter s = (INF g\{g. gauge g}. principal {p. p tagged_division_of s \ g fine p})" proposition eventually_division_filter: "(\\<^sub>F p in division_filter s. P p) \ (\g. gauge g \ (\p. p tagged_division_of s \ g fine p \ P p))" unfolding division_filter_def proof (subst eventually_INF_base; clarsimp) fix g1 g2 :: "'a \ 'a set" show "gauge g1 \ gauge g2 \ \x. gauge x \ {p. p tagged_division_of s \ x fine p} \ {p. p tagged_division_of s \ g1 fine p} \ {p. p tagged_division_of s \ x fine p} \ {p. p tagged_division_of s \ g2 fine p}" by (intro exI[of _ "\x. g1 x \ g2 x"]) (auto simp: fine_Int) qed (auto simp: eventually_principal) lemma division_filter_not_empty: "division_filter (cbox a b) \ bot" unfolding trivial_limit_def eventually_division_filter by (auto elim: fine_division_exists) lemma eventually_division_filter_tagged_division: "eventually (\p. p tagged_division_of s) (division_filter s)" unfolding eventually_division_filter by (intro exI[of _ "\x. ball x 1"]) auto end