# HG changeset patch # User paulson # Date 1501928305 -7200 # Node ID 1072edd475dcaebf7710ca3ec5943698901e88b4 # Parent 91257fbcabeeaa1b4dbabc945aff0f58828f2fbc trying to disentangle bounded_variation_absolutely_integrable_interval diff -r 91257fbcabee -r 1072edd475dc src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Fri Aug 04 23:07:14 2017 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Sat Aug 05 12:18:25 2017 +0200 @@ -1605,7 +1605,7 @@ finally show "(\k\d. norm (integral k f)) \ integral UNIV (\x. norm (f x))" . qed -lemma helplemma: +lemma absdiff_norm_less: assumes "sum (\x. norm (f x - g x)) s < e" and "finite s" shows "\sum (\x. norm(f x)) s - sum (\x. norm(g x)) s\ < e" @@ -1630,13 +1630,13 @@ by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x:?D. ?f x" - show ?thesis - apply (rule absolutely_integrable_onI [OF f has_integral_integrable]) - apply (subst has_integral[of _ ?S]) - apply safe - proof goal_cases - case e: (1 e) - then have "?S - e / 2 < ?S" by simp + have *: "\d. gauge d \ + (\p. p tagged_division_of cbox a b \ + d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - ?S) < e)" + if e: "e > 0" for e + proof - + have "?S - e / 2 < ?S" using \e > 0\ by simp then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] @@ -1645,17 +1645,19 @@ proof fix x have "\da>0. \xa\\{i \ d. x \ i}. da \ dist x xa" - apply (rule separate_point_closed) - apply (rule closed_Union) - apply (rule finite_subset[OF _ d'(1)]) - using d'(4) - apply auto - done + proof (rule separate_point_closed) + show "closed (\{i \ d. x \ i})" + apply (rule closed_Union) + apply (simp add: d'(1)) + using d'(4) apply auto + done + show "x \ \{i \ d. x \ i}" + by auto + qed then show "\e>0. \i\d. x \ i \ ball x e \ i = {}" by force qed - then obtain k where k: "\x. 0 < k x" - "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" + then obtain k where k: "\x. 0 < k x" "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" by metis have "e/2 > 0" using e by auto @@ -1665,7 +1667,7 @@ \ (\(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)" - show ?case + show ?thesis apply (rule_tac x="?g" in exI) apply safe proof - @@ -1683,8 +1685,7 @@ unfolding p'_def fine_def by auto have p'': "p' tagged_division_of (cbox a b)" - apply (rule tagged_division_ofI) - proof - + proof (rule tagged_division_ofI) show "finite p'" apply (rule finite_subset[of _ "(\(k,(x,l)). (x,k \ l)) ` {(k,xl) | k xl. k \ d \ xl \ p}"]) @@ -1705,10 +1706,7 @@ using p'(2-3)[OF il(3)] il by auto show "\a b. k = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] - apply safe - unfolding Int_interval - apply auto - done + by (meson Int_interval) next fix x1 k1 assume "(x1, k1) \ p'" @@ -1722,47 +1720,51 @@ 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)] - unfolding il1 il2 - by auto + 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 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by auto - show "\{k. \x. (x, k) \ p'} = cbox a b" - apply rule - apply (rule Union_least) - unfolding mem_Collect_eq - apply (erule exE) - apply (drule *[rule_format]) - apply safe + have "y \ \{k. \x. (x, k) \ p'}" if y: "y \ cbox a b" for y proof - - fix y - assume y: "y \ cbox a b" - then have "\x l. (x, l) \ p \ y\l" - unfolding p'(6)[symmetric] by auto - then obtain x l where xl: "(x, l) \ p" "y \ l" by metis - then have "\k. k \ d \ y \ k" + 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" using y unfolding d'(6)[symmetric] by auto - then obtain i where i: "i \ d" "y \ i" by metis 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'}" - unfolding p'_def Union_iff - apply (rule_tac x="i \ l" in bexI) - using i xl - apply auto - done + then show ?thesis + apply (rule_tac X="i \ l" in UnionI) + using i xl by (auto simp: p'_def) + qed + show "\{k. \x. (x, k) \ p'} = cbox a b" + proof + show "\{k. \x. (x, k) \ p'} \ cbox a b" + using * by auto + next + show "cbox a b \ \{k. \x. (x, k) \ p'}" + proof + fix y + assume y: "y \ cbox a b" + 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" + using y unfolding d'(6)[symmetric] by auto + 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) + using i xl by (auto simp: p'_def) + qed qed qed - then have "(\(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 - then have **: "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e / 2" + then have XXX: "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e / 2" unfolding split_def - using p'' - by (force intro!: helplemma) + using p'' by (force intro!: absdiff_norm_less) have p'alt: "p' = {(x,(i \ l)) | x i l. (x,l) \ p \ i \ d \ i \ l \ {}}" proof (safe, goal_cases) @@ -1776,8 +1778,6 @@ apply safe apply (rule_tac x=x in exI) apply (rule_tac x="i \ l" in exI) - apply safe - using prems apply auto done then show ?case @@ -1797,247 +1797,251 @@ done note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]] - have *: "\sni sni' sf sf'. \sf' - sni'\ < e / 2 \ ?S - e / 2 < sni \ sni' \ ?S \ - sni \ sni' \ sf' = sf \ \sf - ?S\ < e" + + have *: "\sni sni' sf sf'. \\sf' - sni'\ < e / 2; ?S - e / 2 < sni; sni' \ ?S; + 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 - apply (rule *[rule_format,OF **]) - apply safe - apply(rule d(2)) - proof goal_cases - case 1 - show ?case + proof (rule *) + show "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e / 2" + by (rule XXX) + show "(\(x, k)\p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) - next - case 2 - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = + show "(\k\d. norm (integral k f)) \ (\(x, k)\p'. norm (integral k f))" + proof - + have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` {(k,l)|k l. k \ d \ l \ 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 - 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 "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) - 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))" + 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 + 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 "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) + 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))" + 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 + 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))" + proof - + 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)" + by (metis Int_lower2 interior_mono le_inf_iff that(4)) + 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 + ultimately show ?thesis + using that integral_null + unfolding uv Int_interval content_eq_0_interior + by (metis (mono_tags, lifting) norm_eq_zero) + qed + show ?thesis + unfolding Setcompr_eq_image + apply (rule sum.reindex_nontrivial [unfolded o_def]) + apply (rule finite_imageI) + apply (rule p') + using * by auto + qed + finally show ?case . + qed + also have "\ = (\(i,l) \ d \ (snd ` p). norm (integral (i\l) f))" + by (simp add: sum.cartesian_product) + also have "\ = (\x\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (case_prod op \ x) f))" + by (force simp: split_def Sigma_def intro!: sum.cong) + also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k 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 - 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))" - 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 eq0: " (integral (l1 \ k1) f) = 0" + if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" + "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" + for l1 l2 k1 k2 j1 j2 proof - - have "interior (k \ l) \ interior (l \ y)" - by (metis Int_lower2 interior_mono le_inf_iff that(4)) - 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 - ultimately show ?thesis - using that integral_null - unfolding uv Int_interval content_eq_0_interior - by (metis (mono_tags, lifting) norm_eq_zero) + obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" + using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast + have "l1 \ l2 \ k1 \ k2" + using that by auto + then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" + by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) + moreover have "interior (l1 \ k1) = interior (l2 \ k2)" + by (simp add: that(1)) + ultimately have "interior(l1 \ k1) = {}" + by auto + then show ?thesis + unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed show ?thesis - unfolding Setcompr_eq_image - apply (rule sum.reindex_nontrivial [unfolded o_def]) - apply (rule finite_imageI) - apply (rule p') - using * by auto + unfolding * + apply (rule sum.reindex_nontrivial [symmetric, unfolded o_def]) + apply (rule finite_product_dependent [OF \finite d\]) + apply (rule finite_imageI [OF \finite p\]) + apply clarsimp + by (metis eq0 fst_conv snd_conv) qed - finally show ?case . + also have "\ = (\(x, k)\p'. norm (integral k f))" + proof - + have 0: "integral (ia \ snd (a, b)) f = 0" + 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 + apply (erule_tac x="(a, ia \ b)" in allE) + apply auto + done + 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 + show ?thesis + unfolding sum_p' + apply (rule sum.mono_neutral_right) + apply (subst *) + apply (rule finite_imageI[OF finite_product_dependent]) + apply fact + apply (rule finite_imageI[OF p'(1)]) + using 0 1 by auto + qed + finally show ?thesis . qed - also have "\ = (\(i,l) \ d \ (snd ` p). norm (integral (i\l) f))" - by (simp add: sum.cartesian_product) - also have "\ = (\x\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (case_prod op \ x) f))" - by (force simp: split_def Sigma_def intro!: sum.cong) - also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" + show "(\(x, k)\p'. norm (content k *\<^sub>R f x)) = (\(x, k)\p. content k *\<^sub>R norm (f x))" proof - - have eq0: " (integral (l1 \ k1) f) = 0" - if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" - "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" - for l1 l2 k1 k2 j1 j2 + let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" + have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" + by force + have pdfin: "finite (p \ d)" + using finite_cartesian_product[OF p'(1) d'(1)] by metis + have "(\(x, k)\p'. norm (content k *\<^sub>R f x)) = (\(x, k)\?S. \content k\ * norm (f x))" + unfolding norm_scaleR + apply (rule sum.mono_neutral_left) + apply (subst *) + apply (rule finite_imageI) + apply fact + unfolding p'alt apply blast + by fastforce + also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" + unfolding * + apply (subst sum.reindex_nontrivial) + apply fact + unfolding split_paired_all + unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject + apply (elim conjE) proof - - obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" - using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast - have "l1 \ l2 \ k1 \ k2" - using that by auto + 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 = {}" - by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) - moreover have "interior (l1 \ k1) = interior (l2 \ k2)" - by (simp add: that(1)) - ultimately have "interior(l1 \ k1) = {}" - by auto - then show ?thesis - unfolding uv Int_interval content_eq_0_interior[symmetric] by auto - qed - show ?thesis - unfolding * - apply (rule sum.reindex_nontrivial [symmetric, unfolded o_def]) - apply (rule finite_product_dependent [OF \finite d\]) - apply (rule finite_imageI [OF \finite p\]) - apply clarsimp - by (metis eq0 fst_conv snd_conv) - qed - also have "\ = (\(x, k)\p'. norm (integral k f))" - proof - - have 0: "integral (ia \ snd (a, b)) f = 0" - 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 - apply (erule_tac x="(a, ia \ b)" in allE) + apply - + apply (erule disjE) + apply (rule disjI2) + defer + apply (rule disjI1) + apply (rule d'(5)[OF as(3-4)]) + apply assumption + apply (rule p'(5)[OF as(1-2)]) apply auto done - then show ?thesis + 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 - have 1: "\i l. snd (a, b) = i \ l \ i \ d \ l \ snd ` p" if "(a, b) \ p'" for a b + qed safe + also have "\ = (\(x, k)\p. content k *\<^sub>R norm (f x))" + apply (subst sum_Sigma_product[symmetric]) + apply (rule p') + apply (rule d') + apply (rule sum.cong) + apply (rule refl) + unfolding split_paired_all split_conv 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 - show ?thesis - unfolding sum_p' - apply (rule sum.mono_neutral_right) - apply (subst *) - apply (rule finite_imageI[OF finite_product_dependent]) - apply fact - apply (rule finite_imageI[OF p'(1)]) - using 0 1 by auto - qed - finally show ?case . - next - case 3 - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" - have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" - by force - have pdfin: "finite (p \ d)" - using finite_cartesian_product[OF p'(1) d'(1)] by metis - have "(\(x, k)\p'. norm (content k *\<^sub>R f x)) = (\(x, k)\?S. \content k\ * norm (f x))" - unfolding norm_scaleR - apply (rule sum.mono_neutral_left) - apply (subst *) - apply (rule finite_imageI) - apply fact - unfolding p'alt apply blast - by fastforce - also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" - unfolding * - apply (subst sum.reindex_nontrivial) - apply fact - 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 - - apply (erule disjE) - apply (rule disjI2) - defer - apply (rule disjI1) - apply (rule d'(5)[OF as(3-4)]) - apply assumption - apply (rule p'(5)[OF as(1-2)]) - apply auto - done - 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 - also have "\ = (\(x, k)\p. content k *\<^sub>R norm (f x))" - apply (subst sum_Sigma_product[symmetric]) - apply (rule p') - apply (rule d') - apply (rule sum.cong) - apply (rule refl) - unfolding split_paired_all split_conv - proof - - fix x l - assume as: "(x, l) \ p" - note xl = p'(2-4)[OF this] - from this(3) guess u v by (elim exE) note uv=this - 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}" - unfolding Setcompr_eq_image - apply (rule sum.reindex_nontrivial [unfolded o_def, symmetric]) - apply (rule d') - proof goal_cases - case prems: (1 k y) - from d'(4)[OF this(1)] d'(4)[OF this(2)] - guess u1 v1 u2 v2 by (elim exE) note uv=this - have "{} = interior ((k \ y) \ cbox u v)" - apply (subst interior_Int) - using d'(5)[OF prems(1-3)] - apply auto - done - also have "\ = interior (y \ (k \ cbox u v))" + fix x l + assume as: "(x, l) \ p" + note xl = p'(2-4)[OF this] + from this(3) guess u v by (elim exE) note uv=this + 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}" + proof - + have eq0: "content (k \ cbox u v) = 0" + if "k \ d" "y \ d" "k \ y" and eq: "k \ cbox u v = y \ cbox u v" for k y + proof - + from d'(4)[OF that(1)] d'(4)[OF that(2)] + obtain \ \ where \: "k \ cbox u v = cbox \ \" + by (meson Int_interval) + have "{} = interior ((k \ y) \ cbox u v)" + by (simp add: d'(5) that) + also have "\ = interior (y \ (k \ cbox u v))" + by auto + also have "\ = interior (k \ cbox u v)" + unfolding eq by auto + finally show ?thesis + unfolding \ content_eq_0_interior .. + qed + then show ?thesis + unfolding Setcompr_eq_image + apply (rule sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) by auto - also have "\ = interior (k \ cbox u v)" - unfolding prems(4) by auto - finally show ?case - unfolding uv Int_interval content_eq_0_interior .. qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" apply (rule sum.mono_neutral_right) unfolding Setcompr_eq_image - apply (rule finite_imageI) - apply (rule d') + apply (rule finite_imageI [OF \finite d\]) apply (fastforce simp: inf.commute)+ done finally show "(\i\d. \content (l \ i)\ * norm (f x)) = content l *\<^sub>R norm (f x)" unfolding sum_distrib_right[symmetric] real_scaleR_def apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) - using xl(2)[unfolded uv] - unfolding uv - apply auto + using xl(2)[unfolded uv] unfolding uv apply auto done qed - finally show ?case . - qed - qed + finally show ?thesis . + qed + qed (rule d) + qed qed + then show ?thesis + using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] + by blast qed + lemma bounded_variation_absolutely_integrable: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" @@ -2161,7 +2165,7 @@ proof (rule *[rule_format]) show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e / 2" unfolding split_def - apply (rule helplemma) + apply (rule absdiff_norm_less) using d2(2)[rule_format,of p] using p(1,3) unfolding tagged_division_of_def split_def