# HG changeset patch # User paulson # Date 1502278883 -7200 # Node ID 5db8427fdfd3b362673fd12993a399b71fdb9019 # Parent 962c12353c679549c1135aedfe36f5a71a3d9dab more cleanup of fundamental_theorem_of_calculus_interior diff -r 962c12353c67 -r 5db8427fdfd3 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Wed Aug 09 12:01:16 2017 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Wed Aug 09 13:41:23 2017 +0200 @@ -3815,6 +3815,8 @@ proof - have *: "\s f t e. sum f s = sum f t \ norm (sum f t) \ e \ norm (sum f s) \ e" by auto + + show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b - a) / 2" apply (rule *[where t1="p \ {t. fst t \ {a, b} \ content(snd t) \ 0}"]) @@ -3840,9 +3842,7 @@ {t. t\p \ fst t = a \ content(snd t) \ 0} \ {t. t\p \ fst t = b \ content(snd t) \ 0}" by blast have **: "norm (sum f s) \ e" - if "\x y. x \ s \ y \ s \ x = y" - and "\x. x \ s \ norm (f x) \ e" - and "e > 0" + 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 @@ -3972,73 +3972,43 @@ qed let ?a = a and ?b = b (* a is something else while proofing the next theorem. *) - show "\x. x \ ?B a \ norm ((\(x, k). content k *\<^sub>R f' x - (f (Sup k) - - f (Inf k))) x) \ e * (b - a) / 4" - apply rule - apply rule - unfolding mem_Collect_eq - unfolding split_paired_all fst_conv snd_conv - proof (safe, goal_cases) - case prems: 1 - guess v using pa[OF prems(1)] .. note v = conjunctD2[OF this] - have "?a \ {?a..v}" - using v(2) by auto - then have "v \ ?b" - using p(3)[OF prems(1)] unfolding subset_eq v by auto + have "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) \ e * (b - a) / 4" + 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 as(2) prems(1)] - apply - - apply (subst(asm) if_P) - apply (rule refl) - unfolding subset_eq - apply safe - apply (erule_tac x=" x" in ballE) - apply (auto simp add:subset_eq dist_real_def v) - done - ultimately show ?case - unfolding v interval_bounds_real[OF v(2)] box_real - apply - - apply(rule da[of "v"]) - using prems fineD[OF as(2) prems(1)] - unfolding v content_eq_0 - apply auto - done + 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 - show "\x. x \ ?B b \ norm ((\(x, k). content k *\<^sub>R f' x - - (f (Sup k) - f (Inf k))) x) \ e * (b - a) / 4" - apply rule - apply rule - unfolding mem_Collect_eq - unfolding split_paired_all fst_conv snd_conv - proof (safe, goal_cases) - case prems: 1 - guess v using pb[OF prems(1)] .. note v = conjunctD2[OF this] - have "?b \ {v.. ?b}" - using v(2) by auto - then have "v \ ?a" using p(3)[OF prems(1)] - unfolding subset_eq v by auto + then show "\x. x \ ?B a \ norm ((\(x, k). content k *\<^sub>R f' x - (f (Sup k) - + f (Inf k))) x) \ e * (b - a) / 4" + by auto + + have "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) \ e * (b - a) / 4" + 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 as(2) prems(1)] - apply - - apply (subst(asm) if_P, rule refl) - unfolding subset_eq - apply safe - apply (erule_tac x=" x" in ballE) - using ab - apply (auto simp add:subset_eq v dist_real_def) - done - ultimately show ?case - unfolding v - unfolding interval_bounds_real[OF v(2)] box_real - apply - - apply(rule db[of "v"]) - using prems fineD[OF as(2) prems(1)] - unfolding v content_eq_0 - apply auto - done + 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 + then show "\x. x \ ?B b \ + norm ((\(x, k). content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) x) + \ e * (b - a) / 4" + by auto qed (insert p(1) ab e, auto simp add: field_simps) qed auto + + qed have *: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2) / 2 \ x - s1 \ s2 / 2" by auto