# HG changeset patch # User paulson # Date 1433803868 -3600 # Node ID b699cedd04e8154514b818fe229832c8b65f51b1 # Parent fd42b2f41404788387cfd48cedbb00f806f8f519 tidying messy proofs diff -r fd42b2f41404 -r b699cedd04e8 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Jun 08 18:56:04 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Jun 08 23:51:08 2015 +0100 @@ -241,8 +241,9 @@ and "\t\f. \a b. t = cbox a b" and "\t\f. s \ (interior t) = {}" shows "s \ interior (\f) = {}" -proof (rule ccontr, unfold ex_in_conv[symmetric]) - case goal1 +proof (clarsimp simp only: all_not_in_conv [symmetric]) + fix x + assume x: "x \ s" "x \ interior (\f)" have lem1: "\x e s U. ball x e \ s \ interior U \ ball x e \ s \ U" using interior_subset by auto (meson Topology_Euclidean_Space.open_ball contra_subsetD interior_maximal mem_ball) @@ -308,17 +309,14 @@ let ?z = "x - (e/2) *\<^sub>R k" assume as: "x\k = a\k" have "ball ?z (e / 2) \ i = {}" - apply (rule ccontr) - unfolding ex_in_conv[symmetric] - apply (erule exE) - proof - + proof (clarsimp simp only: all_not_in_conv [symmetric]) fix y - assume "y \ ball ?z (e / 2) \ i" - then have "dist ?z y < e/2" and yi:"y\i" by auto + assume "y \ ball ?z (e / 2)" and yi: "y \ i" + then have "dist ?z y < e/2" by auto then have "\(?z - y) \ k\ < e/2" using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto then have "y\k < a\k" - using e[THEN conjunct1] k + using e k by (auto simp add: field_simps abs_less_iff as inner_simps) then have "y \ i" unfolding ab mem_box by (auto intro!: bexI[OF _ k]) @@ -337,10 +335,8 @@ done also have "\ < \e\ / 2 + \e\ / 2" apply (rule add_strict_left_mono) - using as - unfolding mem_ball dist_norm - using e - apply (auto simp add: field_simps) + using as e + apply (auto simp add: field_simps dist_norm) done finally show "y \ ball x e" unfolding mem_ball dist_norm using e by (auto simp add:field_simps) @@ -354,19 +350,16 @@ let ?z = "x + (e/2) *\<^sub>R k" assume as: "x\k = b\k" have "ball ?z (e / 2) \ i = {}" - apply (rule ccontr) - unfolding ex_in_conv[symmetric] - apply (erule exE) - proof - + proof (clarsimp simp only: all_not_in_conv [symmetric]) fix y - assume "y \ ball ?z (e / 2) \ i" - then have "dist ?z y < e/2" and yi: "y \ i" + assume "y \ ball ?z (e / 2)" and yi: "y \ i" + then have "dist ?z y < e/2" by auto then have "\(?z - y) \ k\ < e/2" using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto then have "y\k > b\k" - using e[THEN conjunct1] k + using e k by (auto simp add:field_simps inner_simps inner_Basis as) then have "y \ i" unfolding ab mem_box by (auto intro!: bexI[OF _ k]) @@ -400,14 +393,14 @@ then obtain x where "ball x (e / 2) \ s \ \f" .. then have "x \ s \ interior (\f)" unfolding lem1[where U="\f", symmetric] - using centre_in_ball e[THEN conjunct1] by auto + using centre_in_ball e by auto then show ?thesis using insert.hyps(3) insert.prems(1) by blast qed qed qed qed - from this[OF assms(1,3) goal1] + from this[OF assms(1,3)] x obtain t x e where "t \ f" "0 < e" "ball x e \ s \ t" by blast then have "x \ s" "x \ interior t" @@ -564,14 +557,15 @@ lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" - apply rule - defer - apply (rule content_pos_lt, assumption) -proof - +proof (rule iffI) assume "0 < content (cbox a b)" then have "content (cbox a b) \ 0" by auto then show "\i\Basis. a\i < b\i" unfolding content_eq_0 not_ex not_le by fastforce +next + assume "\i\Basis. a \ i < b \ i" + then show "0 < content (cbox a b)" + by (metis content_pos_lt) qed lemma content_empty [simp]: "content {} = 0" @@ -593,22 +587,16 @@ have "cbox c d \ {}" using assms False by auto then have cd_ne: "\i\Basis. c \ i \ d \ i" using assms unfolding box_ne_empty by auto - show ?thesis - unfolding content_def - unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] - unfolding if_not_P[OF False] if_not_P[OF `cbox c d \ {}`] - apply (rule setprod_mono) - apply rule - proof - fix i :: 'a - assume i: "i \ Basis" - show "0 \ b \ i - a \ i" - using ab_ne[THEN bspec, OF i] i by auto - show "b \ i - a \ i \ d \ i - c \ i" - using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2),of i] - using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1),of i] - using i by auto - qed + have "\i. i \ Basis \ 0 \ b \ i - a \ i" + using ab_ne by (metis diff_le_iff(1)) + moreover + have "\i. i \ Basis \ b \ i - a \ i \ d \ i - c \ i" + using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)] + assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)] + by (metis diff_mono) + ultimately show ?thesis + unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] + by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF `cbox c d \ {}`]) qed lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" @@ -1069,36 +1057,26 @@ qed obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" using bchoice[OF *] by blast - have "\x. x \ p \ \d. d division_of \(q x - {x})" - apply rule - apply (rule_tac p="q x" in division_of_subset) - proof - - fix x + { fix x assume x: "x \ p" - show "q x division_of \q x" + have "q x division_of \q x" apply (rule division_ofI) using division_ofD[OF q(1)[OF x]] apply auto - done - show "q x - {x} \ q x" - by auto - qed + done } + then have "\x. x \ p \ \d. d division_of \(q x - {x})" + by (meson Diff_subset division_of_subset) then have "\d. d division_of \ ((\i. \(q i - {i})) ` p)" apply - - apply (rule elementary_inters) - apply (rule finite_imageI[OF p(1)]) - unfolding image_is_empty - apply (rule False) - apply auto + apply (rule elementary_inters [OF finite_imageI[OF p(1)]]) + apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]]) done then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. - show ?thesis - apply (rule that[of "d \ p"]) + have "d \ p division_of cbox a b" proof - - have *: "\s f t. s \ {} \ \i\s. f i \ i = t \ t = \(f ` s) \ \s" by auto + have te: "\s f t. s \ {} \ \i\s. f i \ i = t \ t = \(f ` s) \ \s" by auto have *: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" - apply (rule *[OF False]) - proof + proof (rule te[OF False], clarify) fix i assume i: "i \ p" show "\(q i - {i}) \ i = cbox a b" @@ -1135,7 +1113,9 @@ done qed qed - qed auto + qed + then show ?thesis + by (meson Un_upper2 that) qed lemma elementary_bounded[dest]: @@ -1938,11 +1918,7 @@ show "P s" unfolding s apply (rule as[rule_format]) - proof - - case goal1 - then show ?case - using s(2)[of i] using ab[OF `i \ Basis`] by auto - qed + using ab s(2) by force show "\a b. s = cbox a b" unfolding s by auto fix t @@ -1959,17 +1935,8 @@ then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" unfolding euclidean_eq_iff[where 'a='a] by auto then have i: "c\i \ e\i" "d\i \ f\i" - apply - - apply(erule_tac[!] disjE) - proof - - assume "c\i \ e\i" - then show "d\i \ f\i" - using s(2)[OF i'] t(2)[OF i'] by fastforce - next - assume "d\i \ f\i" - then show "c\i \ e\i" - using s(2)[OF i'] t(2)[OF i'] by fastforce - qed + using s(2) t(2) apply fastforce + using t(2)[OF i'] `c \ i \ e \ i \ d \ i \ f \ i` i' s(2) t(2) by fastforce have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto show "interior s \ interior t = {}" @@ -1979,16 +1946,9 @@ assume "x \ box c d" "x \ box e f" then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" unfolding mem_box using i' - apply - - apply (erule_tac[!] x=i in ballE)+ - apply auto - done - show False - using s(2)[OF i'] - apply - - apply (erule_tac disjE) - apply (erule_tac[!] conjE) - proof - + by force+ + show False using s(2)[OF i'] + proof safe assume as: "c \ i = a \ i" "d \ i = (a \ i + b \ i) / 2" show False using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps) @@ -2007,7 +1967,8 @@ "x \ cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast + c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" + by blast show "x\cbox a b" unfolding mem_box proof safe @@ -2056,7 +2017,7 @@ 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" proof case goal1 - then show ?case + show ?case proof - presume "\ P (cbox (fst x) (snd x)) \ ?thesis" then show ?thesis by (cases "P (cbox (fst x) (snd x))") auto