# HG changeset patch # User paulson # Date 1502188621 -7200 # Node ID 92b4f0073eea1e0fb4387b0f7d812b791811dffb # Parent 911f950510e028943c84b0b5cc9f7eedba444ae7 more unknotting diff -r 911f950510e0 -r 92b4f0073eea src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Aug 07 21:43:33 2017 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Aug 08 12:37:01 2017 +0200 @@ -3589,38 +3589,27 @@ apply (auto simp add: algebra_simps) done -lemma fundamental_theorem_of_calculus_interior: +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 - - { - presume *: "a < b \ ?thesis" - show ?thesis - proof (cases "a < b") - case True - then show ?thesis by (rule *) - next - case False - then have "a = b" - using assms(1) by auto - then have *: "cbox a b = {b}" "f b - f a = 0" - by (auto simp add: order_antisym) - show ?thesis - unfolding *(2) - unfolding content_eq_0 - using * \a = b\ - by (auto simp: ex_in_conv) - qed - } - assume ab: "a < b" - let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ +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 + let ?P = "\e. \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})" - { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto } + { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content_real by force } fix e :: real assume e: "e > 0" + then have eba8: "(e * (b - a)) / 8 > 0" + using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" using derf_exp by simp @@ -3639,155 +3628,128 @@ by metis have "bounded (f ` cbox a b)" apply (rule compact_imp_bounded compact_continuous_image)+ - using compact_cbox assms - apply auto - done - from this[unfolded bounded_pos] obtain B + using compact_cbox assms by auto + then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" - by metis - have "\da. 0 < 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)" + 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 "a \ {a .. b}" - using ab by auto - note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] - note * = this[unfolded continuous_within Lim_within,rule_format] - have "(e * (b - a)) / 8 > 0" - using e ab by (auto simp add: field_simps) - from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] - have "\l. 0 < l \ norm(l *\<^sub>R f' a) \ (e * (b - a)) / 8" + 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 - thus ?thesis using ab e by auto + thus ?thesis using ab e that by auto next case False then show ?thesis - apply (rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI) - using ab e - apply (auto simp add: field_simps) + apply (rule_tac l="(e * (b - a)) / 8 / norm (f' a)" in that) + using ab e apply (auto simp add: field_simps) done qed - then obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b - a) / 8" by metis - show ?thesis - apply (rule_tac x="min k l" in exI) - apply safe - unfolding min_less_iff_conj - apply rule - apply (rule l k)+ + 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 - - fix c - assume as: "a \ c" "{a .. c} \ {a .. b}" "{a .. c} \ ball a (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] + 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)" - unfolding norm_scaleR - apply (rule mult_right_mono) - using as' by auto + 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 \0 < e * (b - a) / 8\ by auto + case True then show ?thesis + using eba8 by auto next - case False - show ?thesis - apply (rule k(2)[unfolded dist_norm]) - using as' False - apply (auto simp add: field_simps) - done + 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 as(1)] by auto + 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 - then guess da .. note da=conjunctD2[OF this,rule_format] - - have "\db>0. \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" + + 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 "b \ {a .. b}" - using ab by auto - note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] - note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" - using e ab by (auto simp add: field_simps) - from *[OF this] obtain k - where k: "0 < k" - "\x. \x \ {a..b}; 0 < dist x b \ dist x b < k\ - \ dist (f x) (f b) < e * (b - a) / 8" - by blast + 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 + case True thus ?thesis + using ab e that by auto next - case False - then show ?thesis + case False then show ?thesis apply (rule_tac l="(e * (b - a)) / 8 / norm (f' b)" in that) - using ab e - apply (auto simp add: field_simps) - done + using ab e by (auto simp add: field_simps) qed - show ?thesis - apply (rule_tac x="min k l" in exI) - apply safe - unfolding min_less_iff_conj - apply rule - apply (rule l k)+ + 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 - - fix c - assume as: "c \ b" "{c..b} \ {a..b}" "{c..b} \ ball b (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] + 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 "\c - b\ \ \l\" - using as' by auto - then show "norm ((b - c) *\<^sub>R f' b) \ e * (b - a) / 8" - apply - - apply (rule order_trans[OF _ l(2)]) - unfolding norm_scaleR - apply (rule mult_right_mono) - apply auto - done + 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 - show "norm (f b - f c) \ e * (b - a) / 8" - apply (rule less_imp_le) - apply (cases "b = c") - defer - apply (subst norm_minus_commute) - apply (rule k(2)[unfolded dist_norm]) - using as' e ab - apply (auto simp add: field_simps) - done + have "norm (f b - f c) < e * (b - a) / 8" + proof (cases "b = c") + case True + then show ?thesis + using eba8 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 as(1)] by auto + 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 - then guess db .. note db=conjunctD2[OF this,rule_format] let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "?P e" - apply (rule_tac x="?d" in exI) - proof (safe, goal_cases) - case 1 - show ?case - apply (rule gauge_ball_dependent) - using ab db(1) da(1) d(1) - apply auto - done + proof (intro exI conjI allI impI) + show "gauge ?d" + using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next - case as: (2 p) + fix p + assume as: "p tagged_division_of {a..b}" "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF as(1)] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" @@ -3795,7 +3757,7 @@ note * = additive_tagged_division_1[OF assms(1) as(1), symmetric] have **: "\n1 s1 n2 s2::real. n2 \ s2 / 2 \ n1 - s1 \ s2 / 2 \ n1 + n2 \ s1 + s2" by arith - show ?case + show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" unfolding content_real[OF assms(1)] and *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus unfolding sum_distrib_left apply (subst(2) pA) @@ -4058,7 +4020,7 @@ ultimately show ?case unfolding v interval_bounds_real[OF v(2)] box_real apply - - apply(rule da(2)[of "v"]) + apply(rule da[of "v"]) using prems fineD[OF as(2) prems(1)] unfolding v content_eq_0 apply auto @@ -4091,7 +4053,7 @@ unfolding v unfolding interval_bounds_real[OF v(2)] box_real apply - - apply(rule db(2)[of "v"]) + apply(rule db[of "v"]) using prems fineD[OF as(2) prems(1)] unfolding v content_eq_0 apply auto @@ -4102,7 +4064,7 @@ qed show ?case apply (rule * [OF sum_nonneg]) - using ge0 apply (force simp add: ) + using ge0 apply force unfolding sum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric] unfolding sum_distrib_left[symmetric] apply (subst additive_tagged_division_1[OF _ as(1)]) @@ -6714,10 +6676,10 @@ unfolding sub apply - apply rule - defer + apply simp apply (subst(asm) integral_diff) using assms(1) - apply auto + apply auto apply (rule LIMSEQ_imp_Suc) apply assumption done