# HG changeset patch # User paulson # Date 1501740555 -7200 # Node ID ce50687a700e871a90d74e070274975afbd2b0d6 # Parent ca985e87c123fbbe8fe4ec53df08b052657fd609# Parent 52914a618299e8f509e4e1b42729420299979992 merged diff -r ca985e87c123 -r ce50687a700e src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Thu Aug 03 07:31:25 2017 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Aug 03 08:09:15 2017 +0200 @@ -1379,8 +1379,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 +1390,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)] @@ -2375,64 +2373,49 @@ done 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)