# HG changeset patch # User paulson # Date 1502817290 -3600 # Node ID bc76686f85a3390448112e7ffa65e06a1e46a017 # Parent ad1e8ad9a3aec75a4ba70bff1952ff58bb2f26e1# Parent b868bb15edbe4d60a37579656c82a6329ef0b54a merged diff -r ad1e8ad9a3ae -r bc76686f85a3 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Aug 15 18:15:04 2017 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Aug 15 18:14:50 2017 +0100 @@ -5670,10 +5670,10 @@ assumes "f integrable_on cbox a b" and "e > 0" and "gauge d" - and "(\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 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" + shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?x \ e") proof - { presume "\k. 0 ?x \ e + k" then show ?thesis by (blast intro: field_le_epsilon) } @@ -5720,15 +5720,9 @@ let ?p = "p \ \(qq ` r)" have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e" - apply (rule assms(4)[rule_format]) - proof + proof (rule less_e) show "d fine ?p" - apply (rule fine_Un) - apply (rule p) - apply (rule fine_Union) - using qq - apply auto - done + by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) note * = tagged_partial_division_of_union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" using r @@ -5902,53 +5896,55 @@ lemma henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" - assumes "f integrable_on cbox a b" - and "e > 0" - and "gauge d" - and "\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 tagged_partial_division_of (cbox a b)" + assumes fed: "f integrable_on cbox a b" "e > 0" "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 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" - unfolding split_def - apply (rule sum_norm_allsubsets_bound) - defer - apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)]) - apply safe - apply (rule assms[rule_format,unfolded split_def]) - defer - apply (rule tagged_partial_division_subset) - apply (rule assms) - apply assumption - apply (rule fine_subset) - apply assumption - apply (rule assms) - using assms(5) - apply auto - done +proof - + have "finite p" + using tag by blast + then show ?thesis + unfolding split_def + proof (rule sum_norm_allsubsets_bound) + fix Q + assume Q: "Q \ p" + then have fine: "d fine Q" + by (simp add: \d fine p\ fine_subset) + show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" + apply (rule henstock_lemma_part1[OF fed less_e, unfolded split_def]) + using Q tag tagged_partial_division_subset apply (force simp add: fine)+ + done + qed +qed lemma henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" - assumes "f integrable_on cbox a b" + assumes intf: "f integrable_on cbox a b" and "e > 0" - obtains d where "gauge d" - and "\p. p tagged_partial_division_of (cbox a b) \ d fine p \ - sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" + 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 assms(2) by simp - from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this] - guess d..note d = conjunctD2[OF this] + 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 \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ + norm ((\(x,K)\p. content K *\<^sub>R f x) - integral (cbox a b) f) + < e / (2 * (real DIM('n) + 1))" + by metis show thesis - apply (rule that) - apply (rule d) - proof (safe, goal_cases) - case (1 p) - note * = henstock_lemma_part2[OF assms(1) * d this] - show ?case - apply (rule le_less_trans[OF *]) - using \e > 0\ - apply (auto simp add: field_simps) - done + 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 @@ -6118,30 +6114,19 @@ done fix t assume "t \ {0..s}" - show "norm (\(x, k)\{xk \ p. m (fst xk) = t}. content k *\<^sub>R f (m x) x - - integral k (f (m x))) \ e/2 ^ (t + 2)" - apply (rule order_trans - [of _ "norm (sum (\(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \ p. m (fst xk) = t})"]) - apply (rule eq_refl) - apply (rule arg_cong[where f=norm]) - apply (rule sum.cong) - apply (rule refl) - defer - apply (rule henstock_lemma_part1) - apply (rule assms(1)[rule_format]) - apply (simp add: e) - apply safe - apply (rule c)+ - apply assumption+ - apply (rule tagged_partial_division_subset[of p]) - apply (rule p(1)[unfolded tagged_division_of_def,THEN conjunct1]) - defer - unfolding fine_def - apply safe - apply (drule p(2)[unfolded fine_def,rule_format]) - unfolding d_def - apply auto - done + have "norm (\(x,k)\{xk \ p. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = + norm (\(x,k)\{xk \ p. 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 _ c]) + show "{xk \ p. m (fst xk) = t} tagged_partial_division_of cbox a b" + apply (rule tagged_partial_division_subset[of p]) + using p by (auto simp: tagged_division_of_def) + show "c t fine {xk \ p. m (fst xk) = t}" + using p(2) by (auto simp: fine_def d_def) + qed (auto simp: e) + finally show "norm (\(x, k)\{xk \ p. m (fst xk) = t}. content k *\<^sub>R f (m x) x - + integral k (f (m x))) \ e/2 ^ (t + 2)" . qed qed (insert s, auto) next