# HG changeset patch # User paulson # Date 1601314455 -3600 # Node ID 7bb074cceefed78e74708bb23c878c4cb3b75430 # Parent 564012e31db17aa7d7a56fac38706787b21b6b69 de-applying diff -r 564012e31db1 -r 7bb074cceefe src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Sat Sep 26 14:29:46 2020 +0200 +++ b/src/HOL/Analysis/Polytope.thy Mon Sep 28 18:34:15 2020 +0100 @@ -21,13 +21,9 @@ "((+) a ` T face_of (+) a ` S) \ T face_of S" proof - have *: "\a T S. T face_of S \ ((+) a ` T face_of (+) a ` S)" - apply (simp add: face_of_def Ball_def, clarify) - by (meson imageI open_segment_translation_eq) + by (simp add: face_of_def) show ?thesis - apply (rule iffI) - apply (force simp: image_comp o_def dest: * [where a = "-a"]) - apply (blast intro: *) - done + by (force simp: image_comp o_def dest: * [where a = "-a"] intro: *) qed lemma face_of_linear_image: @@ -83,33 +79,27 @@ using y by (auto simp: rel_interior_cball) have "y \ x" "y \ S" "y \ T" using face_of_imp_subset rel_interior_subset T that by blast+ - then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" + then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" using \x \ S\ \x \ T\ \T face_of S\ unfolding face_of_def - apply clarify - apply (drule_tac x=x in bspec, assumption) - apply (drule_tac x=y in bspec, assumption) - apply (subst (asm) open_segment_commute) - apply (force simp: open_segment_image_interval image_def) - done + by (meson greaterThanLessThan_iff in_segment(2)) have in01: "min (1/2) (e / norm (x - y)) \ {0<..<1}" using \y \ x\ \e > 0\ by simp - show ?thesis - apply (rule zne [OF in01]) - apply (rule e [THEN subsetD]) - apply (rule IntI) - using \y \ x\ \e > 0\ - apply (simp add: cball_def dist_norm algebra_simps) - apply (simp add: Real_Vector_Spaces.scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) - apply (rule mem_affine [OF affine_affine_hull _ x]) - using \y \ T\ apply (auto simp: hull_inc) - done + have \
: "norm (min (1/2) (e / norm (x - y)) *\<^sub>R y - min (1/2) (e / norm (x - y)) *\<^sub>R x) \ e" + using \e > 0\ + by (simp add: scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) + show False + apply (rule zne [OF in01 e [THEN subsetD]]) + using \y \ T\ + apply (simp add: hull_inc mem_affine x) + by (simp add: dist_norm algebra_simps \
) qed show ?thesis - apply (rule subset_antisym) - using assms apply (simp add: hull_subset face_of_imp_subset) - apply (cases "T={}", simp) - apply (force simp: rel_interior_eq_empty [symmetric] \convex T\ intro: *) - done + proof (rule subset_antisym) + show "T \ affine hull T \ S" + using assms by (simp add: hull_subset face_of_imp_subset) + show "affine hull T \ S \ T" + using "*" \convex T\ rel_interior_eq_empty by fastforce + qed qed lemma face_of_imp_closed: @@ -139,8 +129,7 @@ with b show "a \ x \ a \ u" by simp qed show ?thesis - apply (simp add: face_of_def assms) - using "*" open_segment_commute by blast + using "*" open_segment_commute by (fastforce simp add: face_of_def assms) qed lemma face_of_Int_supporting_hyperplane_ge_strong: @@ -200,24 +189,20 @@ apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) apply (rule_tac x="e / (e + norm (b - c))" in bexI) - using False nbc \0 < e\ - apply (auto simp: algebra_simps) - done + using False nbc \0 < e\ by (auto simp: algebra_simps) then have "d \ T \ c \ T" - apply (rule face_ofD [OF \T face_of S\]) - using \d \ u\ \c \ u\ \u \ S\ \b \ T\ apply auto - done + by (meson \b \ T\ \c \ u\ \d \ u\ assms face_ofD subset_iff) then show ?thesis .. qed qed lemma face_of_eq: fixes S :: "'a::real_normed_vector set" - assumes "T face_of S" "u face_of S" "(rel_interior T) \ (rel_interior u) \ {}" - shows "T = u" - apply (rule subset_antisym) - apply (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subsetCE subset_of_face_of) - by (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subset_iff subset_of_face_of) + assumes "T face_of S" "U face_of S" "(rel_interior T) \ (rel_interior U) \ {}" + shows "T = U" + using assms + unfolding disjoint_iff_not_equal + by (metis IntI empty_iff face_of_imp_subset mem_rel_interior_ball subset_antisym subset_of_face_of) lemma face_of_disjoint_rel_interior: fixes S :: "'a::real_normed_vector set" @@ -270,10 +255,11 @@ fixes S :: "'a::euclidean_space set" assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" - apply (rule subset_of_face_of [OF T \U \ S\]) - using face_of_imp_eq_affine_Int [OF \convex S\ T] - using rel_interior_subset [of U] dis - using \U \ S\ disjnt_def by fastforce +proof (rule subset_of_face_of [OF T \U \ S\]) + show "T \ rel_interior U \ {}" + using face_of_imp_eq_affine_Int [OF \convex S\ T] rel_interior_subset [of U] dis \U \ S\ disjnt_def + by fastforce +qed lemma affine_hull_face_of_disjoint_rel_interior: fixes S :: "'a::euclidean_space set" @@ -363,13 +349,14 @@ case False then have sumcf: "sum c T = 1 - k" by (simp add: S k_def sum_diff sumc1) + have ge0: "\x. x \ T \ 0 \ inverse (1 - k) * c x" + by (metis \T \ S\ cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg subsetD sum_nonneg sumcf) + have eq1: "(\x\T. inverse (1 - k) * c x) = 1" + by (metis False eq_iff_diff_eq_0 mult.commute right_inverse sum_distrib_left sumcf) have "(\i\T. c i *\<^sub>R i) /\<^sub>R (1 - k) \ convex hull T" apply (simp add: convex_hull_finite fin) apply (rule_tac x="\i. inverse (1-k) * c i" in exI) - apply auto - apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE) - apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf) - by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong) + by (metis (mono_tags, lifting) eq1 ge0 scaleR_scaleR scale_sum_right sum.cong) with \0 < k\ have "inverse(k) *\<^sub>R (w - sum (\i. c i *\<^sub>R i) T) \ affine hull T" by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD]) moreover have "inverse(k) *\<^sub>R (w - sum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" @@ -390,10 +377,14 @@ qed proposition face_of_convex_hull_insert: - "\finite S; a \ affine hull S; T face_of convex hull S\ \ T face_of convex hull insert a S" - apply (rule face_of_trans, blast) - apply (rule face_of_convex_hulls; force simp: insert_Diff_if) - done + assumes "finite S" "a \ affine hull S" and T: "T face_of convex hull S" + shows "T face_of convex hull insert a S" +proof - + have "convex hull S face_of convex hull insert a S" + by (simp add: assms face_of_convex_hulls insert_Diff_if subset_insertI) + then show ?thesis + using T face_of_trans by blast +qed proposition face_of_affine_trivial: assumes "affine S" "T face_of S" @@ -411,11 +402,12 @@ case True with \a \ T\ show ?thesis by auto next case False - then have "a \ open_segment (2 *\<^sub>R a - b) b" - apply (auto simp: open_segment_def closed_segment_def) + then have "a \ 2 *\<^sub>R a - b" + by (simp add: scaleR_2) + with False have "a \ open_segment (2 *\<^sub>R a - b) b" + apply (clarsimp simp: open_segment_def closed_segment_def) apply (rule_tac x="1/2" in exI) - apply (simp add: algebra_simps) - by (simp add: scaleR_2) + by (simp add: algebra_simps) moreover have "2 *\<^sub>R a - b \ S" by (rule mem_affine [OF \affine S\ \a \ S\ \b \ S\, of 2 "-1", simplified]) moreover note \b \ S\ \a \ T\ @@ -499,10 +491,12 @@ lemma faces_of_translation: "{F. F face_of image (\x. a + x) S} = image (image (\x. a + x)) {F. F face_of S}" -apply (rule subset_antisym, clarify) -apply (auto simp: image_iff) -apply (metis face_of_imp_subset face_of_translation_eq subset_imageE) -done +proof - + have "\F. F face_of (+) a ` S \ \G. G face_of S \ F = (+) a ` G" + by (metis face_of_imp_subset face_of_translation_eq subset_imageE) + then show ?thesis + by (auto simp: image_iff) +qed proposition face_of_Times: assumes "F face_of S" and "F' face_of S'" @@ -531,49 +525,51 @@ corollary face_of_Times_decomp: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - shows "c face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ c = F \ F')" + shows "C face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ C = F \ F')" (is "?lhs = ?rhs") proof - assume c: ?lhs + assume C: ?lhs show ?rhs - proof (cases "c = {}") + proof (cases "C = {}") case True then show ?thesis by auto next case False - have 1: "fst ` c \ S" "snd ` c \ S'" - using c face_of_imp_subset by fastforce+ - have "convex c" - using c by (metis face_of_imp_convex) - have conv: "convex (fst ` c)" "convex (snd ` c)" - by (simp_all add: \convex c\ convex_linear_image linear_fst linear_snd) - have fstab: "a \ fst ` c \ b \ fst ` c" - if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ c" for a b x x' + have 1: "fst ` C \ S" "snd ` C \ S'" + using C face_of_imp_subset by fastforce+ + have "convex C" + using C by (metis face_of_imp_convex) + have conv: "convex (fst ` C)" "convex (snd ` C)" + by (simp_all add: \convex C\ convex_linear_image linear_fst linear_snd) + have fstab: "a \ fst ` C \ b \ fst ` C" + if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ C" for a b x x' proof - have *: "(x,x') \ open_segment (a,x') (b,x')" using that by (auto simp: in_segment) show ?thesis - using face_ofD [OF c *] that face_of_imp_subset [OF c] by force + using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed - have fst: "fst ` c face_of S" + have fst: "fst ` C face_of S" by (force simp: face_of_def 1 conv fstab) - have sndab: "a' \ snd ` c \ b' \ snd ` c" - if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ c" for a' b' x x' + have sndab: "a' \ snd ` C \ b' \ snd ` C" + if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ C" for a' b' x x' proof - have *: "(x,x') \ open_segment (x,a') (x,b')" using that by (auto simp: in_segment) show ?thesis - using face_ofD [OF c *] that face_of_imp_subset [OF c] by force + using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed - have snd: "snd ` c face_of S'" + have snd: "snd ` C face_of S'" by (force simp: face_of_def 1 conv sndab) - have cc: "rel_interior c \ rel_interior (fst ` c) \ rel_interior (snd ` c)" - by (force simp: face_of_Times rel_interior_Times conv fst snd \convex c\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) - have "c = fst ` c \ snd ` c" - apply (rule face_of_eq [OF c]) - apply (simp_all add: face_of_Times rel_interior_Times conv fst snd) - using False rel_interior_eq_empty \convex c\ cc - apply blast - done + have cc: "rel_interior C \ rel_interior (fst ` C) \ rel_interior (snd ` C)" + by (force simp: face_of_Times rel_interior_Times conv fst snd \convex C\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) + have "C = fst ` C \ snd ` C" + proof (rule face_of_eq [OF C]) + show "fst ` C \ snd ` C face_of S \ S'" + by (simp add: face_of_Times rel_interior_Times conv fst snd) + show "rel_interior C \ rel_interior (fst ` C \ snd ` C) \ {}" + using False rel_interior_eq_empty \convex C\ cc + by (auto simp: face_of_Times rel_interior_Times conv fst) + qed with fst snd show ?thesis by metis qed next @@ -617,13 +613,17 @@ show ?thesis proof assume L: ?lhs - have "F \ {x. a \ x \ b} \ F face_of {x. a \ x = b}" - using False - apply (simp add: frontier_halfspace_le [symmetric] rel_frontier_nonempty_interior [OF ine, symmetric]) - apply (rule face_of_subset [OF L]) - apply (simp add: face_of_subset_rel_frontier [OF L]) - apply (force simp: rel_frontier_def closed_halfspace_le) - done + have "F face_of {x. a \ x = b}" if "F \ {x. a \ x \ b}" + proof - + have "F face_of rel_frontier {x. a \ x \ b}" + proof (rule face_of_subset [OF L]) + show "F \ rel_frontier {x. a \ x \ b}" + by (simp add: L face_of_subset_rel_frontier that) + qed (force simp: rel_frontier_def closed_halfspace_le) + then show ?thesis + using False + by (simp add: frontier_halfspace_le rel_frontier_nonempty_interior [OF ine]) + qed with L show ?rhs using affine_hyperplane face_of_affine_eq by blast next @@ -655,10 +655,13 @@ done lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \ convex S" - apply (simp add: exposed_face_of_def face_of_refl_eq, auto) - apply (rule_tac x=0 in exI)+ - apply force - done +proof + assume S: "convex S" + have "S \ {x. 0 \ x \ 0} \ S = S \ {x. 0 \ x = 0}" + by auto + with S show "S exposed_face_of S" + using exposed_face_of_def face_of_refl_eq by blast +qed (simp add: exposed_face_of_def face_of_refl_eq) lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" by simp @@ -739,9 +742,7 @@ using QsubP assms by blast moreover have "Q \ {T. T exposed_face_of S} \ \Q exposed_face_of S" using \finite Q\ False - apply (induction Q rule: finite_induct) - using exposed_face_of_Int apply fastforce+ - done + by (induction Q rule: finite_induct; use exposed_face_of_Int in fastforce) ultimately show ?thesis by (simp add: IntQ) qed @@ -777,25 +778,28 @@ have lez: "u \ (x + y) \ z" if "x \ S" "y \ T" for x y using S that by auto have sef: "S \ {x. u \ x = u \ a0} exposed_face_of S" - apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) - apply (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) - done + proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) + show "\x. x \ S \ u \ x \ u \ a0" + by (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) + qed have tef: "T \ {x. u \ x = u \ b0} exposed_face_of T" - apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) - apply (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) - done + proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) + show "\x. x \ T \ u \ x \ u \ b0" + by (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) + qed have "{x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0} \ F" by (auto simp: feq) (metis inner_right_distrib p z) moreover have "F \ {x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0}" - apply (auto simp: feq) - apply (rename_tac x y) - apply (rule_tac x=x in exI) - apply (rule_tac x=y in exI, simp) - using z p \a0 \ S\ \b0 \ T\ - apply clarify - apply (simp add: inner_right_distrib) - apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) - done + proof - + have "\x y. \z = u \ (x + y); x \ S; y \ T\ + \ u \ x = u \ a0 \ u \ y = u \ b0" + using z p \a0 \ S\ \b0 \ T\ + apply (simp add: inner_right_distrib) + apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) + done + then show ?thesis + using feq by blast + qed ultimately have "F = {x + y |x y. x \ S \ {x. u \ x = u \ a0} \ y \ T \ {x. u \ x = u \ b0}}" by blast then show ?thesis @@ -848,9 +852,7 @@ for P Q using hull_subset by fastforce have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" - apply (rule *) - apply (simp only: le eq) - using Ssub by auto + by (rule *) (use le eq Ssub in auto) then show "S \ {x. - a' \ x \ - b'}" by auto show "S \ {x. a \ x = b} = S \ {x. - a' \ x = - b'}" @@ -890,28 +892,26 @@ lemma extreme_point_not_in_REL_INTERIOR: fixes S :: "'a::real_normed_vector set" shows "\x extreme_point_of S; S \ {x}\ \ x \ rel_interior S" -apply (simp add: face_of_singleton [symmetric]) -apply (blast dest: face_of_disjoint_rel_interior) -done + by (metis disjoint_iff face_of_disjoint_rel_interior face_of_singleton insertI1) lemma extreme_point_not_in_interior: - fixes S :: "'a::{real_normed_vector, perfect_space} set" - shows "x extreme_point_of S \ x \ interior S" -apply (case_tac "S = {x}") -apply (simp add: empty_interior_finite) -by (meson contra_subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) + fixes S :: "'a::{real_normed_vector, perfect_space} set" + assumes "x extreme_point_of S" shows "x \ interior S" +proof (cases "S = {x}") + case False + then show ?thesis + by (meson assms subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) +qed (simp add: empty_interior_finite) lemma extreme_point_of_face: "F face_of S \ v extreme_point_of F \ v extreme_point_of S \ v \ F" by (meson empty_subsetI face_of_face face_of_singleton insert_subset) lemma extreme_point_of_convex_hull: - "x extreme_point_of (convex hull S) \ x \ S" -apply (simp add: extreme_point_of_stillconvex) -using hull_minimal [of S "(convex hull S) - {x}" convex] -using hull_subset [of S convex] -apply blast -done + "x extreme_point_of (convex hull S) \ x \ S" + using hull_minimal [of S "(convex hull S) - {x}" convex] + using hull_subset [of S convex] + by (force simp add: extreme_point_of_stillconvex) proposition extreme_points_of_convex_hull: "{x. x extreme_point_of (convex hull S)} \ S" @@ -945,32 +945,32 @@ lemma extreme_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" -apply (simp add: face_of_singleton [symmetric]) -by (metis face_of_Int_supporting_hyperplane_le_strong convex_singleton) + by (metis convex_singleton face_of_Int_supporting_hyperplane_le_strong face_of_singleton) lemma extreme_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" -apply (simp add: face_of_singleton [symmetric]) -by (metis face_of_Int_supporting_hyperplane_ge_strong convex_singleton) + using extreme_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] + by simp lemma exposed_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" -apply (simp add: exposed_face_of_def face_of_singleton) -apply (force simp: extreme_point_of_Int_supporting_hyperplane_le) -done + unfolding exposed_face_of_def + by (force simp: face_of_singleton extreme_point_of_Int_supporting_hyperplane_le) lemma exposed_point_of_Int_supporting_hyperplane_ge: - "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" -using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] -by simp + "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" + using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] + by simp lemma extreme_point_of_convex_hull_insert: - "\finite S; a \ convex hull S\ \ a extreme_point_of (convex hull (insert a S))" -apply (case_tac "a \ S") -apply (simp add: hull_inc) -using face_of_convex_hulls [of "insert a S" "{a}"] -apply (auto simp: face_of_singleton hull_same) -done + assumes "finite S" "a \ convex hull S" + shows "a extreme_point_of (convex hull (insert a S))" +proof (cases "a \ S") + case False + then show ?thesis + using face_of_convex_hulls [of "insert a S" "{a}"] assms + by (auto simp: face_of_singleton hull_same) +qed (use assms in \simp add: hull_inc\) subsection\Facets\ @@ -1066,41 +1066,11 @@ have noax: "norm a \ norm x" and nobx: "norm b \ norm x" using xsup that by auto have "a \ b" using empty_iff open_segment_idem x by auto - have *: "(1 - u) * na + u * nb < norm x" if "na < norm x" "nb \ norm x" "0 < u" "u < 1" for na nb u - proof - - have "(1 - u) * na + u * nb < (1 - u) * norm x + u * nb" - by (simp add: that) - also have "... \ (1 - u) * norm x + u * norm x" - by (simp add: that) - finally have "(1 - u) * na + u * nb < (1 - u) * norm x + u * norm x" . - then show ?thesis - using scaleR_collapse [symmetric, of "norm x" u] by auto - qed - have "norm x < norm x" if "norm a < norm x" - using x - apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) - apply (rule norm_triangle_lt) - apply (simp add: norm_mult) - using * [of "norm a" "norm b"] nobx that - apply blast - done - moreover have "norm x < norm x" if "norm b < norm x" - using x - apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) - apply (rule norm_triangle_lt) - apply (simp add: norm_mult) - using * [of "norm b" "norm a" "1-u" for u] noax that - apply (simp add: add.commute) - done - ultimately have "\ (norm a < norm x) \ \ (norm b < norm x)" - by auto - then show ?thesis - using different_norm_3_collinear_points noax nobx that(3) by fastforce + show False + by (metis dist_0_norm dist_decreases_open_segment noax nobx not_le x) qed then show ?thesis - apply (rule_tac x=x in that) - apply (force simp: extreme_point_of_def \x \ S\) - done + by (meson \x \ S\ extreme_point_of_def that) qed subsection\Krein-Milman, the weaker form\ @@ -1116,10 +1086,7 @@ have "closed S" by (simp add: \compact S\ compact_imp_closed) have "closure (convex hull {x. x extreme_point_of S}) \ S" - apply (rule closure_minimal [OF hull_minimal \closed S\]) - using assms - apply (auto simp: extreme_point_of_def) - done + by (simp add: \closed S\ assms closure_minimal extreme_point_of_def hull_minimal) moreover have "u \ closure (convex hull {x. x extreme_point_of S})" if "u \ S" for u proof (rule ccontr) @@ -1167,8 +1134,9 @@ proof (induction n arbitrary: S rule: less_induct) case (less n S) show ?case proof (cases "0 \ rel_interior S") - case True with Krein_Milman show ?thesis - by (metis subsetD convex_convex_hull convex_rel_interior_closure less.prems(2) less.prems(3) rel_interior_subset) + case True with Krein_Milman less.prems + show ?thesis + by (metis subsetD convex_convex_hull convex_rel_interior_closure rel_interior_subset) next case False have "rel_interior S \ {}" @@ -1179,8 +1147,7 @@ and less_ay: "\y. y \ rel_interior S \ a \ 0 < a \ y" by (blast intro: supporting_hyperplane_rel_boundary intro!: less False) have face: "S \ {x. a \ x = 0} face_of S" - apply (rule face_of_Int_supporting_hyperplane_ge [OF \convex S\]) - using le_ay by auto + using face_of_Int_supporting_hyperplane_ge le_ay \convex S\ by auto then have co: "compact (S \ {x. a \ x = 0})" "convex (S \ {x. a \ x = 0})" using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+ have "a \ y = 0" if "y \ span (S \ {x. a \ x = 0})" for y @@ -1196,7 +1163,7 @@ then have "0 \ convex hull {x. x extreme_point_of (S \ {x. a \ x = 0})}" by (rule less.IH) (auto simp: co less.prems) then show ?thesis - by (metis (mono_tags, lifting) Collect_mono_iff \S \ {x. a \ x = 0} face_of S\ extreme_point_of_face hull_mono subset_iff) + by (metis (mono_tags, lifting) Collect_mono_iff face extreme_point_of_face hull_mono subset_iff) qed qed @@ -1297,9 +1264,17 @@ fixes S :: "'a::euclidean_space set" assumes "compact S" and T: "T face_of (convex hull S)" obtains s' where "s' \ S" "T = convex hull s'" -apply (rule_tac s' = "{x. x extreme_point_of T}" in that) -using T extreme_point_of_convex_hull extreme_point_of_face apply blast -by (metis (no_types) Krein_Milman_Minkowski assms compact_convex_hull convex_convex_hull face_of_imp_compact face_of_imp_convex) +proof + show "{x. x extreme_point_of T} \ S" + using T extreme_point_of_convex_hull extreme_point_of_face by blast + show "T = convex hull {x. x extreme_point_of T}" + proof (rule Krein_Milman_Minkowski) + show "compact T" + using T assms compact_convex_hull face_of_imp_compact by auto + show "convex T" + using T face_of_imp_convex by blast + qed +qed lemma face_of_convex_hull_aux: @@ -1343,8 +1318,7 @@ then show "F \ convex hull insert a (convex hull T \ convex hull S)" by (simp add: FeqT hull_mono) show "convex hull insert a (convex hull T \ convex hull S) \ F" - apply (rule hull_minimal) - using True by (auto simp: \F = convex hull T\ hull_inc) + by (simp add: FeqT True hull_inc hull_minimal) qed moreover have "convex hull T \ convex hull S face_of convex hull S" by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI) @@ -1390,6 +1364,10 @@ and b: "b \ convex hull S" and c: "c \ convex hull S" and "x \ F" for b c ub uc ux x proof - + have xah: "x \ affine hull S" + using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ by blast + have ah: "b \ affine hull S" "c \ affine hull S" + using b c convex_hull_subset_affine_hull by blast+ obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \ (1 - uc) *\<^sub>R a + uc *\<^sub>R c" and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x = (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" @@ -1402,10 +1380,7 @@ ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x" by (auto simp: algebra_simps) then have "a \ affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \ 0" - apply (rule face_of_convex_hull_aux) - using b c that apply (auto simp: algebra_simps) - using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ apply blast+ - done + by (rule face_of_convex_hull_aux) (use b c xah ah that in \auto simp: algebra_simps\) then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0" using a by blast with 0 have equx: "(1 - v) * ub + v * uc = ux" @@ -1415,26 +1390,22 @@ proof (cases "uc = 0") case True then show ?thesis - using equx 0 \0 \ ub\ \ub \ 1\ \v < 1\ \x \ F\ - apply (auto simp: algebra_simps) - apply (rule_tac x=x in exI, simp) - apply (rule_tac x=ub in exI, auto) - apply (metis add.left_neutral diff_eq_eq less_irrefl mult.commute mult_cancel_right1 real_vector.scale_cancel_left real_vector.scale_left_diff_distrib) - using \x \ F\ \uc \ 1\ apply blast - done + using equx \0 \ ub\ \ub \ 1\ \v < 1\ uxx \x \ F\ by force next case False show ?thesis proof (cases "ub = 0") case True then show ?thesis - using equx 0 \0 \ uc\ \uc \ 1\ \0 < v\ \x \ F\ \uc \ 0\ by (force simp: algebra_simps) + using equx \0 \ uc\ \uc \ 1\ \0 < v\ uxx \x \ F\ by force next case False then have "0 < ub" "0 < uc" using \uc \ 0\ \0 \ ub\ \0 \ uc\ by auto + then have "(1 - v) * ub > 0" "v * uc > 0" + by (simp_all add: \0 < uc\ \0 < v\ \v < 1\) then have "ux \ 0" - by (metis \0 < v\ \v < 1\ diff_ge_0_iff_ge dual_order.strict_implies_order equx leD le_add_same_cancel2 zero_le_mult_iff zero_less_mult_iff) + using equx \0 < v\ by auto have "b \ F \ c \ F" proof (cases "b = c") case True @@ -1489,8 +1460,7 @@ then obtain c where "c \ S" and T: "T = convex hull c" by blast have "affine hull c \ affine hull (S - c) = {}" - apply (rule disjoint_affine_hull [OF assms \c \ S\], auto) - done + by (intro disjoint_affine_hull [OF assms \c \ S\], auto) then have "affine hull c \ convex hull (S - c) = {}" using convex_hull_subset_affine_hull by fastforce then show ?lhs @@ -1533,27 +1503,36 @@ by (force simp: facet_of_def) then have "\ S \ {u}" using \T \ {}\ u by auto - have [simp]: "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" + have "aff_dim (S - {u}) = aff_dim S - 1" using assms \u \ S\ - apply (simp add: aff_dim_convex_hull affine_dependent_def) - apply (drule bspec, assumption) + unfolding affine_dependent_def by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S]) - show ?lhs - apply (subst u) - apply (simp add: \\ S \ {u}\ facet_of_def face_of_convex_hull_affine_independent [OF assms], blast) - done + then have "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" + by (simp add: aff_dim_convex_hull) + then show ?lhs + by (metis Diff_subset \T \ {}\ assms face_of_convex_hull_affine_independent facet_of_def u) qed lemma facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" - shows - "\affine_dependent S - \ (T facet_of (convex hull S) \ - 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" -apply (simp add: facet_of_convex_hull_affine_independent) -apply (auto simp: Set.subset_singleton_iff) -apply (metis Diff_cancel Int_empty_right Int_insert_right_if1 aff_independent_finite card_eq_0_iff card_insert_if card_mono card_subset_eq convex_hull_eq_empty eq_iff equals0D finite_insert finite_subset inf.absorb_iff2 insert_absorb insert_not_empty not_less_eq_eq numeral_2_eq_2) -done + assumes "\ affine_dependent S" + shows "(T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + then obtain x where + "x \ S" and x: "T = convex hull (S - {x})" and "finite S" + using assms facet_of_convex_hull_affine_independent aff_independent_finite by blast + moreover have "Suc (Suc 0) \ card S" + using L x \x \ S\ \finite S\ + by (metis Suc_leI assms card.remove convex_hull_eq_empty card_gt_0_iff facet_of_convex_hull_affine_independent finite_Diff not_less_eq_eq) + ultimately show ?rhs + by auto +next + assume ?rhs then show ?lhs + using assms + by (auto simp: facet_of_convex_hull_affine_independent Set.subset_singleton_iff) +qed lemma segment_face_of: assumes "(closed_segment a b) face_of S" @@ -1584,10 +1563,11 @@ have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast also have "... \ ?rhs" - apply (rule hull_mono) - apply (auto simp: frontier_def extreme_point_not_in_interior) - using closure_subset apply (force simp: extreme_point_of_def) - done + proof (rule hull_mono) + show "{x. x extreme_point_of S} \ frontier S" + using closure_subset + by (auto simp: frontier_def extreme_point_not_in_interior extreme_point_of_def) + qed finally show "?lhs \ ?rhs" . next have "?rhs \ convex hull S" @@ -1603,9 +1583,12 @@ "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" -apply (simp add: polytope_def, safe) -apply (metis convex_hull_translation finite_imageI translation_galois) -by (metis convex_hull_translation finite_imageI) +proof - + have "\a A. polytope A \ polytope ((+) a ` A)" + by (metis (no_types) convex_hull_translation finite_imageI polytope_def) + then show ?thesis + by (metis (no_types) add.left_inverse image_add_0 translation_assoc) +qed lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" unfolding polytope_def using convex_hull_linear_image by blast @@ -1702,9 +1685,9 @@ lemma polyhedron_Int [intro,simp]: "\polyhedron S; polyhedron T\ \ polyhedron (S \ T)" - apply (simp add: polyhedron_def, clarify) - apply (rename_tac F G) - apply (rule_tac x="F \ G" in exI, auto) + apply (clarsimp simp add: polyhedron_def) + subgoal for F G + by (rule_tac x="F \ G" in exI, auto) done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" @@ -1718,21 +1701,17 @@ lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)" proof - - have "\a. a \ 0 \ - (\b. {x. (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b})" - by (rule_tac x="(SOME i. i \ Basis)" in exI) (force simp: SOME_Basis nonzero_Basis) - moreover have "\a b. a \ 0 \ - {x. - (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b}" - apply (rule_tac x="-(SOME i. i \ Basis)" in exI) + define i::'a where "(i \ SOME i. i \ Basis)" + have "\a. a \ 0 \ (\b. {x. i \ x \ -1} = {x. a \ x \ b})" + by (rule_tac x="i" in exI) (force simp: i_def SOME_Basis nonzero_Basis) + moreover have "\a b. a \ 0 \ {x. -i \ x \ - 1} = {x. a \ x \ b}" + apply (rule_tac x="-i" in exI) apply (rule_tac x="-1" in exI) - apply (simp add: SOME_Basis nonzero_Basis) + apply (simp add: i_def SOME_Basis nonzero_Basis) done ultimately show ?thesis unfolding polyhedron_def - apply (rule_tac x="{{x. (SOME i. i \ Basis) \ x \ -1}, - {x. -(SOME i. i \ Basis) \ x \ -1}}" in exI) - apply force - done + by (rule_tac x="{{x. i \ x \ -1}, {x. -i \ x \ -1}}" in exI) force qed lemma polyhedron_halfspace_le: @@ -1770,14 +1749,12 @@ lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S" -apply (simp add: polyhedron_def) -using closed_halfspace_le by fastforce + by (metis closed_Inter closed_halfspace_le polyhedron_def) lemma polyhedron_imp_convex: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ convex S" -apply (simp add: polyhedron_def) -using convex_Inter convex_halfspace_le by fastforce + by (metis convex_Inter convex_halfspace_le polyhedron_def) lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" @@ -1795,16 +1772,10 @@ (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs - apply (simp add: polyhedron_def) - apply (erule ex_forward) - using hull_subset apply force - done + using hull_subset polyhedron_def by fastforce next assume ?rhs then show ?lhs - apply clarify - apply (erule ssubst) - apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le) - done + by (metis polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le) qed proposition rel_interior_polyhedron_explicit: @@ -1846,13 +1817,10 @@ have \_aff: "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ affine hull S" by (metis \x \ S\ add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff) have "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ S" - apply (rule ins [OF _ \_aff]) - apply (simp add: algebra_simps lte) - done + using ins [OF _ \_aff] by (simp add: algebra_simps lte) then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \ S" - apply (rule_tac l = \ in that) - using \e > 0\ \z \ x\ apply (auto simp: \_def) - done + using \e > 0\ \z \ x\ + by (rule_tac l = \ in that) (auto simp: \_def) then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \ i" using seq \i \ F\ by auto have "b i * l + (a i \ x) * (1 - l) < a i \ (l *\<^sub>R z + (1 - l) *\<^sub>R x)" @@ -1870,7 +1838,7 @@ have 1: "\h. h \ F \ x \ interior h" by (metis interior_halfspace_le mem_Collect_eq less faceq) have 2: "\y. \\h\F. y \ interior h; y \ affine hull S\ \ y \ S" - by (metis IntI Inter_iff contra_subsetD interior_subset seq) + by (metis IntI Inter_iff subsetD interior_subset seq) show ?thesis apply (simp add: rel_interior \x \ S\) apply (rule_tac x="\h\F. interior h" in exI) @@ -1930,8 +1898,7 @@ qed next assume ?rhs then show ?lhs - apply (simp add: polyhedron_Int_affine) - by metis + by (metis polyhedron_Int_affine) qed @@ -1969,9 +1936,8 @@ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" by (subst seq) blast - have 2: "\F'. F' \ F \ S \ affine hull S \ \F'" - apply (frule *) - by (metis aff subsetCE subset_iff_psubset_eq) + have 2: "S \ affine hull S \ \F'" if "F' \ F" for F' + using * [OF that] by (metis IntE aff inf.strict_order_iff that) show ?rhs by (metis \finite F\ seq aff psubsetI 1 2) next @@ -1986,26 +1952,25 @@ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}) \ (\F'. F' \ F \ S \ (affine hull S) \ \F'))" -apply (rule iffI) - apply (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) -apply (auto simp: polyhedron_Int_affine elim!: ex_forward) -done + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) +qed (auto simp: polyhedron_Int_affine elim!: ex_forward) proposition facet_of_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" - shows "c facet_of S \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" + shows "C facet_of S \ (\h. h \ F \ C = S \ {x. a h \ x = b h})" proof (cases "S = {}") case True with psub show ?thesis by force next case False have "polyhedron S" - apply (simp add: polyhedron_Int_affine) - apply (rule_tac x=F in exI) - using assms apply force - done + unfolding polyhedron_Int_affine by (metis \finite F\ faceq seq) then have "convex S" by (rule polyhedron_imp_convex) with False rel_interior_eq_empty have "rel_interior S \ {}" by blast @@ -2040,31 +2005,28 @@ have "(1 - l) * (a i \ x) < (1 - l) * b i" by (simp add: \l < 1\ \i \ F\) moreover have "l * (a i \ z) \ l * b i" - apply (rule mult_left_mono) - apply (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) - using \0 < l\ - apply simp - done + proof (rule mult_left_mono) + show "a i \ z \ b i" + by (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) + qed (use \0 < l\ in auto) ultimately show ?thesis by (simp add: w_def algebra_simps) qed have weq: "a h \ w = b h" using xltz unfolding w_def l_def by (simp add: algebra_simps) (simp add: field_simps) + have faceS: "S \ {x. a h \ x = b h} face_of S" + proof (rule face_of_Int_supporting_hyperplane_le) + show "\x. x \ S \ a h \ x \ b h" + using faceq seq that by fastforce + qed fact have "w \ affine hull S" by (simp add: w_def mem_affine xaff zaff) moreover have "w \ \F" using \a h \ w = b h\ awlt faceq less_eq_real_def by blast ultimately have "w \ S" using seq by blast - with weq have "S \ {x. a h \ x = b h} \ {}" by blast - moreover have "S \ {x. a h \ x = b h} face_of S" - apply (rule face_of_Int_supporting_hyperplane_le) - apply (rule \convex S\) - apply (subst (asm) seq) - using faceq that apply fastforce - done - moreover have "affine hull (S \ {x. a h \ x = b h}) = - (affine hull S) \ {x. a h \ x = b h}" + with weq have ne: "S \ {x. a h \ x = b h} \ {}" by blast + moreover have "affine hull (S \ {x. a h \ x = b h}) = (affine hull S) \ {x. a h \ x = b h}" proof show "affine hull (S \ {x. a h \ x = b h}) \ affine hull S \ {x. a h \ x = b h}" apply (intro Int_greatest hull_mono Int_lower1) @@ -2084,8 +2046,7 @@ case False then obtain h' where h': "h' \ F - {h}" by auto let ?body = "(\j. if 0 < a j \ y - a j \ w - then (b j - a j \ w) / (a j \ y - a j \ w) - else 1) ` (F - {h})" + then (b j - a j \ w) / (a j \ y - a j \ w) else 1) ` (F - {h})" define inff where "inff = Inf ?body" from \finite F\ have "finite ?body" by blast @@ -2109,11 +2070,8 @@ proof (cases "a j \ w < a j \ y") case True then have "inff \ (b j - a j \ w) / (a j \ y - a j \ w)" - apply (simp add: inff_def) - apply (rule cInf_le_finite) - using \finite F\ apply blast - apply (simp add: that split: if_split_asm) - done + unfolding inff_def + using \finite F\ by (auto intro: cInf_le_finite simp add: that split: if_split_asm) then show ?thesis using \0 < inff\ awlt [OF that] mult_strict_left_mono by (fastforce simp add: field_split_simps split: if_split_asm) @@ -2128,7 +2086,7 @@ ultimately show ?thesis by (blast intro: that) qed - define c where "c = (1 - T) *\<^sub>R w + T *\<^sub>R y" + define C where "C = (1 - T) *\<^sub>R w + T *\<^sub>R y" have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ j" if "j \ F" for j proof (cases "j = h") case True @@ -2142,69 +2100,67 @@ with faceq [OF that] show ?thesis by simp qed moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ affine hull S" - apply (rule affine_affine_hull [simplified affine_alt, rule_format]) - apply (simp add: \w \ affine hull S\) - using yaff apply blast - done - ultimately have "c \ S" - using seq by (force simp: c_def) - moreover have "a h \ c = b h" - using yaff by (force simp: c_def algebra_simps weq) - ultimately have caff: "c \ affine hull (S \ {y. a h \ y = b h})" + using yaff \w \ affine hull S\ affine_affine_hull affine_alt by blast + ultimately have "C \ S" + using seq by (force simp: C_def) + moreover have "a h \ C = b h" + using yaff by (force simp: C_def algebra_simps weq) + ultimately have caff: "C \ affine hull (S \ {y. a h \ y = b h})" by (simp add: hull_inc) have waff: "w \ affine hull (S \ {y. a h \ y = b h})" using \w \ S\ weq by (blast intro: hull_inc) - have yeq: "y = (1 - inverse T) *\<^sub>R w + c /\<^sub>R T" - using \0 < T\ by (simp add: c_def algebra_simps) + have yeq: "y = (1 - inverse T) *\<^sub>R w + C /\<^sub>R T" + using \0 < T\ by (simp add: C_def algebra_simps) show "y \ affine hull (S \ {y. a h \ y = b h})" by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff]) qed qed - ultimately show ?thesis - apply (simp add: facet_of_def) - apply (subst aff_dim_affine_hull [symmetric]) - using \b h < a h \ z\ zaff - apply (force simp: aff_dim_affine_Int_hyperplane) - done + ultimately have "aff_dim (affine hull (S \ {x. a h \ x = b h})) = aff_dim S - 1" + using \b h < a h \ z\ zaff by (force simp: aff_dim_affine_Int_hyperplane) + then show ?thesis + by (simp add: ne faceS facet_of_def) qed show ?thesis proof - show "\h. h \ F \ c = S \ {x. a h \ x = b h} \ c facet_of S" + show "\h. h \ F \ C = S \ {x. a h \ x = b h} \ C facet_of S" using * by blast next - assume "c facet_of S" - then have "c face_of S" "convex c" "c \ {}" and affc: "aff_dim c = aff_dim S - 1" + assume "C facet_of S" + then have "C face_of S" "convex C" "C \ {}" and affc: "aff_dim C = aff_dim S - 1" by (auto simp: facet_of_def face_of_imp_convex) - then obtain x where x: "x \ rel_interior c" + then obtain x where x: "x \ rel_interior C" by (force simp: rel_interior_eq_empty) - then have "x \ c" + then have "x \ C" by (meson subsetD rel_interior_subset) then have "x \ S" - using \c facet_of S\ facet_of_imp_subset by blast + using \C facet_of S\ facet_of_imp_subset by blast have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF assms]) - have "c \ S" - using \c facet_of S\ facet_of_irrefl by blast + have "C \ S" + using \C facet_of S\ facet_of_irrefl by blast then have "x \ rel_interior S" - by (metis IntI empty_iff \x \ c\ \c \ S\ \c face_of S\ face_of_disjoint_rel_interior) + by (metis IntI empty_iff \x \ C\ \C \ S\ \C face_of S\ face_of_disjoint_rel_interior) with rels \x \ S\ obtain i where "i \ F" and i: "a i \ x \ b i" by force have "x \ {u. a i \ u \ b i}" by (metis IntD2 InterE \i \ F\ \x \ S\ faceq seq) then have "a i \ x \ b i" by simp then have "a i \ x = b i" using i by auto - have "c \ S \ {x. a i \ x = b i}" - apply (rule subset_of_face_of [of _ S]) - apply (simp add: "*" \i \ F\ facet_of_imp_face_of) - apply (simp add: \c face_of S\ face_of_imp_subset) - using \a i \ x = b i\ \x \ S\ x by blast - then have cface: "c face_of (S \ {x. a i \ x = b i})" - by (meson \c face_of S\ face_of_subset inf_le1) + have "C \ S \ {x. a i \ x = b i}" + proof (rule subset_of_face_of [of _ S]) + show "S \ {x. a i \ x = b i} face_of S" + by (simp add: "*" \i \ F\ facet_of_imp_face_of) + show "C \ S" + by (simp add: \C face_of S\ face_of_imp_subset) + show "S \ {x. a i \ x = b i} \ rel_interior C \ {}" + using \a i \ x = b i\ \x \ S\ x by blast + qed + then have cface: "C face_of (S \ {x. a i \ x = b i})" + by (meson \C face_of S\ face_of_subset inf_le1) have con: "convex (S \ {x. a i \ x = b i})" by (simp add: \convex S\ convex_Int convex_hyperplane) - show "\h. h \ F \ c = S \ {x. a h \ x = b h}" + show "\h. h \ F \ C = S \ {x. a h \ x = b h}" apply (rule_tac x=i in exI) - apply (simp add: \i \ F\) by (metis (no_types) * \i \ F\ affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface]) qed qed @@ -2216,28 +2172,26 @@ and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" - and c: "c face_of S" and "c \ {}" "c \ S" - obtains h where "h \ F" "c \ S \ {x. a h \ x = b h}" + and C: "C face_of S" and "C \ {}" "C \ S" + obtains h where "h \ F" "C \ S \ {x. a h \ x = b h}" proof - - have "c \ S" using \c face_of S\ + have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" - apply (simp add: polyhedron_Int_affine) - by (metis \finite F\ faceq seq) + by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ {x. a h \ x = b h}) face_of S" if "h \ F" for h - apply (rule face_of_Int_supporting_hyperplane_le) - using faceq seq that by fastforce - have "rel_interior c \ {}" - using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast - then obtain x where "x \ rel_interior c" by auto + using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce + have "rel_interior C \ {}" + using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast + then obtain x where "x \ rel_interior C" by auto have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "x \ rel_interior S" - by (metis IntI \x \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) + by (metis IntI \x \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "x \ S" - using \c \ S\ \x \ rel_interior c\ rel_interior_subset by auto + using \C \ S\ \x \ rel_interior C\ rel_interior_subset by auto then have xint: "x \ \F" using seq by blast have "F \ {}" using assms @@ -2249,11 +2203,10 @@ have face: "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\) show ?thesis - apply (rule_tac h = i in that) - apply (rule \i \ F\) - apply (rule subset_of_face_of [OF face \c \ S\]) - using \a i \ x = b i\ \x \ rel_interior c\ \x \ S\ apply blast - done + proof + show "C \ S \ {x. a i \ x = b i}" + using subset_of_face_of [OF face \C \ S\] \a i \ x = b i\ \x \ rel_interior C\ \x \ S\ by blast + qed fact qed text\Initial part of proof duplicates that above\ @@ -2263,29 +2216,27 @@ and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" - and c: "c face_of S" and "c \ {}" "c \ S" - shows "c = \{S \ {x. a h \ x = b h} | h. h \ F \ c \ S \ {x. a h \ x = b h}}" + and C: "C face_of S" and "C \ {}" "C \ S" + shows "C = \{S \ {x. a h \ x = b h} | h. h \ F \ C \ S \ {x. a h \ x = b h}}" proof - let ?ab = "\h. {x. a h \ x = b h}" - have "c \ S" using \c face_of S\ + have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" - apply (simp add: polyhedron_Int_affine) - by (metis \finite F\ faceq seq) + by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ ?ab h) face_of S" if "h \ F" for h - apply (rule face_of_Int_supporting_hyperplane_le) - using faceq seq that by fastforce - have "rel_interior c \ {}" - using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast - then obtain z where z: "z \ rel_interior c" by auto + using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce + have "rel_interior C \ {}" + using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast + then obtain z where z: "z \ rel_interior C" by auto have rels: "rel_interior S = {z \ S. \h\F. a h \ z < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "z \ rel_interior S" - by (metis IntI \z \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) + by (metis IntI \z \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "z \ S" - using \c \ S\ \z \ rel_interior c\ rel_interior_subset by auto + using \C \ S\ \z \ rel_interior C\ rel_interior_subset by auto with seq have xint: "z \ \F" by blast have "open (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto simp: \finite F\ open_halfspace_lt open_INT) @@ -2294,20 +2245,15 @@ by (auto intro: openE [of _ z]) then have e: "\h. \h \ F; a h \ z < b h\ \ ball z e \ {w. a h \ w < b h}" by blast - have "c \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h + have "C \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h proof - show "z \ S \ ?ab h \ c \ S \ ?ab h" - apply (rule subset_of_face_of [of _ S]) - using that \c \ S\ \z \ rel_interior c\ - using facet_of_polyhedron_explicit [OF \finite F\ seq faceq psub] - unfolding facet_of_def - apply auto - done + show "z \ S \ ?ab h \ C \ S \ ?ab h" + by (metis "*" Collect_cong IntI \C \ S\ empty_iff subset_of_face_of that z) next - show "c \ S \ ?ab h \ z \ S \ ?ab h" - using \z \ rel_interior c\ rel_interior_subset by force + show "C \ S \ ?ab h \ z \ S \ ?ab h" + using \z \ rel_interior C\ rel_interior_subset by force qed - then have **: "{S \ ?ab h | h. h \ F \ c \ S \ c \ ?ab h} = + then have **: "{S \ ?ab h | h. h \ F \ C \ S \ C \ ?ab h} = {S \ ?ab h |h. h \ F \ z \ S \ ?ab h}" by blast have bsub: "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} @@ -2335,8 +2281,7 @@ then show ?thesis by blast qed have 1: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S" - apply (rule hull_mono) - using that \z \ S\ by auto + using that \z \ S\ by (intro hull_mono) auto have 2: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{?ab h |h. h \ F \ a h \ z = b h}" by (rule hull_minimal) (auto intro: affine_hyperplane) @@ -2346,34 +2291,30 @@ for A B C D E by blast show ?thesis by (intro * 1 2 3) qed - have "\h. h \ F \ c \ ?ab h" - apply (rule face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub]) - using assms by auto - then have fac: "\{S \ ?ab h |h. h \ F \ c \ S \ ?ab h} face_of S" - using * by (force simp: \c \ S\ intro: face_of_Inter) - have red: - "(\a. P a \ T \ S \ \{F x |x. P x}) \ T \ \{S \ F x |x. P x}" - for P T F by blast + have "\h. h \ F \ C \ ?ab h" + using assms + by (metis face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub] le_inf_iff) + then have fac: "\{S \ ?ab h |h. h \ F \ C \ S \ ?ab h} face_of S" + using * by (force simp: \C \ S\ intro: face_of_Inter) + have red: "(\a. P a \ T \ S \ \{F X |X. P X}) \ T \ \{S \ F X |X::'a set. P X}" for P T F + by blast have "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{S \ ?ab h |h. h \ F \ a h \ z = b h}" - apply (rule red) - apply (metis seq bsub) - done + by (rule red) (metis seq bsub) with \0 < e\ have zinrel: "z \ rel_interior (\{S \ ?ab h |h. h \ F \ z \ S \ a h \ z = b h})" by (auto simp: mem_rel_interior_ball \z \ S\) show ?thesis - apply (rule face_of_eq [OF c fac]) - using z zinrel apply (force simp: **) - done + using z zinrel + by (intro face_of_eq [OF C fac]) (force simp: **) qed subsection\More general corollaries from the explicit representation\ corollary facet_of_polyhedron: - assumes "polyhedron S" and "c facet_of S" - obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "c = S \ {x. a \ x = b}" + assumes "polyhedron S" and "C facet_of S" + obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "C = S \ {x. a \ x = b}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" @@ -2381,19 +2322,18 @@ using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis - obtain i where "i \ F" and c: "c = S \ {x. a i \ x = b i}" + obtain i where "i \ F" and C: "C = S \ {x. a i \ x = b i}" using facet_of_polyhedron_explicit [OF \finite F\ seq ab min] assms by force moreover have ssub: "S \ {x. a i \ x \ b i}" - apply (subst seq) - using \i \ F\ ab by auto + using \i \ F\ ab by (subst seq) auto ultimately show ?thesis by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab) qed corollary face_of_polyhedron: - assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" - shows "c = \{F. F facet_of S \ c \ F}" + assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" + shows "C = \{F. F facet_of S \ C \ F}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" @@ -2408,10 +2348,10 @@ qed lemma face_of_polyhedron_subset_facet: - assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" - obtains F where "F facet_of S" "c \ F" -using face_of_polyhedron assms -by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) + assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" + obtains F where "F facet_of S" "C \ F" + using face_of_polyhedron assms + by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) lemma exposed_face_of_polyhedron: @@ -2432,13 +2372,10 @@ by (metis Collect_empty_eq_bot \F face_of S\ assms empty_def face_of_polyhedron_subset_facet) moreover have "\T. \T facet_of S; F \ T\ \ T exposed_face_of S" by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron) - ultimately have "\{fa. - fa facet_of S \ F \ fa} exposed_face_of S" + ultimately have "\{G. G facet_of S \ F \ G} exposed_face_of S" by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter) then show ?thesis - using False - apply (subst face_of_polyhedron [OF assms \F face_of S\], auto) - done + using False \F face_of S\ assms face_of_polyhedron by fastforce qed qed @@ -2464,7 +2401,6 @@ apply clarify apply (rename_tac c) apply (drule face_of_polyhedron_explicit [OF \finite F\ seq ab min, simplified], simp_all) - apply (erule ssubst) apply (rule_tac x="{h \ F. c \ S \ {x. a h \ x = b h}}" in exI, auto) done ultimately show ?thesis @@ -2477,16 +2413,19 @@ lemma finite_polyhedron_extreme_points: fixes S :: "'a :: euclidean_space set" - shows "polyhedron S \ finite {v. v extreme_point_of S}" -apply (simp add: face_of_singleton [symmetric]) -apply (rule finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) -done + assumes "polyhedron S" shows "finite {v. v extreme_point_of S}" +proof - + have "finite {v. {v} face_of S}" + using assms by (intro finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) + then show ?thesis + by (simp add: face_of_singleton) +qed lemma finite_polyhedron_facets: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {F. F facet_of S}" -unfolding facet_of_def -by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) + unfolding facet_of_def + by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) proposition rel_interior_of_polyhedron: @@ -2518,13 +2457,8 @@ using xnot by fastforce then have "F \ Collect ((\) h)" using 2 \x \ S\ facet by blast - with \h \ F\ have "\F \ S \ {x. a h \ x = b h}" by blast with 2 that \x \ \F\ show ?thesis - apply simp - apply (drule_tac x="\F" in spec) - apply (simp add: facet) - apply (drule_tac x=h in spec) - using seq by auto + by blast qed qed moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x @@ -2548,10 +2482,11 @@ lemma rel_frontier_of_polyhedron_alt: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" - shows "rel_frontier S = \ {F. F face_of S \ (F \ S)}" -apply (rule subset_antisym) - apply (force simp: rel_frontier_of_polyhedron facet_of_def assms) -using face_of_subset_rel_frontier by fastforce + shows "rel_frontier S = \ {F. F face_of S \ F \ S}" +proof + show "rel_frontier S \ \ {F. F face_of S \ F \ S}" + by (force simp: rel_frontier_of_polyhedron facet_of_def assms) +qed (use face_of_subset_rel_frontier in fastforce) text\A characterization of polyhedra as having finitely many faces\ @@ -2591,12 +2526,13 @@ with rel_interior_subset have "c \ S" by blast have ccp: "closed_segment c p \ affine hull S" by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE) + have oS: "openin (top_of_set (closed_segment c p)) (closed_segment c p \ rel_interior S)" + by (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) obtain x where xcl: "x \ closed_segment c p" and "x \ S" and xnot: "x \ rel_interior S" using connected_openin [of "closed_segment c p"] apply simp apply (drule_tac x="closed_segment c p \ rel_interior S" in spec) - apply (erule impE) - apply (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) + apply (drule mp [OF _ oS]) apply (drule_tac x="closed_segment c p \ (- S)" in spec) using rel_interior_subset \closed S\ c \p \ S\ apply blast done @@ -2623,9 +2559,7 @@ have sns: "S \ {y. d \ y = d \ x} \ S" by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset) obtain h where "h \ F" "x \ h" - apply (rule_tac h="S \ {y. d \ y = d \ x}" in that) - apply (simp_all add: F_def sex sne sns \x \ S\) - done + using F_def \x \ S\ sex sns by blast have abface: "{y. a h \ y = b h} face_of {y. a h \ y \ b h}" using hyperplane_face_of_halfspace_le by blast then have "c \ h" @@ -2684,8 +2618,7 @@ lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" - by (subst polyhedron_linear_image_eq) - (auto simp: bij_uminus intro!: linear_uminus) + by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ @@ -2699,11 +2632,13 @@ by (simp add: finite_polytope_faces polyhedron_eq_finite_faces polytope_imp_closed polytope_imp_convex polytope_imp_bounded) next - assume ?rhs then show ?lhs - unfolding polytope_def - apply (rule_tac x="{v. v extreme_point_of S}" in exI) - apply (simp add: finite_polyhedron_extreme_points Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) - done + assume R: ?rhs + then have "finite {v. v extreme_point_of S}" + by (simp add: finite_polyhedron_extreme_points) + moreover have "S = convex hull {v. v extreme_point_of S}" + using R by (simp add: Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) + ultimately show ?lhs + unfolding polytope_def by blast qed lemma polytope_Int: @@ -2787,37 +2722,31 @@ then have ccs: "closed (convex hull S)" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) { fix x T - assume "finite T" "T \ S" "int (card T) \ aff_dim S + 1" "x \ convex hull T" + assume "int (card T) \ aff_dim S + 1" "finite T" "T \ S""x \ convex hull T" then have "S \ T" using True \finite S\ aff_dim_le_card affine_independent_iff_card by fastforce then obtain a where "a \ S" "a \ T" using \T \ S\ by blast - then have "x \ (\a\S. convex hull (S - {a}))" + then have "\y\S. x \ convex hull (S - {y})" using True affine_independent_iff_card [of S] - apply simp - apply (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) - done + by (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) } note * = this have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" - apply (subst caratheodory_aff_dim) - apply (blast intro: *) - done + by (subst caratheodory_aff_dim) (blast dest: *) have 2: "\((\a. convex hull (S - {a})) ` S) \ convex hull S" by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE) show ?thesis using True apply (simp add: segment_convex_hull frontier_def) using interior_convex_hull_eq_empty [OF assms] apply (simp add: closure_closed [OF ccs]) - apply (rule subset_antisym) - using 1 apply blast - using 2 apply blast - done + using "1" "2" by auto next case False - then have "frontier (convex hull S) = (convex hull S) - rel_interior(convex hull S)" - apply (simp add: rel_boundary_of_convex_hull [symmetric] frontier_def) - by (metis aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) - also have "... = \{convex hull (S - {a}) |a. a \ S}" + then have "frontier (convex hull S) = closure (convex hull S) - interior (convex hull S)" + by (simp add: rel_boundary_of_convex_hull frontier_def) + also have "\ = (convex hull S) - rel_interior(convex hull S)" + by (metis False aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) + also have "\ = \{convex hull (S - {a}) |a. a \ S}" proof - have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)" by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron) @@ -2997,13 +2926,10 @@ by (meson bounded_pos_less) define C where "C \ {z \ \. \z * e / 2 / real DIM('a)\ \ B}" define I where "I \ \i \ Basis. \j \ C. { (i::'a, j * e / 2 / DIM('a)) }" - have "finite C" - using finite_int_segment [of "-B / (e / 2 / DIM('a))" "B / (e / 2 / DIM('a))"] - apply (simp add: C_def) - apply (erule rev_finite_subset) - using \0 < e\ - apply (auto simp: field_split_simps) - done + have "C \ {x \ \. - B / (e / 2 / real DIM('a)) \ x \ x \ B / (e / 2 / real DIM('a))}" + using \0 < e\ by (auto simp: field_split_simps C_def) + then have "finite C" + using finite_int_segment finite_subset by blast then have "finite I" by (simp add: I_def) obtain \' where eq: "\\' = \\" and "finite \'" @@ -3015,8 +2941,7 @@ and sub1: "\C. C \ \' \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" apply (rule exE [OF cell_subdivision_lemma]) - using assms \finite I\ apply auto - done + using assms \finite I\ by auto show ?thesis proof (rule_tac \'="\'" in that) show "diameter X < e" if "X \ \'" for X @@ -3141,20 +3066,17 @@ lemma zero_simplex_sing: "0 simplex {a}" apply (simp add: simplex_def) - by (metis affine_independent_1 card.empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI) + using affine_independent_1 card_1_singleton_iff convex_hull_singleton by blast lemma simplex_sing [simp]: "n simplex {a} \ n = 0" using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast lemma simplex_zero: "0 simplex S \ (\a. S = {a})" -apply (auto simp: ) - using aff_dim_eq_0 aff_dim_simplex by blast + by (metis aff_dim_eq_0 aff_dim_simplex simplex_sing) lemma one_simplex_segment: "a \ b \ 1 simplex closed_segment a b" - apply (simp add: simplex_def) - apply (rule_tac x="{a,b}" in exI) - apply (auto simp: segment_convex_hull) - done + unfolding simplex_def + by (rule_tac x="{a,b}" in exI) (auto simp: segment_convex_hull) lemma simplex_segment_cases: "(if a = b then 0 else 1) simplex closed_segment a b" @@ -3282,8 +3204,7 @@ ultimately have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu - apply (simp add: between_mem_segment [symmetric]) - by (metis between_commute between_trans_2 between_antisym) + by (meson between_antisym between_mem_segment between_trans_2 ends_in_segment(2)) then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" @@ -3329,25 +3250,23 @@ have "\s. \n \ 1; s \ \\ \ \m. m simplex s" using poly\ aff\ by (force intro: polytope_lowdim_imp_simplex) then show ?thesis - unfolding simplicial_complex_def - apply (rule_tac x="\" in exI) - using True by (auto simp: less.prems) + unfolding simplicial_complex_def using True + by (rule_tac x="\" in exI) (auto simp: less.prems) next case False define \ where "\ \ {C \ \. aff_dim C < n}" have "finite \" "\C. C \ \ \ polytope C" "\C. C \ \ \ aff_dim C \ int (n - 1)" - "\C F. \C \ \; F face_of C\ \ F \ \" "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" - using less.prems - apply (auto simp: \_def) - by (metis aff_dim_subset face_of_imp_subset less_le not_le) - with less.IH [of "n-1" \] False - obtain \ where "simplicial_complex \" + using less.prems by (auto simp: \_def) + moreover have \
: "\C F. \C \ \; F face_of C\ \ F \ \" + using less.prems unfolding \_def + by (metis (no_types, lifting) mem_Collect_eq aff_dim_subset face_of_imp_subset less_le not_le) + ultimately obtain \ where "simplicial_complex \" and aff_dim\: "\K. K \ \ \ aff_dim K \ int (n - 1)" and "\\ = \\" and fin\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and C\: "\K. K \ \ \ \C. C \ \ \ K \ C" - by auto + using less.IH [of "n-1" \] False by auto then have "finite \" and simpl\: "\S. S \ \ \ \n. n simplex S" and face\: "\F S. \S \ \; F face_of S\ \ F \ \" @@ -3392,11 +3311,9 @@ using \_def \_def \C \ \\ \L \ \\ intface\ by (simp add: inf_commute) moreover have "L \ C \ C" using \C \ \\ \L \ \\ - apply (clarsimp simp: \_def \_def) - by (metis aff_dim_subset inf_le1 not_le) + by (metis (mono_tags, lifting) \_def \_def intface\ mem_Collect_eq not_le order_refl \
) moreover have "K \ L \ C" - using \C \ \\ \L \ \\ \K \ C\ \K \ L\ - by (auto simp: \_def \_def) + using \C \ \\ \L \ \\ \K \ C\ \K \ L\ by (auto simp: \_def \_def) ultimately show ?thesis using that by metis qed have "affine hull F \ rel_interior C = {}" @@ -3504,8 +3421,10 @@ qed finally have DC: "D \ rel_interior C = {}" . have eq: "X \ convex hull (insert ?z K) = X \ convex hull K" - apply (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) - using DC by (meson \X \ D\ disjnt_def disjnt_subset1) + proof (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) + show "disjnt X (rel_interior C)" + using DC by (meson \X \ D\ disjnt_def disjnt_subset1) + qed obtain I where I: "\ affine_dependent I" and Keq: "K = convex hull I" and [simp]: "convex hull K = K" using "*" \K \ \\ by force @@ -3558,11 +3477,11 @@ case True then have "L \ rel_frontier C" using \L \ rel_frontier D\ by auto - show ?thesis - apply (simp add: X Y True) - apply (simp add: convex_hull_insert_Int_eq [OF z] \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\) - using face_of_polytope_insert2 - by (metis "*" IntI \C \ \\ \K \ \\ \L \ \\\K \ rel_frontier C\ \L \ rel_frontier C\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_convex_hull z) + have "convex hull insert (SOME z. z \ rel_interior C) (K \ L) face_of + convex hull insert (SOME z. z \ rel_interior C) K" + by (metis face_of_polytope_insert2 "*" IntI \C \ \\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_def z \K \ \\ \L \ \\\K \ rel_frontier C\) + then show ?thesis + using True X Y \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\ convex_hull_insert_Int_eq z by force next case False have "convex D" @@ -3590,11 +3509,11 @@ finally have CD: "C \ (rel_interior D) = {}" . have zKC: "(convex hull insert ?z K) \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed convex\ hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z) - have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = - convex hull (insert ?z K) \ convex hull L" - apply (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) - using zKC CD apply (force simp: disjnt_def) - done + have "disjnt (convex hull insert (SOME z. z \ rel_interior C) K) (rel_interior D)" + using zKC CD by (force simp: disjnt_def) + then have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = + convex hull (insert ?z K) \ convex hull L" + by (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) have ch_id: "convex hull K = K" "convex hull L = L" using "*" \K \ \\ \L \ \\ hull_same by auto have "convex C" @@ -3719,7 +3638,6 @@ ultimately show ?thesis by blast qed - have "(\C. C \ \ \ L \ C) \ aff_dim L \ int n" if "L \ \ \ ?\" for L using that proof @@ -3802,11 +3720,8 @@ have "\(\C\\. {F. F face_of C}) = \\" using face_of_imp_subset face by blast ultimately show ?thesis - apply clarify - apply (rule that, assumption+) - using n apply blast - apply (simp_all add: poly face_of_refl polytope_imp_convex) - using face_of_imp_subset by fastforce + using face_of_imp_subset n + by (fastforce intro!: that simp add: poly face_of_refl polytope_imp_convex) next case False then have m1: "\C. C \ \ \ aff_dim C = -1"