# HG changeset patch # User paulson # Date 1605050464 0 # Node ID aeac6424d3b51b56d9455c72324cf8aa1a1c778e # Parent e1d04777d8b68add55628e33d8150118d2aac09f cleanup diff -r e1d04777d8b6 -r aeac6424d3b5 src/HOL/Analysis/Affine.thy --- a/src/HOL/Analysis/Affine.thy Sat Nov 07 23:20:43 2020 +0000 +++ b/src/HOL/Analysis/Affine.thy Tue Nov 10 23:21:04 2020 +0000 @@ -1483,34 +1483,15 @@ lemma affine_independent_card_dim_diffs: fixes S :: "'a :: euclidean_space set" assumes "\ affine_dependent S" "a \ S" - shows "card S = dim {x - a|x. x \ S} + 1" + shows "card S = dim ((\x. x - a) ` S) + 1" proof - - have 1: "{b - a|b. b \ (S - {a})} \ {x - a|x. x \ S}" by auto - have 2: "x - a \ span {b - a |b. b \ S - {a}}" if "x \ S" for x - proof (cases "x = a") - case True then show ?thesis by (simp add: span_clauses) - next - case False then show ?thesis - using assms by (blast intro: span_base that) - qed - have "\ affine_dependent (insert a S)" + have non: "\ affine_dependent (insert a S)" by (simp add: assms insert_absorb) - then have 3: "independent {b - a |b. b \ S - {a}}" - using dependent_imp_affine_dependent by fastforce - have "{b - a |b. b \ S - {a}} = (\b. b-a) ` (S - {a})" - by blast - then have "card {b - a |b. b \ S - {a}} = card ((\b. b-a) ` (S - {a}))" - by simp - also have "\ = card (S - {a})" - by (metis (no_types, lifting) card_image diff_add_cancel inj_onI) - also have "\ = card S - 1" - by (simp add: aff_independent_finite assms) - finally have 4: "card {b - a |b. b \ S - {a}} = card S - 1" . have "finite S" by (meson assms aff_independent_finite) with \a \ S\ have "card S \ 0" by auto - moreover have "dim {x - a |x. x \ S} = card S - 1" - using 2 by (blast intro: dim_unique [OF 1 _ 3 4]) + moreover have "dim ((\x. x - a) ` S) = card S - 1" + using aff_dim_eq_dim_subtract aff_dim_unique \a \ S\ hull_inc insert_absorb non by fastforce ultimately show ?thesis by auto qed diff -r e1d04777d8b6 -r aeac6424d3b5 src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Sat Nov 07 23:20:43 2020 +0000 +++ b/src/HOL/Analysis/Starlike.thy Tue Nov 10 23:21:04 2020 +0000 @@ -151,15 +151,14 @@ have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) - have "z \ interior (ball c d)" - using y \0 < e\ \e \ 1\ - apply (simp add: interior_open[OF open_ball] z_def dist_norm) - by (simp add: field_simps norm_minus_commute) + have "(1 - e) * norm (x - y) / e < d" + using y \0 < e\ by (simp add: field_simps norm_minus_commute) + then have "z \ interior (ball c d)" + using \0 < e\ \e \ 1\ by (simp add: interior_open[OF open_ball] z_def dist_norm) then have "z \ interior S" using d interiorI interior_ball by blast then show ?thesis - unfolding * - using mem_interior_convex_shrink \y \ S\ assms by blast + unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: @@ -303,19 +302,19 @@ fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" - using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \e > 0\ + using as[THEN subsetD[where c="x - (e/2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next - have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ + have **: "dist x (x + (e/2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) - have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = + have "\i. i \ Basis \ (x + (e/2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) - then have *: "sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = + then have *: "sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) - have "sum ((\) x) Basis < sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" + have "sum ((\) x) Basis < sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force @@ -404,8 +403,8 @@ lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = - {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" - (is "rel_interior (convex hull (insert 0 ?p)) = ?s") + {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" + (is "_ = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast @@ -416,51 +415,48 @@ using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False - have h0: "affine hull (convex hull (insert 0 ?p)) = - {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" + have h0: "affine hull (convex hull (insert 0 D)) = + {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" - assume x: "x \ rel_interior (convex hull (insert 0 ?p))" + assume x: "x \ rel_interior (convex hull (insert 0 D))" then obtain e where "e > 0" and - "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 ?p)" - using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto - then have as [rule_format]: "\y. dist x y < e \ (\i\Basis. i \ D \ y\i = 0) \ - (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" - unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto + "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 D)" + using mem_rel_interior_ball[of x "convex hull (insert 0 D)"] h0 by auto + then have as: "\y. \dist x y < e \ (\i\Basis. i \ D \ y\i = 0)\ \ + (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" + using assms by (force simp: substd_simplex) have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" - then have "\j\D. 0 \ (x - (e / 2) *\<^sub>R i) \ j" + then have "\j\D. 0 \ (x - (e/2) *\<^sub>R i) \ j" using D \e > 0\ x0 - by (force simp: dist_norm inner_simps inner_Basis intro!: as[THEN conjunct1]) + by (intro as[THEN conjunct1]) (force simp: dist_norm inner_simps inner_Basis) then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto - then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" - using \e > 0\ norm_Basis[of a] D - unfolding dist_norm - by auto - have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" + then have **: "dist x (x + (e/2) *\<^sub>R a) < e" + using \e > 0\ norm_Basis[of a] D by (auto simp: dist_norm) + have "\i. i \ Basis \ (x + (e/2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) - then have *: "sum ((\) (x + (e / 2) *\<^sub>R a)) D = - sum (\i. x\i + (if a = i then e/2 else 0)) D" + then have *: "sum ((\) (x + (e/2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto - then have h1: "(\i\Basis. i \ D \ (x + (e / 2) *\<^sub>R a) \ i = 0)" + then have h1: "(\i\Basis. i \ D \ (x + (e/2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) - have "sum ((\) x) D < sum ((\) (x + (e / 2) *\<^sub>R a)) D" + have "sum ((\) x) D < sum ((\) (x + (e/2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" - using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] + using ** h1 as[rule_format, of "x + (e/2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto @@ -476,31 +472,28 @@ ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis - then have h2: "x \ convex hull (insert 0 ?p)" - using as assms - unfolding substd_simplex[OF assms] by fastforce + then have h2: "x \ convex hull (insert 0 D)" + using as assms by (force simp add: substd_simplex) obtain a where a: "a \ D" using \D \ {}\ by auto - let ?d = "(1 - sum ((\) x) D) / real (card D)" - have "0 < card D" using \D \ {}\ \finite D\ - by (simp add: card_gt_0_iff) - have "Min (((\) x) ` D) > 0" - using as \D \ {}\ \finite D\ by (simp) - moreover have "?d > 0" using as using \0 < card D\ by auto - ultimately have h3: "min (Min (((\) x) ` D)) ?d > 0" - by auto - have "\e>0. ball x e \ {x. \i\Basis. i \ D \ x \ i = 0} - \ convex hull insert 0 D" + define d where "d \ (1 - sum ((\) x) D) / real (card D)" + have "\e>0. ball x e \ {x. \i\Basis. i \ D \ x \ i = 0} \ convex hull insert 0 D" unfolding substd_simplex[OF assms] - apply (rule_tac x="min (Min (((\) x) ` D)) ?d" in exI) - apply (rule, rule h3) - apply safe - unfolding mem_ball - proof - + proof (intro exI; safe) + have "0 < card D" using \D \ {}\ \finite D\ + by (simp add: card_gt_0_iff) + have "Min (((\) x) ` D) > 0" + using as \D \ {}\ \finite D\ by (simp) + moreover have "d > 0" + using as \0 < card D\ by (auto simp: d_def) + ultimately show "min (Min (((\) x) ` D)) d > 0" + by auto fix y :: 'a - assume y: "dist x y < min (Min ((\) x ` D)) ?d" assume y2: "\i\Basis. i \ D \ y\i = 0" - have "sum ((\) y) D \ sum (\i. x\i + ?d) D" + assume "y \ ball x (min (Min ((\) x ` D)) d)" + then have y: "dist x y < min (Min ((\) x ` D)) d" + by auto + have "sum ((\) y) D \ sum (\i. x\i + d) D" proof (rule sum_mono) fix i assume "i \ D" @@ -508,13 +501,13 @@ by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) - also have "... < ?d" + also have "... < d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) - finally have "\y\i - x\i\ < ?d" . - then show "y \ i \ x \ i + ?d" by auto + finally have "\y\i - x\i\ < d" . + then show "y \ i \ x \ i + d" by auto qed also have "\ \ 1" - unfolding sum.distrib sum_constant using \0 < card D\ + unfolding sum.distrib sum_constant d_def using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . @@ -524,15 +517,14 @@ proof (cases "i\D") case True have "norm (x - y) < x\i" - using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] - using Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ - by (simp add: card_gt_0_iff) + using y Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ + by (simp add: dist_norm card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (use y2 in auto) qed - then have "x \ rel_interior (convex hull (insert 0 ?p))" + then have "x \ rel_interior (convex hull (insert 0 D))" using h0 h2 rel_interior_ball by force } ultimately have @@ -549,8 +541,7 @@ obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - - let ?D = D - let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D" + let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) D" have "finite D" using assms finite_Basis infinite_super by blast then have d1: "0 < real (card D)" @@ -558,7 +549,7 @@ { fix i assume "i \ D" - have "?a \ i = sum (\j. if i = j then inverse (2 * real (card D)) else 0) ?D" + have "?a \ i = sum (\j. if i = j then inverse (2 * real (card D)) else 0) D" unfolding inner_sum_left using \i \ D\ by (auto simp: inner_Basis subsetD[OF assms(2)] intro: sum.cong) also have "... = inverse (2 * real (card D))" @@ -579,12 +570,12 @@ by auto finally show "0 < ?a \ i" by auto next - have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real (card D))) ?D" + have "sum ((\) ?a) D = sum (\i. inverse (2 * real (card D))) D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) - finally show "sum ((\) ?a) ?D < 1" by auto + finally show "sum ((\) ?a) D < 1" by auto next fix i assume "i \ Basis" and "i \ D" @@ -768,10 +759,7 @@ using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" - apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) - using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \x \ a\ - apply simp - done + using "*" \x \ a\ e1 by force } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto @@ -781,13 +769,7 @@ ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto - next - case True - then have "rel_interior S = {}" by auto - then have "closure (rel_interior S) = {}" - using closure_empty by auto - with True show ?thesis by auto - qed + qed auto qed lemma rel_interior_same_affine_hull: @@ -822,7 +804,7 @@ lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" - obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" + obtains e where "0 < e" "e < 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" @@ -832,14 +814,10 @@ by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def assms - apply (simp add: algebra_simps) - using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] add_divide_distrib[of a b "a+b"] - apply auto - done + by (simp add: divide_simps vector_fraction_eq_iff) (simp add: algebra_simps) finally have "z = y - e *\<^sub>R (y-x)" by auto - moreover have "e > 0" using e_def assms by auto - moreover have "e \ 1" using e_def assms by auto + moreover have "e > 0" "e < 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed @@ -884,12 +862,12 @@ then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) - then obtain e1 where "0 < e1" "e1 \ 1" "z = y - e1 *\<^sub>R (y - x)" + then obtain e1 where "0 < e1" "e1 < 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ - by auto + by fastforce qed } ultimately show ?thesis by auto @@ -1256,11 +1234,7 @@ have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" - using x1_def x2_def - apply (auto simp add: algebra_simps) - using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] - apply auto - done + by (simp add: x1_def x2_def algebra_simps) (simp add: "*" pth_8) then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * @@ -1567,8 +1541,8 @@ and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms -apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) -by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) + unfolding rel_frontier_def frontier_def + using convex_affine_closure_Int convex_affine_rel_interior_Int rel_interior_nonempty_interior by fastforce lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" @@ -2113,13 +2087,10 @@ then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto - then have "x \ ?rhs" - apply auto - apply (rule_tac x = c in exI) - apply (rule_tac x = s in exI) - using * c_def s_def p \x \ S i\ - apply auto - done + moreover have "(\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. s i \ S i)" + using * c_def s_def p \x \ S i\ by auto + ultimately have "x \ ?rhs" + by force } then have "?rhs \ S i" by auto } @@ -2322,17 +2293,18 @@ proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" - apply (rule subsetD [OF rel_interior_subset inint]) - using \l \ 0\ \0 < d\ \0 < \\ by auto + proof (rule subsetD [OF rel_interior_subset inint]) + show "d - min d (\ / 2 / norm l) < d" + using \l \ 0\ \0 < d\ \0 < \\ by auto + qed auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" - apply (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) - using 1 2 \0 < d\ \0 < \\ apply (auto simp: algebra_simps) - done + using 1 2 \0 < d\ \0 < \\ + by (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) (auto simp: algebra_simps) qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) @@ -2348,14 +2320,14 @@ obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - - have "interior S = rel_interior S" + have \
: "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp - then show thesis - apply (rule ray_to_rel_frontier [OF \bounded S\ _ _ \l \ 0\]) - using a affine_hull_nonempty_interior apply blast - by (simp add: \interior S = rel_interior S\ frontier_def rel_frontier_def that) + moreover have "a + l \ affine hull S" + using a affine_hull_nonempty_interior by blast + ultimately show thesis + by (metis \
\bounded S\ \l \ 0\ frontier_def ray_to_rel_frontier rel_frontier_def that) qed @@ -2384,10 +2356,8 @@ moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" - apply (simp add: in_segment) - apply (rule_tac x="1/d" in exI) - apply (auto simp: algebra_simps) - done + unfolding in_segment + by (rule_tac x="1/d" in exI) (auto simp: algebra_simps) next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" proof (rule rel_interior_closure_convex_segment [OF \convex S\ x]) @@ -2691,11 +2661,8 @@ by auto } then have "(1, x) \ rel_interior K0" - using K0_def * set_sum_alt[of I "(\i. rel_interior (K i))"] assms k_def cs - apply simp - apply (rule_tac x = k in exI) - apply (simp add: sum_prod) - done + using * set_sum_alt[of I "(\i. rel_interior (K i))"] assms cs + by (simp add: k_def) (metis (mono_tags, lifting) sum_prod) then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto @@ -2795,20 +2762,6 @@ shows "convex hull T = affine hull T \ convex hull S" proof - have fin: "finite S" "finite T" using assms aff_independent_finite finite_subset by auto - { fix u v x - assume uv: "sum u T = 1" "\x\S. 0 \ v x" "sum v S = 1" - "(\x\S. v x *\<^sub>R x) = (\v\T. u v *\<^sub>R v)" "x \ T" - then have S: "S = (S - T) \ T" \ \split into separate cases\ - using assms by auto - have [simp]: "(\x\T. v x *\<^sub>R x) + (\x\S - T. v x *\<^sub>R x) = (\x\T. u x *\<^sub>R x)" - "sum v T + sum v (S - T) = 1" - using uv fin S - by (auto simp: sum.union_disjoint [symmetric] Un_commute) - have "(\x\S. if x \ T then v x - u x else v x) = 0" - "(\x\S. (if x \ T then v x - u x else v x) *\<^sub>R x) = 0" - using uv fin - by (subst S, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ - } note [simp] = this have "convex hull T \ affine hull T" using convex_hull_subset_affine_hull by blast moreover have "convex hull T \ convex hull S" @@ -2818,11 +2771,31 @@ have 0: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using affine_dependent_explicit_finite assms(1) fin(1) by auto show ?thesis - apply (clarsimp simp add: convex_hull_finite affine_hull_finite fin) - apply (rule_tac x=u in exI) - subgoal for u v - using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ by force - done + proof (clarsimp simp add: affine_hull_finite fin) + fix u + assume S: "(\v\T. u v *\<^sub>R v) \ convex hull S" + and T1: "sum u T = 1" + then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\v\T. u v *\<^sub>R v)" + by (auto simp add: convex_hull_finite fin) + { fix x + assume"x \ T" + then have S: "S = (S - T) \ T" \ \split into separate cases\ + using assms by auto + have [simp]: "(\x\T. v x *\<^sub>R x) + (\x\S - T. v x *\<^sub>R x) = (\x\T. u x *\<^sub>R x)" + "sum v T + sum v (S - T) = 1" + using v fin S + by (auto simp: sum.union_disjoint [symmetric] Un_commute) + have "(\x\S. if x \ T then v x - u x else v x) = 0" + "(\x\S. (if x \ T then v x - u x else v x) *\<^sub>R x) = 0" + using v fin T1 + by (subst S, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ + } note [simp] = this + have "(\x\T. 0 \ u x)" + using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ v(1) by fastforce + then show "(\v\T. u v *\<^sub>R v) \ convex hull T" + using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ T1 + by (fastforce simp add: convex_hull_finite fin) + qed qed ultimately show ?thesis by blast @@ -2841,11 +2814,13 @@ by blast then have fin: "finite T" using assms by (metis finite_insert aff_independent_finite) - then have "UNIV \ (+) a ` span ((\x. x - a) ` T)" - using assms T - apply (intro card_ge_dim_independent Fun.vimage_subsetD) - apply (auto simp: surj_def affine_dependent_iff_dependent card_image inj_on_def dim_subset_UNIV) - done + have "UNIV \ (+) a ` span ((\x. x - a) ` T)" + proof (intro card_ge_dim_independent Fun.vimage_subsetD) + show "independent ((\x. x - a) ` T)" + using T affine_dependent_iff_dependent assms(1) by auto + show "dim ((+) a -` UNIV) \ card ((\x. x - a) ` T)" + using assms T fin by (auto simp: card_image inj_on_def) + qed (use surj_plus in auto) then show ?thesis using T(2) affine_hull_insert_span_gen equalityI by fastforce qed @@ -2896,66 +2871,66 @@ by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma rel_interior_convex_hull_explicit: - fixes s :: "'a::euclidean_space set" - assumes "\ affine_dependent s" - shows "rel_interior(convex hull s) = - {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent S" + shows "rel_interior(convex hull S) = + {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" - proof (cases "\a. s = {a}") + proof (cases "\a. S = {a}") case True then show "?lhs \ ?rhs" by force next case False - have fs: "finite s" + have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix a b and d::real - assume ab: "a \ s" "b \ s" "a \ b" - then have s: "s = (s - {a,b}) \ {a,b}" \ \split into separate cases\ + assume ab: "a \ S" "b \ S" "a \ b" + then have S: "S = (S - {a,b}) \ {a,b}" \ \split into separate cases\ by auto - have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" - "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" + have "(\x\S. if x = a then - d else if x = b then d else 0) = 0" + "(\x\S. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs - by (subst s, subst sum.union_disjoint, auto)+ + by (subst S, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y - assume y: "y \ convex hull s" "y \ ?rhs" - { fix u T a - assume ua: "\x\s. 0 \ u x" "sum u s = 1" "\ 0 < u a" "a \ s" - and yT: "y = (\x\s. u x *\<^sub>R x)" "y \ T" "open T" - and sb: "T \ affine hull s \ {w. \u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = w}" + assume y: "y \ convex hull S" "y \ ?rhs" + have *: False if + ua: "\x\S. 0 \ u x" "sum u S = 1" "\ 0 < u a" "a \ S" + and yT: "y = (\x\S. u x *\<^sub>R x)" "y \ T" "open T" + and sb: "T \ affine hull S \ {w. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = w}" + for u T a + proof - have ua0: "u a = 0" using ua by auto - obtain b where b: "b\s" "a \ b" + obtain b where b: "b\S" "a \ b" using ua False by auto - obtain e where e: "0 < e" "ball (\x\s. u x *\<^sub>R x) e \ T" + obtain e where e: "0 < e" "ball (\x\S. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) - have "(\x\s. u x *\<^sub>R x) \ affine hull s" + have "(\x\S. u x *\<^sub>R x) \ affine hull S" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) - then have "(\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull s" + then have "(\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull S" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) - then have "y - d *\<^sub>R (a - b) \ T \ affine hull s" + then have "y - d *\<^sub>R (a - b) \ T \ affine hull S" using d e yT by auto - then obtain v where "\x\s. 0 \ v x" - "sum v s = 1" - "(\x\s. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" + then obtain v where v: "\x\S. 0 \ v x" + "sum v S = 1" + "(\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto - then have False - using assms - apply (simp add: affine_dependent_explicit_finite fs) - apply (drule_tac x="\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) - using ua b d - apply (auto simp: algebra_simps sum_subtractf sum.distrib) - done - } note * = this - have "y \ rel_interior (convex hull s)" + have aff: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" + using assms by (simp add: affine_dependent_explicit_finite fs) + show False + using ua b d v aff [of "\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)"] + by (auto simp: algebra_simps sum_subtractf sum.distrib) + qed + have "y \ rel_interior (convex hull S)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) @@ -2968,42 +2943,46 @@ qed lemma interior_convex_hull_explicit_minimal: - fixes s :: "'a::euclidean_space set" + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent S" shows - "\ affine_dependent s - ==> interior(convex hull s) = - (if card(s) \ DIM('a) then {} - else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" - apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) - apply (rule trans [of _ "rel_interior(convex hull s)"]) - apply (simp add: affine_independent_span_gt rel_interior_interior) - by (simp add: rel_interior_convex_hull_explicit) + "interior(convex hull S) = + (if card(S) \ DIM('a) then {} + else {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" + (is "_ = (if _ then _ else ?rhs)") +proof (clarsimp simp: aff_independent_finite empty_interior_convex_hull assms) + assume S: "\ card S \ DIM('a)" + have "interior (convex hull S) = rel_interior(convex hull S)" + using assms S by (simp add: affine_independent_span_gt rel_interior_interior) + then show "interior(convex hull S) = ?rhs" + by (simp add: assms S rel_interior_convex_hull_explicit) +qed lemma interior_convex_hull_explicit: - fixes s :: "'a::euclidean_space set" - assumes "\ affine_dependent s" + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent S" shows - "interior(convex hull s) = - (if card(s) \ DIM('a) then {} - else {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" + "interior(convex hull S) = + (if card(S) \ DIM('a) then {} + else {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a - assume "card Basis < card s" and u: "\x. x\s \ 0 < u x" "sum u s = 1" and a: "a \ s" - then have cs: "Suc 0 < card s" + assume "card Basis < card S" and u: "\x. x\S \ 0 < u x" "sum u S = 1" and a: "a \ S" + then have cs: "Suc 0 < card S" by (metis DIM_positive less_trans_Suc) - obtain b where b: "b \ s" "a \ b" - proof (cases "s \ {a}") + obtain b where b: "b \ S" "a \ b" + proof (cases "S \ {a}") case True then show thesis using cs subset_singletonD by fastforce qed blast have "u a + u b \ sum u {a,b}" using a b by simp - also have "... \ sum u s" + also have "... \ sum u S" using a b u by (intro Groups_Big.sum_mono2) (auto simp: less_imp_le aff_independent_finite assms) finally have "u a < 1" - using \b \ s\ u by fastforce + using \b \ S\ u by fastforce } note [simp] = this show ?thesis using assms by (force simp add: not_le interior_convex_hull_explicit_minimal) @@ -3139,11 +3118,9 @@ proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) - have "\u xa v. - \ xa \ S; - u xa = 0; sum u S = 1; \x\S. 0 < v x; - sum v S = 1; - (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)\ + have "\u y v. + \y \ S; u y = 0; sum u S = 1; \x\S. 0 < v x; + sum v S = 1; (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)\ \ \u. sum u S = 0 \ (\v\S. u v \ 0) \ (\v\S. u v *\<^sub>R v) = 0" apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) @@ -3152,8 +3129,8 @@ using fs assms apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit) apply (auto simp: convex_hull_finite) - apply (metis less_eq_real_def) - using affine_dependent_explicit_finite assms by blast + apply (metis less_eq_real_def) + by (simp add: affine_dependent_explicit_finite) qed lemma frontier_convex_hull_explicit: @@ -3210,7 +3187,7 @@ ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) - apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) + apply (auto simp add: convex_hull_finite fs Union_SetCompr_eq) done qed @@ -3237,9 +3214,15 @@ shows "x \ frontier(convex hull S)" proof (cases "affine_dependent S") case True - with assms show ?thesis - apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) - by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) + with assms obtain y where "y \ S" and y: "y \ affine hull (S - {y})" + by (auto simp: affine_dependent_def) + moreover have "x \ closure (convex hull S)" + by (meson closure_subset hull_inc subset_eq \x \ S\) + moreover have "x \ interior (convex hull S)" + using assms + by (metis Suc_mono affine_hull_convex_hull affine_hull_nonempty_interior \y \ S\ y card.remove empty_iff empty_interior_affine_hull finite_Diff hull_redundant insert_Diff interior_UNIV not_less) + ultimately show ?thesis + unfolding frontier_def by blast next case False { assume "card S = Suc (card Basis)" @@ -3483,62 +3466,51 @@ lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "coplanar s" "linear f" shows "coplanar(f ` s)" + assumes "coplanar S" "linear f" shows "coplanar(f ` S)" proof - { fix u v w - assume "s \ affine hull {u, v, w}" - then have "f ` s \ f ` (affine hull {u, v, w})" + assume "S \ affine hull {u, v, w}" + then have "f ` S \ f ` (affine hull {u, v, w})" by (simp add: image_mono) - then have "f ` s \ affine hull (f ` {u, v, w})" + then have "f ` S \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed -lemma coplanar_translation_imp: "coplanar s \ coplanar ((\x. a + x) ` s)" - unfolding coplanar_def - apply clarify - apply (rule_tac x="u+a" in exI) - apply (rule_tac x="v+a" in exI) - apply (rule_tac x="w+a" in exI) - using affine_hull_translation [of a "{u,v,w}" for u v w] - apply (force simp: add.commute) - done - -lemma coplanar_translation_eq: "coplanar((\x. a + x) ` s) \ coplanar s" +lemma coplanar_translation_imp: + assumes "coplanar S" shows "coplanar ((\x. a + x) ` S)" +proof - + obtain u v w where "S \ affine hull {u,v,w}" + by (meson assms coplanar_def) + then have "(+) a ` S \ affine hull {u + a, v + a, w + a}" + using affine_hull_translation [of a "{u,v,w}" for u v w] + by (force simp: add.commute) + then show ?thesis + unfolding coplanar_def by blast +qed + +lemma coplanar_translation_eq: "coplanar((\x. a + x) ` S) \ coplanar S" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" + assumes "linear f" "inj f" shows "coplanar(f ` S) = coplanar S" proof - assume "coplanar s" - then show "coplanar (f ` s)" - unfolding coplanar_def - using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms - by (meson coplanar_def coplanar_linear_image) + assume "coplanar S" + then show "coplanar (f ` S)" + using assms(1) coplanar_linear_image by blast next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast - assume "coplanar (f ` s)" - then obtain u v w where "f ` s \ affine hull {u, v, w}" - by (auto simp: coplanar_def) - then have "g ` f ` s \ g ` (affine hull {u, v, w})" - by blast - then have "s \ g ` (affine hull {u, v, w})" - using g by (simp add: Fun.image_comp) - then show "coplanar s" - unfolding coplanar_def - using affine_hull_linear_image [of g "{u,v,w}" for u v w] \linear g\ linear_conv_bounded_linear - by fastforce -qed -(*The HOL Light proof is simply - MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; -*) - -lemma coplanar_subset: "\coplanar t; s \ t\ \ coplanar s" + assume "coplanar (f ` S)" + then show "coplanar S" + by (metis coplanar_linear_image g(1) g(2) id_apply image_comp image_id) +qed + +lemma coplanar_subset: "\coplanar t; S \ t\ \ coplanar S" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" @@ -3549,11 +3521,10 @@ proof - obtain u x y where "b - a = y *\<^sub>R u" "c - a = x *\<^sub>R u" using assms unfolding collinear_def by auto - with \a \ b\ show ?thesis - apply (simp add: affine_hull_2) - apply (rule_tac x="1 - x/y" in exI) - apply (simp add: algebra_simps) - done + with \a \ b\ have "\v. c = (1 - x / y) *\<^sub>R a + v *\<^sub>R b \ 1 - x / y + v = 1" + by (simp add: algebra_simps) + then show ?thesis + by (simp add: hull_inc mem_affine) qed lemma collinear_3_affine_hull: @@ -3569,10 +3540,15 @@ by (auto simp: insert_commute) next case False - then show ?thesis - apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) - apply (metis collinear_3_affine_hull insert_commute)+ - done + then have "collinear{a,b,c}" if "affine_dependent {a,b,c}" + using that unfolding affine_dependent_def + by (auto simp: insert_Diff_if; metis affine_hull_3_imp_collinear insert_commute) + moreover + have "affine_dependent {a,b,c}" if "collinear{a,b,c}" + using False that by (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) + ultimately + show ?thesis + using False by blast qed lemma affine_dependent_imp_collinear_3: @@ -3617,11 +3593,13 @@ by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed -lemma collinear_midpoint: "collinear{a,midpoint a b,b}" - apply (auto simp: collinear_3 collinear_lemma) - apply (drule_tac x="-1" in spec) - apply (simp add: algebra_simps) - done +lemma collinear_midpoint: "collinear{a, midpoint a b, b}" +proof - + have \
: "\a \ midpoint a b; b - midpoint a b \ - 1 *\<^sub>R (a - midpoint a b)\ \ b = midpoint a b" + by (simp add: algebra_simps) + show ?thesis + by (auto simp: collinear_3 collinear_lemma intro: \
) +qed lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" @@ -3632,13 +3610,21 @@ "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) - have "b = midpoint a c \ collinear{a,b,c} " + have "b = midpoint a c \ collinear{a,b,c}" using collinear_midpoint by blast - moreover have "collinear{a,b,c} \ b = midpoint a c \ dist a b = dist b c" - apply (auto simp: collinear_3_expand assms dist_midpoint) - apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) - apply (simp add: algebra_simps) - done + moreover have "b = midpoint a c \ dist a b = dist b c" if "collinear{a,b,c}" + proof - + consider "a = c" | u where "b = u *\<^sub>R a + (1 - u) *\<^sub>R c" + using \collinear {a,b,c}\ unfolding collinear_3_expand by blast + then show ?thesis + proof cases + case 2 + with assms have "dist a b = dist b c \ b = midpoint a c" + by (simp add: dist_norm * midpoint_def scaleR_add_right del: divide_const_simps) + then show ?thesis + by (auto simp: dist_midpoint) + qed (use assms in auto) + qed ultimately show ?thesis by blast qed @@ -3650,24 +3636,22 @@ case True with assms show ?thesis by (auto simp: dist_commute) next - case False with assms show ?thesis - apply (auto simp: collinear_3 collinear_lemma between_norm) - apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) - apply (simp add: vector_add_divide_simps real_vector.scale_minus_right [symmetric]) - done + case False + then have False if "\c. b - x \ c *\<^sub>R (a - x)" + using that [of "-(norm(b - x) / norm(x - a))"] assms + by (simp add: between_norm vector_add_divide_simps flip: real_vector.scale_minus_right) + then show ?thesis + by (auto simp: collinear_3 collinear_lemma) qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") - case True then show ?thesis - by (auto simp: dist_commute) -next case False show ?thesis using False between_imp_collinear between_midpoint(1) midpoint_collinear by blast -qed +qed (auto simp: dist_commute) lemma collinear_triples: assumes "a \ b" @@ -3682,13 +3666,19 @@ assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) - then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ (insert a (insert b S))" for x + then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ insert a (insert b S)" for x using that assms collinear_3_expand by fastforce+ - show ?lhs - unfolding collinear_def - apply (rule_tac x="b-a" in exI) - apply (clarify dest!: *) - by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff) + have "\c. x - y = c *\<^sub>R (b - a)" + if x: "x \ insert a (insert b S)" and y: "y \ insert a (insert b S)" for x y + proof - + obtain u v where "x = u *\<^sub>R a + (1 - u) *\<^sub>R b" "y = v *\<^sub>R a + (1 - v) *\<^sub>R b" + using "*" x y by presburger + then have "x - y = (v - u) *\<^sub>R (b - a)" + by (simp add: scale_left_diff_distrib scale_right_diff_distrib) + then show ?thesis .. + qed + then show ?lhs + unfolding collinear_def by metis qed lemma collinear_4_3: @@ -3709,28 +3699,34 @@ lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" -apply (simp add: affine_hull_2, safe) -apply (rule_tac x=v in image_eqI) -apply (simp add: algebra_simps) -apply (metis scaleR_add_left scaleR_one, simp) -apply (rule_tac x="1-u" in exI) -apply (simp add: algebra_simps) -done +proof - + have 1: "u *\<^sub>R a + v *\<^sub>R b = a + v *\<^sub>R (b - a)" if "u + v = 1" for u v + using that + by (simp add: algebra_simps flip: scaleR_add_left) + have 2: "a + u *\<^sub>R (b - a) = (1 - u) *\<^sub>R a + u *\<^sub>R b" for u + by (auto simp: algebra_simps) + show ?thesis + by (force simp add: affine_hull_2 dest: 1 intro!: 2) +qed lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" - shows "\\ collinear{a,b,c}; DIM('a) = 2\ - \ interior(convex hull {a,b,c}) = - {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ - x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" -apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) -apply (rule_tac x="u a" in exI, simp) -apply (rule_tac x="u b" in exI, simp) -apply (rule_tac x="u c" in exI, simp) -apply (rename_tac uu x y z) -apply (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) -apply simp -done + assumes "\ collinear{a,b,c}" and 2: "DIM('a) = 2" + shows "interior(convex hull {a,b,c}) = + {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" + (is "?lhs = ?rhs") +proof + have abc: "a \ b" "a \ c" "b \ c" "\ affine_dependent {a, b, c}" + using assms by (auto simp: collinear_3_eq_affine_dependent) + with 2 show "?lhs \ ?rhs" + by (fastforce simp add: interior_convex_hull_explicit_minimal) + show "?rhs \ ?lhs" + using abc 2 + apply (clarsimp simp add: interior_convex_hull_explicit_minimal) + subgoal for x y z + by (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) auto + done +qed subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ @@ -3992,18 +3988,20 @@ assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" - apply (rule connected_chain) - using S apply blast +proof (rule connected_chain) + show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis image_iff le_cases nest) +qed (use S in blast) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" - apply (rule connected_chain_gen [of "S k"]) - using S apply auto - by (meson le_cases nest subsetCE) +proof (rule connected_chain_gen [of "S k"]) + show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" + by (metis imageE le_cases nest) +qed (use S in auto) subsection\Proper maps, including projections out of compact sets\ @@ -4040,6 +4038,7 @@ by metis then have fX: "\n. f (X n) = h n" by metis + define \ where "\ \ \n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))}" have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n proof (intro closed_Int_compact [OF \closed C\ com] compact_sequence_with_limit) show "insert y (range (\i. f (X (n + i)))) \ T" @@ -4047,31 +4046,28 @@ show "(\i. f (X (n + i))) \ y" by (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) qed - then have comf: "compact {a \ K. f a \ insert y (range (\i. f(X(n + i))))}" for n - by (simp add: Keq Int_def conj_commute) + then have comf: "compact (\ n)" for n + by (simp add: Keq Int_def \_def conj_commute) have ne: "\\ \ {}" if "finite \" - and \: "\t. t \ \ \ - (\n. t = {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" + and \: "\t. t \ \ \ (\n. t = \ n)" for \ proof - - obtain m where m: "\t. t \ \ \ \k\m. t = {a \ K. f a \ insert y (range (\i. f (X (k + i))))}" + obtain m where m: "\t. t \ \ \ \k\m. t = \ k" by (rule exE [OF finite_indexed_bound [OF \finite \\ \]], force+) have "X m \ \\" - using X le_Suc_ex by (fastforce dest: m) + using X le_Suc_ex by (fastforce simp: \_def dest: m) then show ?thesis by blast qed - have "\{{a. a \ K \ f a \ insert y (range (\i. f(X(n + i))))} |n. n \ UNIV} - \ {}" - apply (rule compact_fip_Heine_Borel) - using comf apply force - using ne apply (simp add: subset_iff del: insert_iff) - done - then have "\x. x \ (\n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" - by blast + have "(\n. \ n) \ {}" + proof (rule compact_fip_Heine_Borel) + show "\\'. \finite \'; \' \ range \\ \ \ \' \ {}" + by (meson ne rangeE subset_eq) + qed (use comf in blast) + then obtain x where "x \ K" "\n. (f x = y \ (\u. f x = h (n + u)))" + by (force simp add: \_def fX) then show ?thesis - apply (simp add: image_iff fX) - by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) + unfolding image_iff by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) @@ -4144,8 +4140,10 @@ by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" - apply (clarsimp simp: dist_norm continuous_on_iff diff) - by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) + proof (clarsimp simp: dist_norm continuous_on_iff diff) + show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" + by (metis \z \ 0\ mult_pos_pos real_mult_less_iff1 zero_less_norm_iff) + qed then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis @@ -4169,11 +4167,11 @@ using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - - have "{.. f ` {x,y} \ {}" "{f z<..} \ f ` {x,y} \ {}" - using fz fz_notin \x \ S\ \y \ S\ - apply (auto simp: closed_segment_eq_real_ivl less_eq_real_def split: if_split_asm) - apply (metis image_eqI)+ - done + consider "f x \ f z \ f z \ f y" | "f y \ f z \ f z \ f x" + using fz + by (auto simp add: closed_segment_eq_real_ivl split: if_split_asm) + then have "{.. f ` {x,y} \ {} \ {f z<..} \ f ` {x,y} \ {}" + by cases (use fz_notin \x \ S\ \y \ S\ in \auto simp: image_iff\) then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed @@ -4212,8 +4210,10 @@ by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" - apply (clarsimp simp: dist_norm continuous_on_iff diff) - by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) + proof (clarsimp simp: dist_norm continuous_on_iff diff) + show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" + by (metis \z \ 0\ mult_pos_pos real_mult_less_iff1 zero_less_norm_iff) + qed then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" @@ -4236,9 +4236,8 @@ then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" - apply safe - apply (metis (mono_tags, hide_lams) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_inc subsetCE) - by (metis (mono_tags, lifting) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE) + unfolding image_def using \a \ S\ \b \ S\ + by (safe; metis (mono_tags, lifting) convex_affine_hull convex_contains_segment gf hull_subset subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis @@ -4500,6 +4499,7 @@ proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" + (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) @@ -4508,12 +4508,13 @@ then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") - case True show ?thesis - using span_zero [of S] - apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane - del: One_nat_def) - apply (auto simp add: \c = 0\) - done + case True + have "?lhs \ (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" + by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) + also have "... \ ?rhs" + using span_zero [of S] True \c \ S\ affine_hull_span_0 hull_inc + by (fastforce simp add: affine_hull_span_gen [of c] \c = 0\) + finally show ?thesis . next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x @@ -4537,12 +4538,11 @@ intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed - show ?thesis - apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane - del: One_nat_def cong: image_cong_simp, safe) - apply (fastforce simp add: inner_distrib intro: xc_im) - apply (force intro!: 2) - done + have "?lhs = (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" + by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) + also have "... = ?rhs" + by (fastforce simp add: affine_hull_span_gen [of c] \c \ S\ hull_inc inner_distrib intro: xc_im intro!: 2) + finally show ?thesis . qed qed @@ -4859,39 +4859,34 @@ subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma - fixes s :: "'a::euclidean_space set" - assumes "\ affine_dependent(s \ t)" - shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) - and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent(S \ T)" + shows convex_hull_Int_subset: "convex hull S \ convex hull T \ convex hull (S \ T)" (is ?C) + and affine_hull_Int_subset: "affine hull S \ affine hull T \ affine hull (S \ T)" (is ?A) proof - - have [simp]: "finite s" "finite t" + have [simp]: "finite S" "finite T" using aff_independent_finite assms by blast+ - have "sum u (s \ t) = 1 \ - (\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" - if [simp]: "sum u s = 1" - "sum v t = 1" - and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v + have "sum u (S \ T) = 1 \ + (\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" + if [simp]: "sum u S = 1" + "sum v T = 1" + and eq: "(\x\T. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" for u v proof - - define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x - have "sum f (s \ t) = 0" - apply (simp add: f_def sum_Un sum_subtractf) - apply (simp add: Int_commute flip: sum.inter_restrict) - done - moreover have "(\x\(s \ t). f x *\<^sub>R x) = 0" - apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) - apply (simp add: if_smult Int_commute eq flip: sum.inter_restrict - cong del: if_weak_cong) - done - ultimately have "\v. v \ s \ t \ f v = 0" - using aff_independent_finite assms unfolding affine_dependent_explicit - by blast - then have u [simp]: "\x. x \ s \ u x = (if x \ t then v x else 0)" - by (simp add: f_def) presburger - have "sum u (s \ t) = sum u s" - by (simp add: sum.inter_restrict) - then have "sum u (s \ t) = 1" - using that by linarith - moreover have "(\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" + define f where "f x = (if x \ S then u x else 0) - (if x \ T then v x else 0)" for x + have "sum f (S \ T) = 0" + by (simp add: f_def sum_Un sum_subtractf flip: sum.inter_restrict) + moreover have "(\x\(S \ T). f x *\<^sub>R x) = 0" + by (simp add: eq f_def sum_Un scaleR_left_diff_distrib sum_subtractf if_smult flip: sum.inter_restrict cong: if_cong) + ultimately have "\v. v \ S \ T \ f v = 0" + using aff_independent_finite assms unfolding affine_dependent_explicit + by blast + then have u [simp]: "\x. x \ S \ u x = (if x \ T then v x else 0)" + by (simp add: f_def) presburger + have "sum u (S \ T) = sum u S" + by (simp add: sum.inter_restrict) + then have "sum u (S \ T) = 1" + using that by linarith + moreover have "(\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force @@ -4902,36 +4897,36 @@ proposition\<^marker>\tag unimportant\ affine_hull_Int: - fixes s :: "'a::euclidean_space set" - assumes "\ affine_dependent(s \ t)" - shows "affine hull (s \ t) = affine hull s \ affine hull t" + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent(S \ T)" + shows "affine hull (S \ T) = affine hull S \ affine hull T" by (simp add: affine_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ convex_hull_Int: - fixes s :: "'a::euclidean_space set" - assumes "\ affine_dependent(s \ t)" - shows "convex hull (s \ t) = convex hull s \ convex hull t" + fixes S :: "'a::euclidean_space set" + assumes "\ affine_dependent(S \ T)" + shows "convex hull (S \ T) = convex hull S \ convex hull T" by (simp add: convex_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ - fixes s :: "'a::euclidean_space set set" - assumes "\ affine_dependent (\s)" - shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") - and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") + fixes S :: "'a::euclidean_space set set" + assumes "\ affine_dependent (\S)" + shows affine_hull_Inter: "affine hull (\S) = (\T\S. affine hull T)" (is "?A") + and convex_hull_Inter: "convex hull (\S) = (\T\S. convex hull T)" (is "?C") proof - - have "finite s" + have "finite S" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms - proof (induction s rule: finite_induct) + proof (induction S rule: finite_induct) case empty then show ?case by auto next - case (insert t F) + case (insert T F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False - with "insert.prems" have [simp]: "\ affine_dependent (t \ \F)" + with "insert.prems" have [simp]: "\ affine_dependent (T \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce @@ -5182,18 +5177,19 @@ show ?thesis proof have "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" - apply (rule card_image [OF inj_onI]) - by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) + proof (rule card_image) + show "inj_on (\a. affine hull (c - {a})) (c - b)" + unfolding inj_on_def + by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) + qed then show "int (card ?\) + aff_dim S = int DIM('a)" by (simp add: imeq card_cb) show "affine hull S = \ ?\" - using \b \ c\ False - apply (subst affine_hull_Inter [OF ind, symmetric]) - apply (simp add: eq double_diff) - done - show "\h. h \ ?\ \ \a b. a \ 0 \ h = {x. a \ x = b}" - apply clarsimp - by (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) + by (metis Diff_eq_empty_iff INT_simps(4) UN_singleton order_refl \b \ c\ False eq double_diff affine_hull_Inter [OF ind]) + have "\a. \a \ c; a \ b\ \ aff_dim (c - {a}) = int (DIM('a) - Suc 0)" + by (metis "*" DIM_ge_Suc0 One_nat_def add_diff_cancel_left' int_ops(2) of_nat_diff) + then show "\h. h \ ?\ \ \a b. a \ 0 \ h = {x. a \ x = b}" + by (auto simp only: One_nat_def aff_dim_eq_hyperplane [symmetric]) qed (use \finite c\ in auto) qed qed @@ -5255,22 +5251,27 @@ using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) - let ?U = "(+) (c+c) ` {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" - have "u + v \ (+) (c + c) ` {x + v |x v. x \ (+) (- c) ` S \ a \ v = 0}" - if "u \ S" "b = a \ v" for u v - apply (rule_tac x="u+v-c-c" in image_eqI) - apply (simp_all add: algebra_simps) - apply (rule_tac x="u-c" in exI) - apply (rule_tac x="v-c" in exI) - apply (simp add: algebra_simps that c) - done + define U where "U \ {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" + have "u + v \ (+) (c+c) ` U" + if "u \ S" "b = a \ v" for u v + proof + show "u + v = c + c + (u + v - c - c)" + by (simp add: algebra_simps) + have "\x y. u + v - c - c = x + y \ (\xa\S. x = xa - c) \ a \ y = 0" + proof (intro exI conjI) + show "u + v - c - c = (u-c) + (v-c)" "a \ (v - c) = 0" + by (simp_all add: algebra_simps c that) + qed (use that in auto) + then show "u + v - c - c \ U" + by (auto simp: U_def image_def) + qed moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) - ultimately have "{x + y |x y. x \ S \ a \ y = b} = ?U" - by (fastforce simp: algebra_simps) + ultimately have "{x + y |x y. x \ S \ a \ y = b} = (+) (c+c) ` U" + by (fastforce simp: algebra_simps U_def) also have "... = range ((+) (c + c))" - by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) + by (simp only: U_def affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . @@ -5568,10 +5569,6 @@ by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) - have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" - by (metis subspace_add a' span_eq_iff sub) - then have Sclo: "\w. w \ S \ (w + a') \ S" - by fastforce show ?thesis proof (cases "a' = 0") case True @@ -5580,17 +5577,23 @@ next case False show ?thesis - apply (rule_tac a' = "a'" and b' = "a' \ z" in that) - apply (auto simp: a ba'' inner_left_distrib False Sclo) - done + proof + show "S \ {x. a' \ x \ a' \ z} = S \ {x. a \ x \ b}" + "S \ {x. a' \ x = a' \ z} = S \ {x. a \ x = b}" + by (auto simp: a ba'' inner_left_distrib) + have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" + by (metis subspace_add a' span_eq_iff sub) + then show "\w. w \ S \ (w + a') \ S" + by fastforce + qed (use False in auto) qed qed lemma diffs_affine_hull_span: assumes "a \ S" - shows "{x - a |x. x \ affine hull S} = span {x - a |x. x \ S}" + shows "(\x. x - a) ` (affine hull S) = span ((\x. x - a) ` S)" proof - - have *: "((\x. x - a) ` (S - {a})) = {x. x + a \ S} - {0}" + have *: "((\x. x - a) ` (S - {a})) = ((\x. x - a) ` S) - {0}" by (auto simp: algebra_simps) show ?thesis by (auto simp add: algebra_simps affine_hull_span2 [OF assms] *) @@ -5599,7 +5602,7 @@ lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" - shows "aff_dim S = dim {x - a |x. x \ S}" + shows "aff_dim S = dim ((\x. x - a) ` S)" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" @@ -5612,22 +5615,21 @@ by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) - have *: "{x - c |x. x \ S} = {x - a |x. x \ S}" - apply safe - apply (simp_all only: xy) - using mem_affine_3_minus [OF \affine S\] \a \ S\ \c \ S\ apply blast+ - done + have *: "(\x. x - c) ` S = (\x. x - a) ` S" + using assms \c \ S\ + by (auto simp: image_iff xy; metis mem_affine_3_minus pth_1) have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp - also have "... = dim {x - c |x. x \ B}" + also have "... = dim ((\x. x - c) ` B)" + using affine_independent_card_dim_diffs [OF ind \c \ B\] by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) - also have "... = dim {x - c | x. x \ affine hull B}" - by (simp add: diffs_affine_hull_span \c \ B\) - also have "... = dim {x - a |x. x \ S}" - by (simp add: affS aff *) - finally show ?thesis . + also have "... = dim ((\x. x - c) ` (affine hull B))" + by (simp add: diffs_affine_hull_span \c \ B\) + also have "... = dim ((\x. x - a) ` S)" + by (simp add: affS aff *) + finally show ?thesis . qed lemma aff_dim_linear_image_le: @@ -5642,13 +5644,13 @@ then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto - have 2: "{x - f a| x. x \ f ` T} = f ` {x - a| x. x \ T}" + have 2: "{x - f a| x. x \ f ` T} = f ` ((\x. x - a) ` T)" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) - also have "... = int (dim (f ` {x - a| x. x \ T}))" + also have "... = int (dim (f ` ((\x. x - a) ` T)))" by (force simp: linear_diff [OF assms] 2) - also have "... \ int (dim {x - a| x. x \ T})" + also have "... \ int (dim ((\x. x - a) ` T))" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) @@ -5863,9 +5865,8 @@ have eq: "((\x. Pair x (f x)) ` S) = (S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis - apply (subst eq) - apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) - by auto + unfolding eq + by (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) auto qed lemma continuous_closed_graph_eq: @@ -6076,22 +6077,26 @@ lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" - shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" + shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" (is "?lhs = ?rhs") proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) - have *: "open_segment a c \ insert b (open_segment a b \ open_segment b c)" - if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ open_segment a c" + have *: "?rhs \ insert b (open_segment a b \ open_segment b c)" + if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ ?rhs" proof - - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (open_segment a c))" + have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (?rhs))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis - apply (rule antisym) - using Un_closed_segment [OF b] assms * - by (simp_all add: closed_segment_eq_open b subset_open_segment insert_commute) + proof + show "?lhs \ ?rhs" + by (simp add: assms b subset_open_segment) + show "?rhs \ ?lhs" + using Un_closed_segment [OF b] * + by (simp add: closed_segment_eq_open insert_commute) + qed qed subsection\Covering an open set by a countable chain of compact sets\ @@ -6152,9 +6157,7 @@ \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" - apply (intro UN_mono order_refl) - apply (simp add: cball_subset_ball_iff field_split_simps) - done + by (simp add: cball_subset_ball_iff field_split_simps UN_mono) finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed @@ -6309,9 +6312,13 @@ fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" - apply (auto simp: orthogonal_comp_def orthogonal_def) - apply (simp add: adjoint_works assms(1) inner_commute) - by (metis adjoint_works all_zero_iff assms(1) inner_commute) +proof - + have "\x. \\y. y \ f x = 0\ \ f x = 0" + using assms inner_commute all_zero_iff by metis + then show ?thesis + using assms + by (auto simp: orthogonal_comp_def orthogonal_def adjoint_works inner_commute) +qed subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ diff -r e1d04777d8b6 -r aeac6424d3b5 src/HOL/Set.thy --- a/src/HOL/Set.thy Sat Nov 07 23:20:43 2020 +0000 +++ b/src/HOL/Set.thy Tue Nov 10 23:21:04 2020 +0000 @@ -1364,6 +1364,12 @@ using assms by blast+ qed +lemma Un_Int_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" + by auto + +lemma Int_Un_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" + by auto + text \\<^medskip> Set complement\ lemma Compl_disjoint [simp]: "A \ - A = {}"