# HG changeset patch # User paulson # Date 1501963961 -7200 # Node ID 455ca98d9de3e7391d68c6ce3cb69242deb2c2f9 # Parent ff60679dc21df4bd05d179e2397563140f7eae93 final tidying up of lemma bounded_variation_absolutely_integrable_interval diff -r ff60679dc21d -r 455ca98d9de3 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Sat Aug 05 18:16:35 2017 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Sat Aug 05 22:12:41 2017 +0200 @@ -1599,23 +1599,22 @@ apply (rule norm_triangle_ineq3) done -text\FIXME: needs refactoring and use of Sigma\ -lemma bounded_variation_absolutely_integrable_interval: +proposition bounded_variation_absolutely_integrable_interval: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes f: "f integrable_on cbox a b" - and *: "\d. d division_of (cbox a b) \ sum (\k. norm(integral k f)) d \ B" + and *: "\d. d division_of (cbox a b) \ sum (\K. norm(integral K f)) d \ B" shows "f absolutely_integrable_on cbox a b" proof - - let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}" + let ?f = "\d. \K\d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x:?D. ?f x" - have *: "\d. gauge d \ + have *: "\\. gauge \ \ (\p. p tagged_division_of cbox a b \ - d fine p \ + \ fine p \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e)" if e: "e > 0" for e proof - @@ -1645,11 +1644,11 @@ have "e/2 > 0" using e by auto with henstock_lemma[OF f] - obtain g where g: "gauge g" - "\p. \p tagged_partial_division_of cbox a b; g fine p\ + 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" by (metis (no_types, lifting)) - let ?g = "\x. g x \ ball x (k x)" + let ?g = "\x. \ x \ ball x (k x)" show ?thesis proof (intro exI conjI allI impI) show "gauge ?g" @@ -1657,11 +1656,11 @@ next fix p assume "p tagged_division_of (cbox a b) \ ?g fine p" - then have p: "p tagged_division_of cbox a b" "g fine p" "(\x. ball x (k x)) fine p" + then have p: "p tagged_division_of cbox a b" "\ fine p" "(\x. ball x (k x)) fine p" by (auto simp: fine_Int) note p' = tagged_division_ofD[OF p(1)] define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" - have gp': "g fine p'" + have gp': "\ fine p'" using p(2) by (auto simp: p'_def fine_def) have p'': "p' tagged_division_of (cbox a b)" proof (rule tagged_division_ofI) @@ -1673,36 +1672,36 @@ by (simp add: d'(1) p'(1)) qed next - fix x k - assume "(x, k) \ p'" - then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" + fix x K + assume "(x, K) \ p'" + then have "\i l. x \ i \ i \ d \ (x, l) \ p \ K = i \ l" unfolding p'_def by auto - then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "k = i \ l" by blast - show "x \ k" and "k \ cbox a b" + then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" by blast + show "x \ K" and "K \ cbox a b" using p'(2-3)[OF il(3)] il by auto - show "\a b. k = cbox a b" + show "\a b. K = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) next - fix x1 k1 - assume "(x1, k1) \ p'" - then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ k1 = i \ l" + fix x1 K1 + assume "(x1, K1) \ p'" + then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ K1 = i \ l" unfolding p'_def by auto - then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "k1 = i1 \ l1" by blast - fix x2 k2 - assume "(x2,k2) \ p'" - then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ k2 = i \ l" + then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "K1 = i1 \ l1" by blast + fix x2 K2 + assume "(x2,K2) \ p'" + then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ K2 = i \ l" unfolding p'_def by auto - then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "k2 = i2 \ l2" by blast - assume "(x1, k1) \ (x2, k2)" + then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "K2 = i2 \ l2" by blast + assume "(x1, K1) \ (x2, K2)" then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) - then show "interior k1 \ interior k2 = {}" + then show "interior K1 \ interior K2 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" - unfolding p'_def using d' by auto - have "y \ \{k. \x. (x, k) \ p'}" if y: "y \ cbox a b" for y + unfolding p'_def using d' by blast + have "y \ \{K. \x. (x, K) \ p'}" if y: "y \ cbox a b" for y proof - obtain x l where xl: "(x, l) \ p" "y \ l" using y unfolding p'(6)[symmetric] by auto @@ -1714,7 +1713,7 @@ unfolding p'_def by (rule_tac X="i \ l" in UnionI) (use i xl in auto) qed - show "\{k. \x. (x, k) \ p'} = cbox a b" + show "\{K. \x. (x, K) \ p'} = cbox a b" proof show "\{k. \x. (x, k) \ p'} \ cbox a b" using * by auto @@ -1723,23 +1722,23 @@ proof fix y assume y: "y \ cbox a b" - obtain x l where xl: "(x, l) \ p" "y \ l" + obtain x L where xl: "(x, L) \ p" "y \ L" using y unfolding p'(6)[symmetric] by auto - obtain i where i: "i \ d" "y \ i" + obtain I where i: "I \ d" "y \ I" using y unfolding d'(6)[symmetric] by auto - have "x \ i" + have "x \ I" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show "y \ \{k. \x. (x, k) \ p'}" - apply (rule_tac X="i \ l" in UnionI) + apply (rule_tac X="I \ L" in UnionI) using i xl by (auto simp: p'_def) qed qed qed - then have sum_less_e2: "(\(x,k) \ p'. norm (content k *\<^sub>R f x - integral k f)) < e/2" + then have sum_less_e2: "(\(x,K) \ p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" using g(2) gp' tagged_division_of_def by blast - have p'alt: "p' = {(x,(i \ l)) | x i l. (x,l) \ p \ i \ d \ i \ l \ {}}" + have p'alt: "p' = {(x, I \ L) | x I L. (x,L) \ p \ I \ d \ I \ L \ {}}" proof (safe, goal_cases) case prems: (2 _ _ x i l) have "x \ i" @@ -1756,29 +1755,30 @@ then show ?case using prems(3) by auto next - fix x k - assume "(x, k) \ p'" - then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" + fix x K + assume "(x, K) \ p'" + then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" unfolding p'_def by auto - then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "k = i \ l" by blast - then show "\y i l. (x, k) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" + then show "\y i l. (x, K) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" using p'(2) by fastforce qed - have sum_p': "(\(x,k) \ p'. norm (integral k f)) = (\k\snd ` p'. norm (integral k f))" + have sum_p': "(\(x,K) \ p'. norm (integral K f)) = (\k\snd ` p'. norm (integral k f))" apply (subst sum.over_tagged_division_lemma[OF p'',of "\k. norm (integral k f)"]) apply (auto intro: integral_null simp: content_eq_0_interior) done - note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]] + have snd_p_div: "snd ` p division_of cbox a b" + by (rule division_of_tagged_division[OF p(1)]) + note snd_p = division_ofD[OF snd_p_div] have fin_d_sndp: "finite (d \ snd ` p)" by (simp add: d'(1) snd_p(1)) have *: "\sni sni' sf sf'. \\sf' - sni'\ < e/2; ?S - e/2 < sni; sni' \ ?S; - sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" + sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" by arith show "norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e" unfolding real_norm_def proof (rule *) - show "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" + show "\(\(x,K)\p'. norm (content K *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) show "(\(x,k) \ p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) @@ -1786,49 +1786,44 @@ proof - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` (d \ snd ` p)" by auto - have "(\k\d. norm (integral k f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" - proof (rule sum_mono, goal_cases) - case k: (1 k) - from d'(4)[OF this] obtain u v where uv: "k = cbox u v" by metis + have "(\K\d. norm (integral K f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" + proof (rule sum_mono) + fix K assume k: "K \ d" + from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" - note uvab = d'(2)[OF k[unfolded uv]] + have uvab: "cbox u v \ cbox a b" + using d(1) k uv by blast have "d' division_of cbox u v" - apply (subst d'_def) - apply (rule division_inter_1) - apply (rule division_of_tagged_division[OF p(1)]) - apply (rule uvab) + unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) + moreover then have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" + by (simp add: sum_norm_le) + ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" + apply (subst integral_combine_division_topdown[of _ _ d']) + apply (auto simp: uv intro: integrable_on_subcbox[OF assms(1) uvab]) done - then have "norm (integral k f) \ sum (\k. norm (integral k f)) d'" - unfolding uv - apply (subst integral_combine_division_topdown[of _ _ d']) - apply (rule integrable_on_subcbox[OF assms(1) uvab]) - apply assumption - apply (rule sum_norm_le) - apply auto - done - also have "\ = (\k\{k \ l |l. l \ snd ` p}. norm (integral k f))" + also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof - - have *: "norm (integral i f) = 0" - if "i \ {cbox u v \ l |l. l \ snd ` p}" - "i \ {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" for i + have *: "norm (integral I f) = 0" + if "I \ {cbox u v \ l |l. l \ snd ` p}" + "I \ {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" for I using that by auto show ?thesis apply (rule sum.mono_neutral_left) apply (simp add: snd_p(1)) unfolding d'_def uv using * by auto qed - also have "\ = (\l\snd ` p. norm (integral (k \ l) f))" + also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" proof - - have *: "norm (integral (k \ l) f) = 0" - if "l \ snd ` p" "y \ snd ` p" "l \ y" "k \ l = k \ y" for l y + have *: "norm (integral (K \ l) f) = 0" + if "l \ snd ` p" "y \ snd ` p" "l \ y" "K \ l = K \ y" for l y proof - - have "interior (k \ l) \ interior (l \ y)" + have "interior (K \ l) \ interior (l \ y)" by (metis Int_lower2 interior_mono le_inf_iff that(4)) - then have "interior (k \ l) = {}" + then have "interior (K \ l) = {}" by (simp add: snd_p(5) that) moreover from d'(4)[OF k] snd_p(4)[OF that(1)] obtain u1 v1 u2 v2 - where uv: "k = cbox u1 u2" "l = cbox v1 v2" by metis + where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis ultimately show ?thesis using that integral_null unfolding uv Int_interval content_eq_0_interior @@ -1841,7 +1836,7 @@ apply (rule p') using * by auto qed - finally show ?case . + finally show "norm (integral K f) \ (\l\snd ` p. norm (integral (K \ l) f))" . qed also have "\ = (\(i,l) \ d \ snd ` p. norm (integral (i\l) f))" by (simp add: sum.cartesian_product) @@ -1879,21 +1874,16 @@ if "ia \ snd (a, b) \ snd ` p'" "ia \ d" "(a, b) \ p" for ia a b proof - have "ia \ b = {}" - using that - unfolding p'alt image_iff Bex_def not_ex + using that unfolding p'alt image_iff Bex_def not_ex apply (erule_tac x="(a, ia \ b)" in allE) apply auto done - then show ?thesis - by auto + then show ?thesis by auto qed have 1: "\i l. snd (a, b) = i \ l \ i \ d \ l \ snd ` p" if "(a, b) \ p'" for a b - proof - - have "\n N Na. (a, b) = (n, N \ Na) \ (n, Na) \ p \ N \ d \ N \ Na \ {}" - using that p'alt by blast - then show "\N Na. snd (a, b) = N \ Na \ N \ d \ Na \ snd ` p" - by (metis (no_types) imageI prod.sel(2)) - qed + using that + apply (clarsimp simp: p'_def image_iff) + by (metis (no_types, hide_lams) snd_conv) show ?thesis unfolding sum_p' apply (rule sum.mono_neutral_right) @@ -1917,37 +1907,40 @@ unfolding p'alt apply auto by fastforce also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" - unfolding * - apply (subst sum.reindex_nontrivial [OF fin_pd]) - unfolding split_paired_all - unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject - apply (elim conjE) proof - - fix x1 l1 k1 x2 l2 k2 - assume as: "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" - "x1 = x2" "l1 \ k1 = l2 \ k2" "\ ((x1 = x2 \ l1 = l2) \ k1 = k2)" - from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this - from as have "l1 \ l2 \ k1 \ k2" - by auto - then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" - apply (rule disjE) - using as(1) as(2) p'(5) as(3) as(4) d'(5) by auto - moreover have "interior (l1 \ k1) = interior (l2 \ k2)" - unfolding as .. - ultimately have "interior (l1 \ k1) = {}" - by auto - then show "\content (l1 \ k1)\ * norm (f x1) = 0" - unfolding uv Int_interval - unfolding content_eq_0_interior[symmetric] - by auto - qed safe + have "\content (l1 \ k1)\ * norm (f x1) = 0" + if "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" + "x1 = x2" "l1 \ k1 = l2 \ k2" "x1 \ x2 \ l1 \ l2 \ k1 \ k2" + for x1 l1 k1 x2 l2 k2 + proof - + obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" + by (meson \(x1, l1) \ p\ \k1 \ d\ d(1) division_ofD(4) p'(4)) + have "l1 \ l2 \ k1 \ k2" + using that by auto + then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" + apply (rule disjE) + using that p'(5) d'(5) by auto + moreover have "interior (l1 \ k1) = interior (l2 \ k2)" + unfolding that .. + ultimately have "interior (l1 \ k1) = {}" + by auto + then show "\content (l1 \ k1)\ * norm (f x1) = 0" + unfolding uv Int_interval content_eq_0_interior[symmetric] by auto + qed + then show ?thesis + unfolding * + apply (subst sum.reindex_nontrivial [OF fin_pd]) + unfolding split_paired_all o_def split_def prod.inject + apply force+ + done + qed also have "\ = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - have sumeq: "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" if "(x, l) \ p" for x l proof - note xl = p'(2-4)[OF that] - from this(3) guess u v by (elim exE) note uv=this + then obtain u v where uv: "l = cbox u v" by blast have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" by (simp add: Int_commute uv) also have "\ = sum content {k \ cbox u v| k. k \ d}" @@ -2135,10 +2128,7 @@ apply (rule refl) unfolding split_paired_all split_conv apply (drule tagged_division_ofD(4)[OF p(1)]) - unfolding norm_scaleR - apply (subst abs_of_nonneg) - apply auto - done + by simp show "(\(x,k) \ p. norm (integral k f)) \ ?S" using partial_division_of_tagged_division[of p "cbox a b"] p(1) apply (subst sum.over_tagged_division_lemma[OF p(1)]) @@ -2612,11 +2602,12 @@ using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) then have F_le_T: "a \ x \ F x \ T" for x by (intro tendsto_lowerbound[OF lim]) - (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder) + (auto simp: eventually_at_top_linorder) have "(SUP i::nat. ?f i x) = ?fR x" for x proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) - from reals_Archimedean2[of "x - a"] guess n .. + obtain n where "x - a < real n" + using reals_Archimedean2[of "x - a"] .. then have "eventually (\n. ?f n x = ?fR x) sequentially" by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) then show "(\n. ?f n x) \ ?fR x"