# HG changeset patch # User paulson # Date 1503575146 -3600 # Node ID 97fc319d608980bae580bf774b53d4e90f3dbf2a # Parent 8645dc296dca97e549ac361d82626c703761027e# Parent 18a6478a574ca072d9b918d4e1e73d5f6a59fa40 Merge (non-trivial) diff -r 8645dc296dca -r 97fc319d6089 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Aug 24 10:47:56 2017 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Aug 24 12:45:46 2017 +0100 @@ -1643,7 +1643,7 @@ by metis have "e/2 > 0" using e by auto - with henstock_lemma[OF f] + with Henstock_lemma[OF f] obtain \ where g: "gauge \" "\p. \p tagged_partial_division_of cbox a b; \ fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" @@ -2101,7 +2101,7 @@ obtain d2 where "gauge d2" and d2: "\p. \p tagged_partial_division_of (cbox a b); d2 fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" - by (blast intro: henstock_lemma [OF f(1) \e/2>0\]) + by (blast intro: Henstock_lemma [OF f(1) \e/2>0\]) obtain p where p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_Int [OF \gauge d1\ \gauge d2\], of a b]) diff -r 8645dc296dca -r 97fc319d6089 src/HOL/Analysis/Harmonic_Numbers.thy --- a/src/HOL/Analysis/Harmonic_Numbers.thy Thu Aug 24 10:47:56 2017 +0200 +++ b/src/HOL/Analysis/Harmonic_Numbers.thy Thu Aug 24 12:45:46 2017 +0100 @@ -38,7 +38,7 @@ unfolding harm_def by simp lemma abs_harm [simp]: "(abs (harm n) :: real) = harm n" - using harm_nonneg[of n] by (rule abs_of_nonneg) + using harm_nonneg[of n] by (rule abs_of_nonneg) lemma norm_harm: "norm (harm n) = harm n" by (subst of_real_harm [symmetric]) (simp add: harm_nonneg) @@ -99,12 +99,12 @@ show "eventually (\n. harm n \ ln (real (Suc n))) at_top" using ln_le_harm by (intro always_eventually allI) (simp_all add: add_ac) show "filterlim (\n. ln (real (Suc n))) at_top sequentially" - by (intro filterlim_compose[OF ln_at_top] filterlim_compose[OF filterlim_real_sequentially] + by (intro filterlim_compose[OF ln_at_top] filterlim_compose[OF filterlim_real_sequentially] filterlim_Suc) qed -subsection \The Euler--Mascheroni constant\ +subsection \The Euler-Mascheroni constant\ text \ The limit of the difference between the partial harmonic sum and the natural logarithm @@ -250,7 +250,7 @@ qed -subsection \Bounds on the Euler--Mascheroni constant\ +subsection \Bounds on the Euler-Mascheroni constant\ (* TODO: Move? *) lemma ln_inverse_approx_le: @@ -295,7 +295,7 @@ finally show "inverse t \ (t - x) * f' + inverse x" . qed hence "integral {x..x+a} inverse \ integral {x..x+a} ?f" using f'_nonpos assms - by (intro integral_le has_integral_integrable[OF int1] has_integral_integrable[OF int2] ineq) + by (blast intro: integral_le has_integral_integrable[OF int1] has_integral_integrable[OF int2] ineq) also have "\ = ?A" using int1 by (rule integral_unique) finally show ?thesis . qed @@ -332,7 +332,7 @@ thus "(t - m) * f' + inverse m \ inverse t" by (simp add: algebra_simps) qed hence "integral {x..y} inverse \ integral {x..y} (\t. (t - m) * f' + inverse m)" - using int1 int2 by (intro integral_le has_integral_integrable) + using int1 int2 by (blast intro: integral_le has_integral_integrable) also have "integral {x..y} (\t. (t - m) * f' + inverse m) = ?A" using integral_unique[OF int1] by simp finally show ?thesis . @@ -359,7 +359,7 @@ from euler_mascheroni_sum_real have "euler_mascheroni = (\k. D k)" by (simp add: sums_iff D_def) also have "\ = (\k. D (k + Suc n)) + (\k\n. D k)" - by (subst suminf_split_initial_segment[OF summable, of "Suc n"], + by (subst suminf_split_initial_segment[OF summable, of "Suc n"], subst lessThan_Suc_atMost) simp finally have sum: "(\k\n. D k) - euler_mascheroni = -(\k. D (k + Suc n))" by simp @@ -524,7 +524,7 @@ text \ - Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015; + Approximation of the Euler-Mascheroni constant. The lower bound is accurate to about 0.0015; the upper bound is accurate to about 0.015. \ lemma euler_mascheroni_gt_19_over_33: "(euler_mascheroni :: real) > 19/33" (is ?th1) diff -r 8645dc296dca -r 97fc319d6089 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Aug 24 10:47:56 2017 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Aug 24 12:45:46 2017 +0100 @@ -253,26 +253,25 @@ lemma has_integral: "(f has_integral y) (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 - y) < e))" + (\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. \d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ - norm (sum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" - unfolding box_real[symmetric] - by (rule has_integral) + (\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 d - where "gauge d" - and "\p. p tagged_division_of (cbox a b) \ d fine p \ - norm ((\(x,k)\p. content k *\<^sub>R f x) - y) < e" + 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: @@ -914,28 +913,28 @@ subsection \Cauchy-type criterion for integrability.\ -lemma integrable_Cauchy: +proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ - (\p1 p2. p1 tagged_division_of (cbox a b) \ \ fine p1 \ - p2 tagged_division_of (cbox a b) \ \ fine p2 \ - norm ((\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)) < e))" + (\\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 \ \ - (\p. p tagged_division_of cbox a b \ \ fine p \ - norm ((\(x,K) \ p. content K *\<^sub>R f x) - y) < e)" + (\\. \ 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 \: "\p. p tagged_division_of cbox a b \ \ fine p \ - norm ((\(x,K)\p. content K *\<^sub>R f x) - y) < e/2" + 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 \\) @@ -947,9 +946,9 @@ by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" - "\m p1 p2. \p1 tagged_division_of cbox a b; - \ m fine p1; p2 tagged_division_of cbox a b; \ m fine p2\ - \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) + "\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 "\n. gauge (\x. \{\ i x |i. i \ {0..n}})" @@ -993,8 +992,8 @@ 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 \ \ - (\p. p tagged_division_of (cbox a b) \ \ fine p \ - norm ((\(x,K) \ p. content K *\<^sub>R f x) - y) < e)" + (\\. \ 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 @@ -1059,15 +1058,15 @@ by auto obtain \1 where \1: "gauge \1" and \1norm: - "\p. \p tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine p\ - \ norm ((\(x,K) \ p. content K *\<^sub>R f x) - i) < e/2" + "\\. \\ 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: - "\p. \p tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine p\ - \ norm ((\(x, k) \ p. content k *\<^sub>R f x) - j) < e/2" + "\\. \\ 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 @@ -1075,8 +1074,8 @@ have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ - (\p. p tagged_division_of cbox a b \ \ fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e)" + (\\. \ 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" @@ -1285,7 +1284,7 @@ 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_union_interval[OF tdiv1 tdiv2] + 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" @@ -4140,7 +4139,7 @@ using as by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" - apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"]) + apply (rule tagged_division_Un_interval_real[of _ _ _ 1 "t"]) using \t \ c\ by (auto simp: * ptag tagged_division_of_self_real) moreover have "d1 fine p \ {(c, {t..c})}" @@ -4762,9 +4761,9 @@ lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" - assumes "(f has_integral i) s" - and "(g has_integral j) s" - and "\x\s. f x \ g x" + assumes "(f has_integral i) S" + and "(g has_integral j) S" + and "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ assms(1-2), of 1] using assms(3) @@ -4772,27 +4771,27 @@ lemma integral_le: fixes f :: "'n::euclidean_space \ real" - assumes "f integrable_on s" - and "g integrable_on s" - and "\x\s. f x \ g x" - shows "integral s f \ integral s g" + 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\s. 0 \ f x" + 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] + 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 integrable_on s" - and "\x\s. 0 \ f x" - shows "0 \ integral s f" + assumes "f integrable_on S" + and "\x. x \ S \ 0 \ f x" + shows "0 \ integral S f" by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)]) @@ -5686,9 +5685,9 @@ subsection \Henstock's lemma\ -lemma henstock_lemma_part1: +lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on cbox a b" + 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\ \ @@ -5711,59 +5710,52 @@ have r: "finite r" using q' unfolding r_def by auto - have "\i\r. \p. p tagged_division_of i \ d fine p \ - norm (sum (\(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)" - apply safe - proof goal_cases - case (1 i) - then have i: "i \ q" - unfolding r_def by auto - from q'(4)[OF this] guess u v by (elim exE) note uv=this + have "\p. p tagged_division_of i \ d fine p \ + norm (sum (\(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)" + if "i\r" for i + proof - have *: "k / (real (card r) + 1) > 0" using k by simp - have "f integrable_on cbox u v" - apply (rule integrable_subinterval[OF assms(1)]) - using q'(2)[OF i] - unfolding uv - apply auto - done + 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]) note integrable_integral[OF this, unfolded has_integral[of f]] from this[rule_format,OF *] guess dd..note dd=conjunctD2[OF this,rule_format] note gauge_Int[OF \gauge d\ dd(1)] from fine_division_exists[OF this,of u v] guess qq . - then show ?case + then show ?thesis apply (rule_tac x=qq in exI) using dd(2)[of qq] unfolding fine_Int uv apply auto done qed - from bchoice[OF this] guess qq..note qq=this[rule_format] + then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ + d fine qq i \ + norm + ((\(x, j) \ qq i. content j *\<^sub>R f x) - + integral i f) + < k / (real (card r) + 1)" + by metis let ?p = "p \ \(qq ` r)" have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - 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 * = tagged_partial_division_of_union_self[OF p(1)] + note * = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" using r - proof (rule tagged_division_union[OF * tagged_division_unions], goal_cases) - case 1 - then show ?case + proof (rule tagged_division_Un[OF * tagged_division_Union]) + show "\i. i \ r \ qq i tagged_division_of i" using qq by auto - next - case 2 - then show ?case - apply rule - apply rule - apply rule - apply(rule q'(5)) - unfolding r_def - apply auto - done - next - case 3 - then show ?case + show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" + by (simp add: q'(5) r_def) + show "interior (UNION p snd) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (UNION p snd))" by blast @@ -5780,9 +5772,7 @@ qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" - unfolding Union_Un_distrib[symmetric] r_def - using q - by auto + using q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed @@ -5915,11 +5905,11 @@ finally show "?x \ e + k" . qed -lemma henstock_lemma_part2: +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: "\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: "\\. \\ 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" @@ -5934,13 +5924,13 @@ 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]) + 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: +lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" @@ -5951,8 +5941,8 @@ 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) + 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 @@ -5961,7 +5951,7 @@ 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 + using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally @@ -5972,14 +5962,12 @@ subsection \Monotone convergence (bounded interval first)\ -(* TODO: is this lemma necessary? *) 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) -(*FIXME: why so much " \ 1"? *) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" @@ -5988,57 +5976,53 @@ 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 + case True then show ?thesis by auto next case False - have fg1: "(f k x) \ 1 \ (g x) \ 1" if x: "x \ cbox a b" for x k + have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - - have "\\<^sub>F xa in sequentially. f k x \ 1 \ f xa x \ 1" - unfolding eventually_sequentially - apply (rule_tac x=k in exI) + have "\\<^sub>F j in sequentially. f k x \ f j x" + apply (rule eventually_sequentiallyI [of k]) using le x apply (force intro: transitive_stepwise_le) done - with Lim_component_ge [OF fg] x - show "f k x \ 1 \ g x \ 1" - using trivial_limit_sequentially by blast + 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 assms(1-2)) + 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) \ 1" + 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\1" - using Lim_component_ge [OF i] trivial_limit_sequentially by blast + 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. (\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ - abs ((\(x, ka)\p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" + 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 p. \p tagged_division_of cbox a b; c x fine p\ \ - abs ((\(u, ka)\p. content ka *\<^sub>R f x u) - integral (cbox a b) (f x)) + 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\1 - (integral (cbox a b) (f k)) \ i\1 - (integral (cbox a b) (f k)) < e/4" + 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 \ 1 - integral (cbox a b) (f k)" - "\k. r \ k \ i \ 1 - integral (cbox a b) (f k) < e/4" + 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)\1 - (f k x)\1 \ (g x)\1 - (f k x)\1 < e/(4 * content(cbox a b))" + 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" @@ -6048,8 +6032,8 @@ by metis then show "\n\r. \k\n. - 0 \ g x \ 1 - f k x \ 1 \ - g x \ 1 - f k x \ 1 + 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] apply (auto simp add: field_simps) @@ -6057,127 +6041,119 @@ 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 \ 1 - f k x \ 1 \ g x \ 1 - f k x \ 1 < e/(4 * content (cbox a b))" + \ 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 \ \ - (\p. p tagged_division_of cbox a b \ - \ fine p \ abs ((\(x,K)\p. content K *\<^sub>R g x) - i) < e)" + (\\. \ 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 p - assume p: "p tagged_division_of (cbox a b)" "d fine p" - note p'=tagged_division_ofD[OF p(1)] - obtain s where s: "\x. x \ p \ m (fst x) \ s" + 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 "abs ((\(x, k)\p. content k *\<^sub>R g x) - i) < e" + show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) - show "\(\(x, k)\p. content k *\<^sub>R g x) - (\(x, K)\p. content K *\<^sub>R f (m x) x)\ - \ e/4" - apply (rule order_trans[of _ "\(x, k)\p. content k * (e/(4 * content (cbox a b)))"]) - unfolding sum_subtractf[symmetric] - apply (rule order_trans) - apply (rule sum_abs) - apply (rule sum_mono) - unfolding split_paired_all split_conv - unfolding split_def sum_distrib_right[symmetric] scaleR_diff_right[symmetric] - unfolding additive_content_tagged_division[OF p(1), unfolded split_def] - proof - + 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) \ p" - then have x: "x \ cbox a b" - using p'(2-3)[OF xk] by auto - with p'(4)[OF xk] obtain u v where "K = cbox u v" by metis - then show "abs (content K *\<^sub>R (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" - unfolding abs_scaleR using m[OF x] - by (metis real_inner_1_right real_scaleR_def abs_of_nonneg inner_real_def less_eq_real_def measure_nonneg mult_left_mono order_refl) - qed (insert False, auto) + 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)\p. content K *\<^sub>R f (m x) x) - (\(x, k)\p. integral k (f (m x)))) - \ norm - (\j = 0..s. - \(x, K)\{xk \ p. m (fst xk) = j}. - content K *\<^sub>R f (m x) x - integral K (f (m x)))" + 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 \ p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) + 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 \ 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))" + 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 \ 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) + proof (rule Henstock_lemma_part1 [OF intf]) + show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" + apply (rule tagged_partial_division_subset[of \]) + using ptag by (auto simp: tagged_division_of_def) + 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 \ p. m (fst xk) = t}. content K *\<^sub>R f (m x) x - + 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 \ p. + 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)\p. content K *\<^sub>R f (m x) x) - (\(x,K)\p. integral K (f (m x)))\ < e/2" + finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next - note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)] - have *: "\sr sx ss ks kr. \kr = sr; ks = ss; - ks \ i; sr \ sx; sx \ ss; 0 \ i\1 - kr\1; i\1 - kr\1 < e/4\ \ \sx - i\ < e/4" - by auto - show "\(\(x, k)\p. integral k (f (m x))) - i\ < e/4" - unfolding real_norm_def - apply (rule *) - apply (rule comb[of r]) - apply (rule comb[of s]) - apply (rule i'[unfolded real_inner_1_right]) - apply (rule_tac[1-2] sum_mono) - unfolding split_paired_all split_conv - apply (rule_tac[1-2] integral_le[OF ]) - proof safe - show "0 \ i\1 - (integral (cbox a b) (f r))\1" "i\1 - (integral (cbox a b) (f r))\1 < e/4" - using r by auto + 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) \ p" - from p'(4)[OF this] guess u v by (elim exE) note uv=this - show "f r integrable_on K" - and "f s integrable_on K" - and "f (m x) integrable_on K" - and "f (m x) integrable_on K" - unfolding uv - apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]]) - using p'(3)[OF xk] - unfolding uv - apply auto - done - fix y - assume "y \ K" - then have "y \ cbox a b" - using p'(3)[OF xk] by auto - then have *: "\m n. n\m \ f m y \ f n y" - using assms(2) by (auto intro: transitive_stepwise_le) - show "f r y \ f (m x) y" "f (m x) y \ f s y" - using s[rule_format,OF xk] r_le_m[of x] p'(2-3)[OF xk] - apply (auto intro!: *) - done + 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 @@ -6201,23 +6177,19 @@ 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)\1 \ (g x)\1" if "x \ S" for x k - apply (rule Lim_component_ge [OF lim [OF that] trivial_limit_sequentially]) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - using le by (force intro: transitive_stepwise_le that) - + have fg: "(f k x) \ (g x)" if "x \ S" for x k + apply (rule tendsto_lowerbound [OF lim [OF that]]) + apply (rule eventually_sequentiallyI [of k]) + using le by (force intro: transitive_stepwise_le that)+ 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))\1 \ i\1" for k + have i': "(integral S (f k)) \ i" for k proof - - have *: "\k. \x\S. \n\k. f k x \ f n x" + have "\k. \x\S. \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) - show ?thesis - apply (rule Lim_component_ge [OF i trivial_limit_sequentially]) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - using * by (simp add: int_f integral_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)" @@ -6277,7 +6249,7 @@ 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)\1 \ (f n x)\1" if "x \ S" "n \ m" for m n + have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n apply (rule transitive_stepwise_le [OF \n \ m\ order_refl]) using dual_order.trans apply blast by (simp add: le \x \ S\) @@ -6416,41 +6388,41 @@ by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" - "\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" + "\\. \ 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 \" - "\p. p tagged_division_of cbox a b \ \ fine p - \ norm ((\(x, k)\p. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" + "\\. \ 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 p - where p: "p tagged_division_of cbox a b" "(\x. \ x \ \ x) fine p" + with fine_division_exists obtain \ + where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis - have "\ fine p" "\ fine p" + 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) \ p" for x K + 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\ p(1) by blast+ + using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" - using \(x, K) \ p\ p(1) by blast + 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)\p. content k *\<^sub>R f x) \ (\(x, k)\p. content k *\<^sub>R g x)" + 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)\p. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" - using \\ fine p\ \ p(1) by simp - show "\(\(x, k)\p. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" - using \\ fine p\ \ p(1) by simp + 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 @@ -7184,13 +7156,13 @@ \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this - { fix k::real and p and u::'a and v w and z::'b and t1 t2 l + { 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: "p tagged_division_of cbox (a,c) (b,d)" - and fine: "(\x. ball x k) fine p" "((t1,t2), l) \ p" + 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)+ diff -r 8645dc296dca -r 97fc319d6089 src/HOL/Analysis/Improper_Integral.thy --- a/src/HOL/Analysis/Improper_Integral.thy Thu Aug 24 10:47:56 2017 +0200 +++ b/src/HOL/Analysis/Improper_Integral.thy Thu Aug 24 12:45:46 2017 +0100 @@ -658,7 +658,7 @@ 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]) + 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 \" @@ -778,7 +778,7 @@ 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))" - apply (rule sum.over_tagged_division_lemma [OF tagged_partial_division_of_union_self [OF S]]) + 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)))" @@ -789,7 +789,7 @@ \ 2 * content (cbox a b)" 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) + by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division) show "UNION S snd \ cbox a b" using S by force show "a \ i \ c \ i" "c \ i \ b \ i" @@ -861,7 +861,7 @@ 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]) + 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" @@ -985,7 +985,7 @@ 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 UNION A snd" - using A_tagged tagged_partial_division_of_union_self by auto + 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 @@ -1014,7 +1014,7 @@ 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 UNION B snd" - using B_tagged tagged_partial_division_of_union_self by auto + 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 diff -r 8645dc296dca -r 97fc319d6089 src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Thu Aug 24 10:47:56 2017 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Aug 24 12:45:46 2017 +0100 @@ -1122,7 +1122,7 @@ unfolding box_real[symmetric] by (rule tagged_division_of_self) -lemma tagged_division_union: +lemma tagged_division_Un: assumes "p1 tagged_division_of s1" and "p2 tagged_division_of s2" and "interior s1 \ interior s2 = {}" @@ -1150,13 +1150,13 @@ 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_unions: +lemma tagged_division_Union: assumes "finite I" - and "\i\I. pfn i tagged_division_of i" - and "\i1\I. \i2\I. i1 \ i2 \ interior(i1) \ interior(i2) = {}" + 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 assms(2)[rule_format]] + 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)" @@ -1175,16 +1175,18 @@ 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) - using assms(3)[rule_format] interior_mono - by blast + using i(1) i'(1) disj interior_mono by blast show "interior k \ interior k' = {}" - apply (cases "i = i'") - using assm(5) i' i(2) xk'(2) apply blast + 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: +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) @@ -1200,7 +1202,7 @@ apply auto done -lemma tagged_division_union_interval: +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})" @@ -1211,14 +1213,14 @@ by auto show ?thesis apply (subst *) - apply (rule tagged_division_union[OF assms(1-2)]) + 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_union_interval_real: +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})" @@ -1226,7 +1228,7 @@ shows "(p1 \ p2) tagged_division_of {a .. b}" using assms unfolding box_real[symmetric] - by (rule tagged_division_union_interval) + by (rule tagged_division_Un_interval) lemma tagged_division_split_left_inj: assumes d: "d tagged_division_of i" @@ -1539,7 +1541,7 @@ done by (simp add: interval_split k interval_doublesplit) qed - + paragraph \Operative\ locale operative = comm_monoid_set + @@ -1571,7 +1573,7 @@ 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 + 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" @@ -1754,7 +1756,7 @@ using assms box_empty_imp by (rule over_tagged_division_lemma) then show ?thesis unfolding assms [THEN division_of_tagged_division, THEN division] . -qed + qed end @@ -1779,14 +1781,14 @@ 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 [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) + 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 qed show "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" @@ -1796,17 +1798,17 @@ 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 + proof (cases "c = a \ c = b") + case False with that have "a < c" "c < b" by auto - then show ?thesis + then show ?thesis by (rule coalesce_less) -next - case True + next + case True with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis by safe simp_all -qed + qed end @@ -1823,7 +1825,7 @@ 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 qed @@ -1901,7 +1903,7 @@ 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_unions apply force + 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 @@ -2252,7 +2254,7 @@ \ (\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_union fine_Un) + apply (metis tagged_division_Un fine_Un) apply auto done obtain e where e: "e > 0" "ball x e \ g x" @@ -2343,7 +2345,7 @@ 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_union) + 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 @@ -2356,7 +2358,7 @@ apply (rule_tac x="q2 \ q1" in exI) apply (intro conjI) unfolding * uv - apply (rule tagged_division_union q2 q1 int fine_Un)+ + 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