# HG changeset patch # User paulson # Date 1501794396 -7200 # Node ID ea6cbb69dda2c54082611d19547c3aeb0be16845 # Parent 604616b627f266d752843fe4e010a64bea7c335e# Parent 9786b06c7b5ad1d1a4fe67e6e6490d38156f35a2 merged diff -r 604616b627f2 -r ea6cbb69dda2 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Aug 03 23:03:44 2017 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Aug 03 23:06:36 2017 +0200 @@ -7,6 +7,24 @@ imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral begin +lemma finite_product_dependent: (*FIXME DELETE*) + assumes "finite s" + and "\x. x \ s \ finite (t x)" + shows "finite {(i, j) |i j. i \ s \ j \ t i}" + using assms +proof induct + case (insert x s) + have *: "{(i, j) |i j. i \ insert x s \ j \ t i} = + (\y. (x,y)) ` (t x) \ {(i, j) |i j. i \ s \ j \ t i}" by auto + show ?case + unfolding * + apply (rule finite_UnI) + using insert + apply auto + done +qed auto + + lemma le_left_mono: "x \ y \ y \ a \ x \ (a::'a::preorder)" by (auto intro: order_trans) @@ -1633,11 +1651,16 @@ then show "\e>0. \i\d. x \ i \ ball x e \ i = {}" by force qed - from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format] - + 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 - from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format] + from henstock_lemma[OF assms(1) this] + obtain g where g: "gauge g" + "\p. \p tagged_partial_division_of cbox a b; g 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)" show ?case apply (rule_tac x="?g" in exI) @@ -1716,15 +1739,12 @@ assume y: "y \ cbox a b" then have "\x l. (x, l) \ p \ y\l" unfolding p'(6)[symmetric] by auto - then guess x l by (elim exE) note xl=conjunctD2[OF this] + then obtain x l where xl: "(x, l) \ p" "y \ l" by metis then have "\k. k \ d \ y \ k" using y unfolding d'(6)[symmetric] by auto - then guess i .. note i = conjunctD2[OF this] + then obtain i where i: "i \ d" "y \ i" by metis have "x \ i" - using fineD[OF p(3) xl(1)] - using k(2)[OF i(1), of x] - using i(2) xl(2) - by auto + 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) @@ -1735,12 +1755,7 @@ qed then have "(\(x, k)\p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2" - apply - - apply (rule g(2)[rule_format]) - unfolding tagged_division_of_def - apply safe - apply (rule gp') - done + 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" unfolding split_def using p'' diff -r 604616b627f2 -r ea6cbb69dda2 src/HOL/Analysis/Lebesgue_Integral_Substitution.thy --- a/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy Thu Aug 03 23:03:44 2017 +0200 +++ b/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy Thu Aug 03 23:06:36 2017 +0200 @@ -12,6 +12,7 @@ imports Interval_Integral begin + lemma nn_integral_substitution_aux: fixes f :: "real \ ennreal" assumes Mf: "f \ borel_measurable borel" diff -r 604616b627f2 -r ea6cbb69dda2 src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Thu Aug 03 23:03:44 2017 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Aug 03 23:06:36 2017 +0200 @@ -10,23 +10,6 @@ Topology_Euclidean_Space begin -lemma finite_product_dependent: (*FIXME DELETE*) - assumes "finite s" - and "\x. x \ s \ finite (t x)" - shows "finite {(i, j) |i j. i \ s \ j \ t i}" - using assms -proof induct - case (insert x s) - have *: "{(i, j) |i j. i \ insert x s \ j \ t i} = - (\y. (x,y)) ` (t x) \ {(i, j) |i j. i \ s \ j \ t i}" by auto - show ?case - unfolding * - apply (rule finite_UnI) - using insert - apply auto - done -qed auto - lemma sum_Sigma_product: assumes "finite S" and "\i. i \ S \ finite (T i)" @@ -210,18 +193,18 @@ subsection \The notion of a gauge --- simply an open set containing the point.\ -definition "gauge d \ (\x. x \ d x \ open (d x))" +definition "gauge \ \ (\x. x \ \ x \ open (\ x))" lemma gaugeI: - assumes "\x. x \ g x" - and "\x. open (g x)" - shows "gauge g" + assumes "\x. x \ \ x" + and "\x. open (\ x)" + shows "gauge \" using assms unfolding gauge_def by auto lemma gaugeD[dest]: - assumes "gauge d" - shows "x \ d x" - and "open (d x)" + assumes "gauge \" + shows "x \ \ x" + and "open (\ x)" using assms unfolding gauge_def by auto lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" @@ -233,7 +216,7 @@ lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" by (rule gauge_ball) auto -lemma gauge_Int[intro]: "gauge d1 \ gauge d2 \ gauge (\x. d1 x \ d2 x)" +lemma gauge_Int[intro]: "gauge \1 \ gauge \2 \ gauge (\x. \1 x \ \2 x)" unfolding gauge_def by auto lemma gauge_reflect: @@ -244,10 +227,10 @@ lemma gauge_Inter: assumes "finite S" - and "\d. d\S \ gauge (f d)" - shows "gauge (\x. \{f d x | d. d \ S})" + and "\u. u\S \ gauge (f u)" + shows "gauge (\x. \{f u x | u. u \ S})" proof - - have *: "\x. {f d x |d. d \ S} = (\d. f d x) ` S" + have *: "\x. {f u x |u. u \ S} = (\u. f u x) ` S" by auto show ?thesis unfolding gauge_def unfolding * @@ -314,18 +297,15 @@ assume "s = {{a}}" "K\s" then have "\x y. K = cbox x y" apply (rule_tac x=a in exI)+ - apply (force simp: cbox_sing) + apply force done } ultimately show ?l unfolding division_of_def cbox_sing by auto next assume ?l - { - fix x - assume x: "x \ s" have "x = {a}" - by (metis \s division_of cbox a a\ cbox_sing division_ofD(2) division_ofD(3) subset_singletonD x) - } + have "x = {a}" if "x \ s" for x + by (metis \s division_of cbox a a\ cbox_sing division_ofD(2) division_ofD(3) subset_singletonD that) moreover have "s \ {}" using \s division_of cbox a a\ by auto ultimately show ?r @@ -401,11 +381,8 @@ have *: "\s. \{x\s. x \ {}} = \s" by auto show "\?A = s1 \ s2" - apply (rule set_eqI) - unfolding * and UN_iff - using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] - apply auto - done + unfolding * + using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto { fix k assume "k \ ?A" @@ -461,14 +438,14 @@ unfolding * by auto qed -lemma elementary_inter: +lemma elementary_Int: fixes s t :: "'a::euclidean_space set" assumes "p1 division_of s" and "p2 division_of t" shows "\p. p division_of (s \ t)" using assms division_inter by blast -lemma elementary_inters: +lemma elementary_Inter: assumes "finite f" and "f \ {}" and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" @@ -488,7 +465,7 @@ moreover obtain px where "px division_of x" using insert(5)[rule_format,OF insertI1] .. ultimately show ?thesis - by (simp add: elementary_inter Inter_insert) + by (simp add: elementary_Int Inter_insert) qed qed auto @@ -548,13 +525,9 @@ show "cbox c d \ p" unfolding p_def by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) - { - fix i :: 'a - assume "i \ Basis" - with incl nonempty have "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" - unfolding box_eq_empty subset_box by (auto simp: not_le) - } - note ord = this + have ord: "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" if "i \ Basis" for i + using incl nonempty that + unfolding box_eq_empty subset_box by (auto simp: not_le) show "p division_of (cbox a b)" proof (rule division_ofI) @@ -665,8 +638,8 @@ then have di: "\x. x \ p \ \d. d division_of \(q x - {x})" by (meson Diff_subset division_of_subset) have "\d. d division_of \((\i. \(q i - {i})) ` p)" - apply (rule elementary_inters [OF finite_imageI[OF p(1)]]) - apply (auto dest: di simp: False elementary_inters [OF finite_imageI[OF p(1)]]) + apply (rule elementary_Inter [OF finite_imageI[OF p(1)]]) + apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]]) done then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. have "d \ p division_of cbox a b" @@ -719,12 +692,7 @@ obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" proof (cases "cbox c d = {}") case True - show ?thesis - apply (rule that[of "{}"]) - unfolding True - using assms - apply auto - done + with assms that show ?thesis by force next case False show ?thesis @@ -812,14 +780,13 @@ by auto show "finite ?D" using "*" pdiv(1) q(1) by auto - have lem1: "\f s. \\(f ` s) = \((\x. \(f x)) ` s)" - by auto - have lem2: "\f s. f \ {} \ \{s \ t |t. t \ f} = s \ \f" + have "\?D = (\x\p. \insert (cbox a b) (q x))" by auto - show "\?D = cbox a b \ \p" - unfolding * lem1 - unfolding lem2[OF \p \ {}\, of "cbox a b", symmetric] - using q(6) by auto + also have "... = \{cbox a b \ t |t. t \ p}" + using q(6) by auto + also have "... = cbox a b \ \p" + using \p \ {}\ by auto + finally show "\?D = cbox a b \ \p" . show "K \ cbox a b \ \p" "K \ {}" if "K \ ?D" for K using q that by blast+ show "\a b. K = cbox a b" if "K \ ?D" for K @@ -920,7 +887,7 @@ apply auto done then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" - by (metis assms(2) divq(6) elementary_inter) + by (metis assms(2) divq(6) elementary_Int) { fix x assume x: "x \ T" "x \ S" @@ -1130,9 +1097,8 @@ assumes "s tagged_partial_division_of i" and "t \ s" shows "t tagged_partial_division_of i" - using assms + using assms finite_subset[OF assms(2)] unfolding tagged_partial_division_of_def - using finite_subset[OF assms(2)] by blast lemma tag_in_interval: "p tagged_division_of i \ (x, k) \ p \ x \ i" @@ -1183,28 +1149,28 @@ qed lemma tagged_division_unions: - assumes "finite iset" - and "\i\iset. pfn i tagged_division_of i" - and "\i1\iset. \i2\iset. i1 \ i2 \ interior(i1) \ interior(i2) = {}" - shows "\(pfn ` iset) tagged_division_of (\iset)" + assumes "finite I" + and "\i\I. pfn i tagged_division_of i" + and "\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]] - show "finite (\(pfn ` iset))" + show "finite (\(pfn ` I))" using assms by auto - have "\{k. \x. (x, k) \ \(pfn ` iset)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` iset)" + have "\{k. \x. (x, k) \ \(pfn ` I)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` I)" by blast - also have "\ = \iset" + also have "\ = \I" using assm(6) by auto - finally show "\{k. \x. (x, k) \ \(pfn ` iset)} = \iset" . + finally show "\{k. \x. (x, k) \ \(pfn ` I)} = \I" . fix x k - assume xk: "(x, k) \ \(pfn ` iset)" - then obtain i where i: "i \ iset" "(x, k) \ pfn i" + assume xk: "(x, k) \ \(pfn ` I)" + then obtain i where i: "i \ I" "(x, k) \ pfn i" by auto - show "x \ k" "\a b. k = cbox a b" "k \ \iset" + show "x \ k" "\a b. k = cbox a b" "k \ \I" using assm(2-4)[OF i] using i(1) by auto fix x' k' - assume xk': "(x', k') \ \(pfn ` iset)" "(x, k) \ (x', k')" - then obtain i' where i': "i' \ iset" "(x', k') \ pfn i'" + assume xk': "(x', k') \ \(pfn ` I)" "(x, k) \ (x', k')" + 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) @@ -1379,8 +1345,8 @@ lemma division_points_psubset: fixes a :: "'a::euclidean_space" - assumes "d division_of (cbox a b)" - and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + assumes d: "d division_of (cbox a b)" + and altb: "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and "l \ d" and disj: "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k: "k \ Basis" @@ -1390,14 +1356,12 @@ division_points (cbox a b) d" (is "?D2 \ ?D") proof - have ab: "\i\Basis. a\i \ b\i" - using assms(2) by (auto intro!:less_imp_le) - guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this + using altb by (auto intro!:less_imp_le) + obtain u v where l: "l = cbox u v" + using d \l \ d\ by blast have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" - using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty - using subset_box(1) - apply auto - apply blast+ - done + apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l) + by (metis assms(5) box_ne_empty(1) cbox_division_memE d l subset_box(1)) have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" unfolding l interval_split[OF k] interval_bounds[OF uv(1)] @@ -1406,11 +1370,7 @@ have "\x. x \ ?D - ?D1" using assms(3-) unfolding division_points_def interval_bounds[OF ab] - apply - - apply (erule disjE) - apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) - apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) - done + by (force simp add: *) moreover have "?D1 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D1 \ ?D" @@ -1423,11 +1383,7 @@ have "\x. x \ ?D - ?D2" using assms(3-) unfolding division_points_def interval_bounds[OF ab] - apply - - apply (erule disjE) - apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) - apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) - done + by (force simp add: *) moreover have "?D2 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D2 \ ?D" @@ -1769,39 +1725,28 @@ done next fix a b c :: real - assume "\a b. b \ a \ g {a .. b} = \<^bold>1" - and "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" + assume eq1: "\a b. b \ a \ g {a .. b} = \<^bold>1" + and eqg: "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" and "a \ c" and "c \ b" - note as = this[rule_format] show "g {a..c} \<^bold>* g {c..b} = g {a..b}" proof (cases "c = a \ c = b") case False then show ?thesis - apply - - apply (subst as(2)) - using as(3-) - apply auto - done + using eqg \a \ c\ \c \ b\ by auto next case True then show ?thesis proof assume *: "c = a" then have "g {a .. c} = \<^bold>1" - apply - - apply (rule as(1)[rule_format]) - apply auto - done + using eq1 by blast then show ?thesis unfolding * by auto next assume *: "c = b" then have "g {c .. b} = \<^bold>1" - apply - - apply (rule as(1)[rule_format]) - apply auto - done + using eq1 by blast then show ?thesis unfolding * by auto qed @@ -1934,18 +1879,18 @@ subsection \Some basic combining lemmas.\ lemma tagged_division_Union_exists: - assumes "finite iset" - and "\i\iset. \p. p tagged_division_of i \ d fine p" - and "\i1\iset. \i2\iset. i1 \ i2 \ interior i1 \ interior i2 = {}" - and "\iset = i" + assumes "finite I" + and "\i\I. \p. p tagged_division_of i \ d fine p" + and "\i1\I. \i2\I. i1 \ i2 \ interior i1 \ interior i2 = {}" + and "\I = i" obtains p where "p tagged_division_of i" and "d fine p" proof - obtain pfn where pfn: - "\x. x \ iset \ pfn x tagged_division_of x" - "\x. x \ iset \ d fine pfn x" + "\x. x \ I \ pfn x tagged_division_of x" + "\x. x \ I \ d fine pfn x" using bchoice[OF assms(2)] by auto show thesis - apply (rule_tac p="\(pfn ` iset)" in that) + apply (rule_tac p="\(pfn ` I)" in that) using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force by (metis (mono_tags, lifting) fine_Union imageE pfn(2)) qed @@ -1972,25 +1917,22 @@ using assms(1,3) by metis then have ab: "\i. i\Basis \ a \ i \ b \ i" by (force simp: mem_box) - { fix f - have "\finite f; + have UN_cases: "\finite f; \s. s\f \ P s; \s. s\f \ \a b. s = cbox a b; - \s t. s\f \ t\f \ s \ t \ interior s \ interior t = {}\ \ P (\f)" - proof (induct f rule: finite_induct) - case empty - show ?case - using assms(1) by auto - next - case (insert x f) - show ?case - unfolding Union_insert - apply (rule assms(2)[rule_format]) - using Int_interior_Union_intervals [of f "interior x"] - apply (auto simp: insert) - by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff) - qed - } note UN_cases = this + \s t. s\f \ t\f \ s \ t \ interior s \ interior t = {}\ \ P (\f)" for f + proof (induct f rule: finite_induct) + case empty + show ?case + using assms(1) by auto + next + case (insert x f) + show ?case + unfolding Union_insert + apply (rule assms(2)[rule_format]) + using Int_interior_Union_intervals [of f "interior x"] + by (metis (no_types, lifting) insert insert_iff open_interior) + qed let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ (c\i = (a\i + b\i) / 2) \ (d\i = b\i)}" let ?PP = "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" @@ -2110,13 +2052,12 @@ then show "\c d. ?P i c d" by blast qed + then obtain \ \ where + "\i\Basis. (\ \ i = a \ i \ \ \ i = (a \ i + b \ i) / 2 \ + \ \ i = (a \ i + b \ i) / 2 \ \ \ i = b \ i) \ \ \ i \ x \ i \ x \ i \ \ \ i" + by (auto simp: choice_Basis_iff) then show "x\\?A" - unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff - apply auto - apply (rule_tac x="cbox xa xaa" in exI) - unfolding mem_box - apply auto - done + by (force simp add: mem_box) qed finally show False using assms by auto @@ -2148,10 +2089,7 @@ 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" by (rule interval_bisection_step[of P, OF assms(1-2) as]) then show ?thesis - apply - - apply (rule_tac x="(c,d)" in exI) - apply auto - done + by (rule_tac x="(c,d)" in exI) auto qed qed then obtain f where f: @@ -2162,11 +2100,7 @@ fst x \ i \ fst (f x) \ i \ fst (f x) \ i \ snd (f x) \ i \ snd (f x) \ i \ snd x \ i \ - 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" - apply - - apply (drule choice) - apply blast - done + 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" by metis define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ @@ -2179,10 +2113,7 @@ proof (induct n) case 0 then show ?case - unfolding S - apply (rule f[rule_format]) using assms(3) - apply auto - done + unfolding S using \\ P (cbox a b)\ f by auto next case (Suc n) show ?case @@ -2230,8 +2161,7 @@ next case (Suc n) have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" - using AB(3) that - using AB(4)[of i n] using i by auto + using AB(3) that AB(4)[of i n] using i by auto also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" using Suc by (auto simp add: field_simps) finally show ?case . @@ -2295,13 +2225,13 @@ fixes a b :: "'a::euclidean_space" assumes "gauge g" obtains p where "p tagged_division_of (cbox a b)" "g fine p" -proof - - presume "\ (\p. p tagged_division_of (cbox a b) \ g fine p) \ False" - then obtain p where "p tagged_division_of (cbox a b)" "g fine p" - by blast - then show thesis .. +proof (cases "\p. p tagged_division_of (cbox a b) \ g fine p") + case True + then show ?thesis + using that by auto next - assume as: "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" + case False + assume "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" obtain x where x: "x \ (cbox a b)" "\e. 0 < e \ @@ -2310,10 +2240,10 @@ cbox c d \ ball x e \ cbox c d \ (cbox a b) \ \ (\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 _ _ as]) + 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 (auto simp: ) + apply auto done obtain e where e: "e > 0" "ball x e \ g x" using gaugeD[OF assms, of x] unfolding open_contains_ball by auto @@ -2325,7 +2255,7 @@ by blast have "g fine {(x, cbox c d)}" unfolding fine_def using e using c_d(2) by auto - then show False + then show ?thesis using tagged_division_of_self[OF c_d(1)] using c_d by auto qed @@ -2369,70 +2299,52 @@ proof (induct p) case empty show ?case - apply (rule_tac x="{}" in exI) - unfolding fine_def - apply auto - done + by (force simp add: fine_def) next case (insert xk p) - guess x k using surj_pair[of xk] by (elim exE) note xk=this - note tagged_partial_division_subset[OF insert(4) subset_insertI] - from insert(3)[OF this insert(5)] + obtain x k where xk: "xk = (x, k)" + using surj_pair[of xk] by metis obtain q1 where q1: "q1 tagged_division_of \{k. \x. (x, k) \ p}" and "d fine q1" and q1I: "\x k. \(x, k)\p; k \ d x\ \ (x, k) \ q1" - by (force simp add: ) + using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI] + by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2)) have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" unfolding xk by auto note p = tagged_partial_division_ofD[OF insert(4)] - from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this - + obtain u v where uv: "k = cbox u v" + using p(4)[unfolded xk, OF insertI1] by blast have "finite {k. \x. (x, k) \ p}" apply (rule finite_subset[of _ "snd ` p"]) - using p - apply safe - apply (metis image_iff snd_conv) - apply auto - done + using image_iff apply fastforce + using insert.hyps(1) by blast then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" - apply (rule Int_interior_Union_intervals) - apply (rule open_interior) - unfolding mem_Collect_eq - apply (erule_tac[!] exE) - apply (drule p(4)[OF insertI2]) - apply assumption - apply (rule p(5)) - unfolding uv xk - apply (rule insertI1) - apply (rule insertI2) - apply assumption - using insert(2) - unfolding uv xk - apply auto - done + proof (rule Int_interior_Union_intervals) + show "open (interior (cbox u v))" + by auto + show "\T. T \ {k. \x. (x, k) \ p} \ \a b. T = cbox a b" + using p(4) by auto + show "\T. T \ {k. \x. (x, k) \ p} \ interior (cbox u v) \ interior T = {}" + by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk) + qed show ?case proof (cases "cbox u v \ d x") case True - then show ?thesis + 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) + with True show ?thesis apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) - apply rule - unfolding * uv - apply (rule tagged_division_union) - apply (rule tagged_division_of_self) - apply (rule p[unfolded xk uv] insertI1)+ - apply (rule q1) - apply (rule int) - apply rule - apply (rule fine_Un) - apply (subst fine_def) - apply (auto simp add: \d fine q1\ q1I uv xk) + using \d fine q1\ fine_def q1I uv xk apply fastforce done next case False - from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this + obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2" + using fine_division_exists[OF assms(2)] by blast show ?thesis apply (rule_tac x="q2 \ q1" in exI) - apply rule + apply (intro conjI) unfolding * uv apply (rule tagged_division_union q2 q1 int fine_Un)+ apply (auto intro: q1 q2 fine_Un \d fine q1\ simp add: False q1I uv xk)