# HG changeset patch # User paulson # Date 1462885484 -3600 # Node ID e49dc94eb6241e35c6e958f79df7361bff3ad4b5 # Parent 844725394a37479a5ae81215dd5ae78cfd8c8ba0 Theory of polyhedra: faces, extreme points, polytopes, and the Krein–Milman Minkowski theorem diff -r 844725394a37 -r e49dc94eb624 NEWS --- a/NEWS Tue May 10 11:56:23 2016 +0100 +++ b/NEWS Tue May 10 14:04:44 2016 +0100 @@ -158,7 +158,10 @@ INCOMPATIBILITY. * More complex analysis including Cauchy's inequality, Liouville theorem, -open mapping theorem, maximum modulus principle, Schwarz Lemma. +open mapping theorem, maximum modulus principle, Residue theorem, Schwarz Lemma. + +* Theory of polyhedra: faces, extreme points, polytopes, and the Krein–Milman +Minkowski theorem. * "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f ` A)". diff -r 844725394a37 -r e49dc94eb624 src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Tue May 10 11:56:23 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Tue May 10 14:04:44 2016 +0100 @@ -6,6 +6,7 @@ Ordered_Euclidean_Space Bounded_Continuous_Function Weierstrass + Polytope Conformal_Mappings Generalised_Binomial_Theorem Gamma diff -r 844725394a37 -r e49dc94eb624 src/HOL/Multivariate_Analysis/Polytope.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Polytope.thy Tue May 10 14:04:44 2016 +0100 @@ -0,0 +1,2613 @@ +section \Faces, Extreme Points, Polytopes, Polyhedra etc.\ + +text\Ported from HOL Light by L C Paulson\ + +theory Polytope +imports Cartesian_Euclidean_Space +begin + +subsection \Faces of a (usually convex) set\ + +definition face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) + where + "T face_of S \ + T \ S \ convex T \ + (\a \ S. \b \ S. \x \ T. x \ open_segment a b \ a \ T \ b \ T)" + +lemma face_ofD: "\T face_of S; x \ open_segment a b; a \ S; b \ S; x \ T\ \ a \ T \ b \ T" + unfolding face_of_def by blast + +lemma face_of_translation_eq [simp]: + "(op + a ` T face_of op + a ` S) \ T face_of S" +proof - + have *: "\a T S. T face_of S \ (op + a ` T face_of op + a ` S)" + apply (simp add: face_of_def Ball_def, clarify) + apply (drule open_segment_translation_eq [THEN iffD1]) + using inj_image_mem_iff inj_add_left apply metis + done + show ?thesis + apply (rule iffI) + apply (force simp: image_comp o_def dest: * [where a = "-a"]) + apply (blast intro: *) + done +qed + +lemma face_of_linear_image: + assumes "linear f" "inj f" + shows "(f ` c face_of f ` S) \ c face_of S" +by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms) + +lemma face_of_refl: "convex S \ S face_of S" + by (auto simp: face_of_def) + +lemma face_of_refl_eq: "S face_of S \ convex S" + by (auto simp: face_of_def) + +lemma empty_face_of [iff]: "{} face_of S" + by (simp add: face_of_def) + +lemma face_of_empty [simp]: "S face_of {} \ S = {}" + by (meson empty_face_of face_of_def subset_empty) + +lemma face_of_trans [trans]: "\S face_of T; T face_of u\ \ S face_of u" + unfolding face_of_def by (safe; blast) + +lemma face_of_face: "T face_of S \ (f face_of T \ f face_of S \ f \ T)" + unfolding face_of_def by (safe; blast) + +lemma face_of_subset: "\F face_of S; F \ T; T \ S\ \ F face_of T" + unfolding face_of_def by (safe; blast) + +lemma face_of_slice: "\F face_of S; convex T\ \ (F \ T) face_of (S \ T)" + unfolding face_of_def by (blast intro: convex_Int) + +lemma face_of_Int: "\t1 face_of S; t2 face_of S\ \ (t1 \ t2) face_of S" + unfolding face_of_def by (blast intro: convex_Int) + +lemma face_of_Inter: "\A \ {}; \T. T \ A \ T face_of S\ \ (\ A) face_of S" + unfolding face_of_def by (blast intro: convex_Inter) + +lemma face_of_Int_Int: "\F face_of T; F' face_of t'\ \ (F \ F') face_of (T \ t')" + unfolding face_of_def by (blast intro: convex_Int) + +lemma face_of_imp_subset: "T face_of S \ T \ S" + unfolding face_of_def by blast + +lemma face_of_imp_eq_affine_Int: + fixes S :: "'a::euclidean_space set" + assumes S: "convex S" "closed S" and T: "T face_of S" + shows "T = (affine hull T) \ S" +proof - + have "convex T" using T by (simp add: face_of_def) + have *: False if x: "x \ affine hull T" and "x \ S" "x \ T" and y: "y \ rel_interior T" for x y + proof - + obtain e where "e>0" and e: "cball y e \ affine hull T \ T" + 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" + 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 + 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 + 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 +qed + +lemma face_of_imp_closed: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "closed S" "T face_of S" shows "closed T" + by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms) + +lemma face_of_Int_supporting_hyperplane_le_strong: + assumes "convex(S \ {x. a \ x = b})" and aleb: "\x. x \ S \ a \ x \ b" + shows "(S \ {x. a \ x = b}) face_of S" +proof - + have *: "a \ u = a \ x" if "x \ open_segment u v" "u \ S" "v \ S" and b: "b = a \ x" + for u v x + proof (rule antisym) + show "a \ u \ a \ x" + using aleb \u \ S\ \b = a \ x\ by blast + next + obtain \ where "b = a \ ((1 - \) *\<^sub>R u + \ *\<^sub>R v)" "0 < \" "\ < 1" + using \b = a \ x\ \x \ open_segment u v\ in_segment + by (auto simp: open_segment_image_interval split: if_split_asm) + then have "b + \ * (a \ u) \ a \ u + \ * b" + using aleb [OF \v \ S\] by (simp add: algebra_simps) + then have "(1 - \) * b \ (1 - \) * (a \ u)" + by (simp add: algebra_simps) + then have "b \ a \ u" + using \\ < 1\ by auto + 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 +qed + +lemma face_of_Int_supporting_hyperplane_ge_strong: + "\convex(S \ {x. a \ x = b}); \x. x \ S \ a \ x \ b\ + \ (S \ {x. a \ x = b}) face_of S" + using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp + +lemma face_of_Int_supporting_hyperplane_le: + "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" + by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong) + +lemma face_of_Int_supporting_hyperplane_ge: + "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" + by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong) + +lemma face_of_imp_convex: "T face_of S \ convex T" + using face_of_def by blast + +lemma face_of_imp_compact: + fixes S :: "'a::euclidean_space set" + shows "\convex S; compact S; T face_of S\ \ compact T" + by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset) + +lemma face_of_Int_subface: + "c1 \ c2 face_of c1 \ c1 \ c2 face_of c2 \ d1 face_of c1 \ d2 face_of c2 + \ (d1 \ d2) face_of d1 \ (d1 \ d2) face_of d2" + by (meson face_of_Int_Int face_of_face inf_le1 inf_le2) + +lemma subset_of_face_of: + fixes S :: "'a::real_normed_vector set" + assumes "T face_of S" "u \ S" "T \ (rel_interior u) \ {}" + shows "u \ T" +proof + fix c + assume "c \ u" + obtain b where "b \ T" "b \ rel_interior u" using assms by auto + then obtain e where "e>0" "b \ u" and e: "cball b e \ affine hull u \ u" + by (auto simp: rel_interior_cball) + show "c \ T" + proof (cases "b=c") + case True with \b \ T\ show ?thesis by blast + next + case False + def d \ "b + (e / norm(b - c)) *\<^sub>R (b - c)" + have "d \ cball b e \ affine hull u" + using \e > 0\ \b \ u\ \c \ u\ + by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False) + with e have "d \ u" by blast + have nbc: "norm (b - c) + e > 0" using \e > 0\ + by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero) + then have [simp]: "d \ c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c] + by (simp add: algebra_simps d_def) (simp add: divide_simps) + have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" + using False nbc + apply (simp add: algebra_simps divide_simps) + by (metis mult_eq_0_iff norm_eq_zero norm_imp_pos_and_ge norm_pths(2) real_scaleR_def scaleR_left.add zero_less_norm_iff) + have "b \ open_segment d c" + 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 + 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 + 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) + +lemma face_of_disjoint_rel_interior: + fixes S :: "'a::real_normed_vector set" + assumes "T face_of S" "T \ S" + shows "T \ rel_interior S = {}" + by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym) + +lemma face_of_disjoint_interior: + fixes S :: "'a::real_normed_vector set" + assumes "T face_of S" "T \ S" + shows "T \ interior S = {}" +proof - + have "T \ interior S \ rel_interior S" + by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans) + thus ?thesis + by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty) +qed + +lemma face_of_subset_rel_boundary: + fixes S :: "'a::real_normed_vector set" + assumes "T face_of S" "T \ S" + shows "T \ (S - rel_interior S)" +by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI) + +lemma face_of_subset_rel_frontier: + fixes S :: "'a::real_normed_vector set" + assumes "T face_of S" "T \ S" + shows "T \ rel_frontier S" + using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce + +lemma face_of_aff_dim_lt: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "T face_of S" "T \ S" + shows "aff_dim T < aff_dim S" +proof - + have "aff_dim T \ aff_dim S" + by (simp add: face_of_imp_subset aff_dim_subset assms) + moreover have "aff_dim T \ aff_dim S" + proof (cases "T = {}") + case True then show ?thesis + by (metis aff_dim_empty \T \ S\) + next case False then show ?thesis + by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI) + qed + ultimately show ?thesis + by simp +qed + + +lemma affine_diff_divide: + assumes "affine S" "k \ 0" "k \ 1" and xy: "x \ S" "y /\<^sub>R (1 - k) \ S" + shows "(x - y) /\<^sub>R k \ S" +proof - + have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x" + using assms + by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] divide_simps) + then show ?thesis + using \affine S\ xy by (auto simp: affine_alt) +qed + +lemma face_of_convex_hulls: + assumes S: "finite S" "T \ S" and disj: "affine hull T \ convex hull (S - T) = {}" + shows "(convex hull T) face_of (convex hull S)" +proof - + have fin: "finite T" "finite (S - T)" using assms + by (auto simp: finite_subset) + have *: "x \ convex hull T" + if x: "x \ convex hull S" and y: "y \ convex hull S" and w: "w \ convex hull T" "w \ open_segment x y" + for x y w + proof - + have waff: "w \ affine hull T" + using convex_hull_subset_affine_hull w by blast + obtain a b where a: "\i. i \ S \ 0 \ a i" and asum: "setsum a S = 1" and aeqx: "(\i\S. a i *\<^sub>R i) = x" + and b: "\i. i \ S \ 0 \ b i" and bsum: "setsum b S = 1" and beqy: "(\i\S. b i *\<^sub>R i) = y" + using x y by (auto simp: assms convex_hull_finite) + obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \ convex hull T" "x \ y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y" + and u01: "0 < u" "u < 1" + using w by (auto simp: open_segment_image_interval split: if_split_asm) + def c \ "\i. (1 - u) * a i + u * b i" + have cge0: "\i. i \ S \ 0 \ c i" + using a b u01 by (simp add: c_def) + have sumc1: "setsum c S = 1" + by (simp add: c_def setsum.distrib setsum_right_distrib [symmetric] asum bsum) + have sumci_xy: "(\i\S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y" + apply (simp add: c_def setsum.distrib scaleR_left_distrib) + by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] aeqx beqy) + show ?thesis + proof (cases "setsum c (S - T) = 0") + case True + have ci0: "\i. i \ (S - T) \ c i = 0" + using True cge0 by (simp add: \finite S\ setsum_nonneg_eq_0_iff) + have a0: "a i = 0" if "i \ (S - T)" for i + using ci0 [OF that] u01 a [of i] b [of i] that + by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff) + have [simp]: "setsum a T = 1" + using assms by (metis setsum.mono_neutral_cong_right a0 asum) + show ?thesis + apply (simp add: convex_hull_finite \finite T\) + apply (rule_tac x=a in exI) + using a0 assms + apply (auto simp: cge0 a aeqx [symmetric] setsum.mono_neutral_right) + done + next + case False + def k \ "setsum c (S - T)" + have "k > 0" using False + unfolding k_def by (metis DiffD1 antisym_conv cge0 setsum_nonneg not_less) + have weq_sumsum: "w = setsum (\x. c x *\<^sub>R x) T + setsum (\x. c x *\<^sub>R x) (S - T)" + by (metis (no_types) add.commute S(1) S(2) setsum.subset_diff sumci_xy weq) + show ?thesis + proof (cases "k = 1") + case True + then have "setsum c T = 0" + by (simp add: S k_def setsum_diff sumc1) + then have [simp]: "setsum c (S - T) = 1" + by (simp add: S setsum_diff sumc1) + have ci0: "\i. i \ T \ c i = 0" + by (meson `finite T` `setsum c T = 0` \T \ S\ cge0 setsum_nonneg_eq_0_iff subsetCE) + then have [simp]: "(\i\S-T. c i *\<^sub>R i) = w" + by (simp add: weq_sumsum) + have "w \ convex hull (S - T)" + apply (simp add: convex_hull_finite fin) + apply (rule_tac x=c in exI) + apply (auto simp: cge0 weq True k_def) + done + then show ?thesis + using disj waff by blast + next + case False + then have sumcf: "setsum c T = 1 - k" + by (simp add: S k_def setsum_diff sumc1) + 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) setsum_nonneg subsetCE) + apply (metis False mult.commute right_inverse right_minus_eq setsum_right_distrib sumcf) + by (metis (mono_tags, lifting) scaleR_right.setsum scaleR_scaleR setsum.cong) + with `0 < k` have "inverse(k) *\<^sub>R (w - setsum (\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 - setsum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" + apply (simp add: weq_sumsum convex_hull_finite fin) + apply (rule_tac x="\i. inverse k * c i" in exI) + using \k > 0\ cge0 + apply (auto simp: scaleR_right.setsum setsum_right_distrib [symmetric] k_def [symmetric]) + done + ultimately show ?thesis + using disj by blast + qed + qed + qed + have [simp]: "convex hull T \ convex hull S" + by (simp add: \T \ S\ hull_mono) + show ?thesis + using open_segment_commute by (auto simp: face_of_def intro: *) +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 + +proposition face_of_affine_trivial: + assumes "affine S" "T face_of S" + shows "T = {} \ T = S" +proof (rule ccontr, clarsimp) + assume "T \ {}" "T \ S" + then obtain a where "a \ T" by auto + then have "a \ S" + using \T face_of S\ face_of_imp_subset by blast + have "S \ T" + proof + fix b assume "b \ S" + show "b \ T" + proof (cases "a = b") + 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) + apply (rule_tac x="1/2" in exI) + apply (simp add: algebra_simps) + by (simp add: scaleR_2) + 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\ + ultimately show ?thesis + by (rule face_ofD [OF \T face_of S\, THEN conjunct2]) + qed + qed + then show False + using `T \ S` \T face_of S\ face_of_imp_subset by blast +qed + + +lemma face_of_affine_eq: + "affine S \ (T face_of S \ T = {} \ T = S)" +using affine_imp_convex face_of_affine_trivial face_of_refl by auto + + +lemma Inter_faces_finite_altbound: + fixes T :: "'a::euclidean_space set set" + assumes cfaI: "\c. c \ T \ c face_of S" + shows "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ \F' = \T" +proof (cases "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ (\c. c \ T \ c \ (\F') \ (\F'))") + case True + then obtain c where c: + "\F'. \finite F'; F' \ T; card F' \ DIM('a) + 2\ \ c F' \ T \ c F' \ (\F') \ (\F')" + by metis + def d \ "rec_nat {c{}} (\n r. insert (c r) r)" + have [simp]: "d 0 = {c {}}" + by (simp add: d_def) + have dSuc [simp]: "\n. d (Suc n) = insert (c (d n)) (d n)" + by (simp add: d_def) + have dn_notempty: "d n \ {}" for n + by (induction n) auto + have dn_le_Suc: "d n \ T \ finite(d n) \ card(d n) \ Suc n" if "n \ DIM('a) + 2" for n + using that + proof (induction n) + case 0 + then show ?case by (simp add: c) + next + case (Suc n) + then show ?case by (auto simp: c card_insert_if) + qed + have aff_dim_le: "aff_dim(\(d n)) \ DIM('a) - int n" if "n \ DIM('a) + 2" for n + using that + proof (induction n) + case 0 + then show ?case + by (simp add: aff_dim_le_DIM) + next + case (Suc n) + have fs: "\d (Suc n) face_of S" + by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE) + have condn: "convex (\d n)" + using Suc.prems nat_le_linear not_less_eq_eq + by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc) + have fdn: "\d (Suc n) face_of \d n" + by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI) + have ne: "\d (Suc n) \ \d n" + by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans) + have *: "\m::int. \d. \d'::int. d < d' \ d' \ m - n \ d \ m - of_nat(n+1)" + by arith + have "aff_dim (\d (Suc n)) < aff_dim (\d n)" + by (rule face_of_aff_dim_lt [OF condn fdn ne]) + moreover have "aff_dim (\d n) \ int (DIM('a)) - int n" + using Suc by auto + ultimately + have "aff_dim (\d (Suc n)) \ int (DIM('a)) - (n+1)" by arith + then show ?case by linarith + qed + have "aff_dim (\d (DIM('a) + 2)) \ -2" + using aff_dim_le [OF order_refl] by simp + with aff_dim_geq [of "\d (DIM('a) + 2)"] show ?thesis + using order.trans by fastforce +next + case False + then show ?thesis + apply simp + apply (erule ex_forward) + by blast +qed + +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 + +proposition face_of_Times: + assumes "F face_of S" and "F' face_of S'" + shows "(F \ F') face_of (S \ S')" +proof - + have "F \ F' \ S \ S'" + using assms [unfolded face_of_def] by blast + moreover + have "convex (F \ F')" + using assms [unfolded face_of_def] by (blast intro: convex_Times) + moreover + have "a \ F \ a' \ F' \ b \ F \ b' \ F'" + if "a \ S" "b \ S" "a' \ S'" "b' \ S'" "x \ F \ F'" "x \ open_segment (a,a') (b,b')" + for a b a' b' x + proof (cases "b=a \ b'=a'") + case True with that show ?thesis + using assms + by (force simp: in_segment dest: face_ofD) + next + case False with assms [unfolded face_of_def] that show ?thesis + by (blast dest!: open_segment_PairD) + qed + ultimately show ?thesis + unfolding face_of_def by blast +qed + +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')" + (is "?lhs = ?rhs") +proof + assume c: ?lhs + show ?rhs + 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 fst_linear snd_linear) + 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 + qed + 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' + 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 + qed + 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\ fst_linear snd_linear 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 + with fst snd show ?thesis by metis + qed +next + assume ?rhs with face_of_Times show ?lhs by auto +qed + +lemma face_of_Times_eq: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + shows "(F \ F') face_of (S \ S') \ + F = {} \ F' = {} \ F face_of S \ F' face_of S'" +by (auto simp: face_of_Times_decomp times_eq_iff) + +lemma hyperplane_face_of_halfspace_le: "{x. a \ x = b} face_of {x. a \ x \ b}" +proof - + have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" + by auto + with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b] + show ?thesis by auto +qed + +lemma hyperplane_face_of_halfspace_ge: "{x. a \ x = b} face_of {x. a \ x \ b}" +proof - + have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" + by auto + with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a] + show ?thesis by auto +qed + +lemma face_of_halfspace_le: + fixes a :: "'n::euclidean_space" + shows "F face_of {x. a \ x \ b} \ + F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" + (is "?lhs = ?rhs") +proof (cases "a = 0") + case True then show ?thesis + using face_of_affine_eq affine_UNIV by auto +next + case False + then have ine: "interior {x. a \ x \ b} \ {}" + using halfspace_eq_empty_lt interior_halfspace_le by blast + 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 + with L show ?rhs + using affine_hyperplane face_of_affine_eq by blast + next + assume ?rhs + then show ?lhs + by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le) + qed +qed + +lemma face_of_halfspace_ge: + fixes a :: "'n::euclidean_space" + shows "F face_of {x. a \ x \ b} \ + F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" +using face_of_halfspace_le [of F "-a" "-b"] by simp + +subsection\Exposed faces\ + +text\That is, faces that are intersection with supporting hyperplane\ + +definition exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" + (infixr "(exposed'_face'_of)" 50) + where "T exposed_face_of S \ + T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" + +lemma empty_exposed_face_of [iff]: "{} exposed_face_of S" + apply (simp add: exposed_face_of_def) + apply (rule_tac x=0 in exI) + apply (rule_tac x=1 in exI, force) + 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 + +lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" + by simp + +lemma exposed_face_of: + "T exposed_face_of S \ + T face_of S \ + (T = {} \ T = S \ + (\a b. a \ 0 \ S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b}))" +proof (cases "T = {}") + case True then show ?thesis + by simp +next + case False + show ?thesis + proof (cases "T = S") + case True then show ?thesis + by (simp add: face_of_refl_eq) + next + case False + with \T \ {}\ show ?thesis + apply (auto simp: exposed_face_of_def) + apply (metis inner_zero_left) + done + qed +qed + +lemma exposed_face_of_Int_supporting_hyperplane_le: + "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" +by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le) + +lemma exposed_face_of_Int_supporting_hyperplane_ge: + "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" +using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp + +proposition exposed_face_of_Int: + assumes "T exposed_face_of S" + and "u exposed_face_of S" + shows "(T \ u) exposed_face_of S" +proof - + obtain a b where T: "S \ {x. a \ x = b} face_of S" + and S: "S \ {x. a \ x \ b}" + and teq: "T = S \ {x. a \ x = b}" + using assms by (auto simp: exposed_face_of_def) + obtain a' b' where u: "S \ {x. a' \ x = b'} face_of S" + and s': "S \ {x. a' \ x \ b'}" + and ueq: "u = S \ {x. a' \ x = b'}" + using assms by (auto simp: exposed_face_of_def) + have tu: "T \ u face_of S" + using T teq u ueq by (simp add: face_of_Int) + have ss: "S \ {x. (a + a') \ x \ b + b'}" + using S s' by (force simp: inner_left_distrib) + show ?thesis + apply (simp add: exposed_face_of_def tu) + apply (rule_tac x="a+a'" in exI) + apply (rule_tac x="b+b'" in exI) + using S s' + apply (fastforce simp: ss inner_left_distrib teq ueq) + done +qed + +proposition exposed_face_of_Inter: + fixes P :: "'a::euclidean_space set set" + assumes "P \ {}" + and "\T. T \ P \ T exposed_face_of S" + shows "\P exposed_face_of S" +proof - + obtain Q where "finite Q" and QsubP: "Q \ P" "card Q \ DIM('a) + 2" and IntQ: "\Q = \P" + using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of] + by force + show ?thesis + proof (cases "Q = {}") + case True then show ?thesis + by (metis Inf_empty Inf_lower IntQ assms ex_in_conv subset_antisym top_greatest) + next + case False + have "Q \ {T. T exposed_face_of S}" + 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 + ultimately show ?thesis + by (simp add: IntQ) + qed +qed + +proposition exposed_face_of_sums: + assumes "convex S" and "convex T" + and "F exposed_face_of {x + y | x y. x \ S \ y \ T}" + (is "F exposed_face_of ?ST") + obtains k l + where "k exposed_face_of S" "l exposed_face_of T" + "F = {x + y | x y. x \ k \ y \ l}" +proof (cases "F = {}") + case True then show ?thesis + using that by blast +next + case False + show ?thesis + proof (cases "F = ?ST") + case True then show ?thesis + using assms exposed_face_of_refl_eq that by blast + next + case False + obtain p where "p \ F" using \F \ {}\ by blast + moreover + obtain u z where T: "?ST \ {x. u \ x = z} face_of ?ST" + and S: "?ST \ {x. u \ x \ z}" + and feq: "F = ?ST \ {x. u \ x = z}" + using assms by (auto simp: exposed_face_of_def) + ultimately obtain a0 b0 + where p: "p = a0 + b0" and "a0 \ S" "b0 \ T" and z: "u \ p = z" + by auto + 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 + 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 + 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 + 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 + by (rule that [OF sef tef]) + qed +qed + +subsection\Extreme points of a set: its singleton faces\ + +definition extreme_point_of :: "['a::real_vector, 'a set] \ bool" + (infixr "(extreme'_point'_of)" 50) + where "x extreme_point_of S \ + x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" + +lemma extreme_point_of_stillconvex: + "convex S \ (x extreme_point_of S \ x \ S \ convex(S - {x}))" + by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def) + +lemma face_of_singleton: + "{x} face_of S \ x extreme_point_of S" +by (fastforce simp add: extreme_point_of_def face_of_def) + +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 + +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) + +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 + +lemma extreme_points_of_convex_hull: + "{x. x extreme_point_of (convex hull S)} \ S" +using extreme_point_of_convex_hull by auto + +lemma extreme_point_of_empty [simp]: "~ (x extreme_point_of {})" + by (simp add: extreme_point_of_def) + +lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" + using extreme_point_of_stillconvex by auto + +lemma extreme_point_of_translation_eq: + "(a + x) extreme_point_of (image (\x. a + x) S) \ x extreme_point_of S" +by (auto simp: extreme_point_of_def) + +lemma extreme_points_of_translation: + "{x. x extreme_point_of (image (\x. a + x) S)} = + (\x. a + x) ` {x. x extreme_point_of S}" +using extreme_point_of_translation_eq +by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel) + +lemma extreme_point_of_Int: + "\x extreme_point_of S; x extreme_point_of T\ \ x extreme_point_of (S \ T)" +by (simp add: extreme_point_of_def) + +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) + +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) + +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 + +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 + +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 + +subsection\Facets\ + +definition facet_of :: "['a::euclidean_space set, 'a set] \ bool" + (infixr "(facet'_of)" 50) + where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" + +lemma facet_of_empty [simp]: "~ S facet_of {}" + by (simp add: facet_of_def) + +lemma facet_of_irrefl [simp]: "~ S facet_of S " + by (simp add: facet_of_def) + +lemma facet_of_imp_face_of: "F facet_of S \ F face_of S" + by (simp add: facet_of_def) + +lemma facet_of_imp_subset: "F facet_of S \ F \ S" + by (simp add: face_of_imp_subset facet_of_def) + +lemma hyperplane_facet_of_halfspace_le: + "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" +unfolding facet_of_def hyperplane_eq_empty +by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le + DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_le) + +lemma hyperplane_facet_of_halfspace_ge: + "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" +unfolding facet_of_def hyperplane_eq_empty +by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge + DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_ge) + +lemma facet_of_halfspace_le: + "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" + (is "?lhs = ?rhs") +proof + assume c: ?lhs + with c facet_of_irrefl show ?rhs + by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm) +next + assume ?rhs then show ?lhs + by (simp add: hyperplane_facet_of_halfspace_le) +qed + +lemma facet_of_halfspace_ge: + "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" +using facet_of_halfspace_le [of F "-a" "-b"] by simp + +subsection \Edges: faces of affine dimension 1\ + +definition edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) + where "e edge_of S \ e face_of S \ aff_dim e = 1" + +lemma edge_of_imp_subset: + "S edge_of T \ S \ T" +by (simp add: edge_of_def face_of_imp_subset) + +subsection\Existence of extreme points\ + +lemma different_norm_3_collinear_points: + fixes a :: "'a::euclidean_space" + assumes "x \ open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)" + shows False +proof - + obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b" + and "a \ b" + and u01: "0 < u" "u < 1" + using assms by (auto simp: open_segment_image_interval if_splits) + then have "(1 - u) *\<^sub>R a \ (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \ u *\<^sub>R b = + (1 - u * u) *\<^sub>R (a \ a)" + using assms by (simp add: norm_eq algebra_simps inner_commute) + then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \ a + (2 * u) *\<^sub>R a \ b) = + (1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \ a))" + by (simp add: algebra_simps) + then have "(1 - u) *\<^sub>R (a \ a) + (2 * u) *\<^sub>R (a \ b) = (1 + u) *\<^sub>R (a \ a)" + using u01 by auto + then have "a \ b = a \ a" + using u01 by (simp add: algebra_simps) + then have "a = b" + using \norm(a) = norm(b)\ norm_eq vector_eq by fastforce + then show ?thesis + using \a \ b\ by force +qed + +proposition extreme_point_exists_convex: + fixes S :: "'a::euclidean_space set" + assumes "compact S" "convex S" "S \ {}" + obtains x where "x extreme_point_of S" +proof - + obtain x where "x \ S" and xsup: "\y. y \ S \ norm y \ norm x" + using distance_attains_sup [of S 0] assms by auto + have False if "a \ S" "b \ S" and x: "x \ open_segment a b" for a b + proof - + 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 + qed + then show ?thesis + apply (rule_tac x=x in that) + apply (force simp: extreme_point_of_def \x \ S\) + done +qed + +subsection\Krein-Milman, the weaker form\ + +proposition Krein_Milman: + fixes S :: "'a::euclidean_space set" + assumes "compact S" "convex S" + shows "S = closure(convex hull {x. x extreme_point_of S})" +proof (cases "S = {}") + case True then show ?thesis by simp +next + case False + 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 + moreover have "u \ closure (convex hull {x. x extreme_point_of S})" + if "u \ S" for u + proof (rule ccontr) + assume unot: "u \ closure(convex hull {x. x extreme_point_of S})" + then obtain a b where "a \ u < b" + and ab: "\x. x \ closure(convex hull {x. x extreme_point_of S}) \ b < a \ x" + using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"] + by blast + have "continuous_on S (op \ a)" + by (rule continuous_intros)+ + then obtain m where "m \ S" and m: "\y. y \ S \ a \ m \ a \ y" + using continuous_attains_inf [of S "\x. a \ x"] \compact S\ \u \ S\ + by auto + def T \ "S \ {x. a \ x = a \ m}" + have "m \ T" + by (simp add: T_def \m \ S\) + moreover have "compact T" + by (simp add: T_def compact_Int_closed [OF \compact S\ closed_hyperplane]) + moreover have "convex T" + by (simp add: T_def convex_Int [OF \convex S\ convex_hyperplane]) + ultimately obtain v where v: "v extreme_point_of T" + using extreme_point_exists_convex [of T] by auto + then have "{v} face_of T" + by (simp add: face_of_singleton) + also have "T face_of S" + by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \convex S\]) + finally have "v extreme_point_of S" + by (simp add: face_of_singleton) + then have "b < a \ v" + using closure_subset by (simp add: closure_hull hull_inc ab) + then show False + using \a \ u < b\ \{v} face_of T\ face_of_imp_subset m T_def that by fastforce + qed + ultimately show ?thesis + by blast +qed + +text\Now the sharper form.\ + +lemma Krein_Milman_Minkowski_aux: + fixes S :: "'a::euclidean_space set" + assumes n: "dim S = n" and S: "compact S" "convex S" "0 \ S" + shows "0 \ convex hull {x. x extreme_point_of S}" +using n S +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) + next + case False + have "rel_interior S \ {}" + by (simp add: rel_interior_convex_nonempty_aux less) + then obtain c where c: "c \ rel_interior S" by blast + obtain a where "a \ 0" + and le_ay: "\y. y \ S \ a \ 0 \ a \ y" + 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 + 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 + proof - + have "y \ span {x. a \ x = 0}" + by (metis inf.cobounded2 span_mono subsetCE that) + then have "y \ {x. a \ x = 0}" + by (rule span_induct [OF _ subspace_hyperplane]) + then show ?thesis by simp + qed + then have "dim (S \ {x. a \ x = 0}) < n" + by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff + inf_le1 \dim S = n\ not_le rel_interior_subset span_0 span_clauses(1)) + 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) + qed +qed + + +theorem Krein_Milman_Minkowski: + fixes S :: "'a::euclidean_space set" + assumes "compact S" "convex S" + shows "S = convex hull {x. x extreme_point_of S}" +proof + show "S \ convex hull {x. x extreme_point_of S}" + proof + fix a assume [simp]: "a \ S" + have 1: "compact (op + (- a) ` S)" + by (simp add: \compact S\ compact_translation) + have 2: "convex (op + (- a) ` S)" + by (simp add: \convex S\ convex_translation) + show a_invex: "a \ convex hull {x. x extreme_point_of S}" + using Krein_Milman_Minkowski_aux [OF refl 1 2] + convex_hull_translation [of "-a"] + by (auto simp: extreme_points_of_translation translation_assoc) + qed +next + show "convex hull {x. x extreme_point_of S} \ S" + proof - + have "{a. a extreme_point_of S} \ S" + using extreme_point_of_def by blast + then show ?thesis + by (simp add: \convex S\ hull_minimal) + qed +qed + + +subsection\Applying it to convex hulls of explicitly indicated finite sets\ + +lemma Krein_Milman_polytope: + fixes S :: "'a::euclidean_space set" + shows + "finite S + \ convex hull S = + convex hull {x. x extreme_point_of (convex hull S)}" +by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull) + +lemma extreme_points_of_convex_hull_eq: + fixes S :: "'a::euclidean_space set" + shows + "\compact S; \T. T \ S \ convex hull T \ convex hull S\ + \ {x. x extreme_point_of (convex hull S)} = S" +by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI) + + +lemma extreme_point_of_convex_hull_eq: + fixes S :: "'a::euclidean_space set" + shows + "\compact S; \T. T \ S \ convex hull T \ convex hull S\ + \ (x extreme_point_of (convex hull S) \ x \ S)" +using extreme_points_of_convex_hull_eq by auto + +lemma extreme_point_of_convex_hull_convex_independent: + fixes S :: "'a::euclidean_space set" + assumes "compact S" and S: "\a. a \ S \ a \ convex hull (S - {a})" + shows "(x extreme_point_of (convex hull S) \ x \ S)" +proof - + have "convex hull T \ convex hull S" if "T \ S" for T + proof - + obtain a where "T \ S" "a \ S" "a \ T" using \T \ S\ by blast + then show ?thesis + by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE) + qed + then show ?thesis + by (rule extreme_point_of_convex_hull_eq [OF \compact S\]) +qed + +lemma extreme_point_of_convex_hull_affine_independent: + fixes S :: "'a::euclidean_space set" + shows + "~ affine_dependent S + \ (x extreme_point_of (convex hull S) \ x \ S)" +by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) + +text\Elementary proofs exist, not requiring Euclidean spaces and all this development\ +lemma extreme_point_of_convex_hull_2: + fixes x :: "'a::euclidean_space" + shows "x extreme_point_of (convex hull {a,b}) \ x = a \ x = b" +proof - + have "x extreme_point_of (convex hull {a,b}) \ x \ {a,b}" + by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2) + then show ?thesis + by simp +qed + +lemma extreme_point_of_segment: + fixes x :: "'a::euclidean_space" + shows + "x extreme_point_of closed_segment a b \ x = a \ x = b" +by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull) + +lemma face_of_convex_hull_subset: + 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) + + +proposition face_of_convex_hull_affine_independent: + fixes S :: "'a::euclidean_space set" + assumes "~ affine_dependent S" + shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (meson \T face_of convex hull S\ aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact) +next + assume ?rhs + 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 + then have "affine hull c \ convex hull (S - c) = {}" + using convex_hull_subset_affine_hull by fastforce + then show ?lhs + by (metis face_of_convex_hulls \c \ S\ aff_independent_finite assms T) +qed + +lemma facet_of_convex_hull_affine_independent: + fixes S :: "'a::euclidean_space set" + assumes "~ affine_dependent S" + shows "T facet_of (convex hull S) \ + T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then have "T face_of (convex hull S)" "T \ {}" + and afft: "aff_dim T = aff_dim (convex hull S) - 1" + by (auto simp: facet_of_def) + then obtain c where "c \ S" and c: "T = convex hull c" + by (auto simp: face_of_convex_hull_affine_independent [OF assms]) + then have affs: "aff_dim S = aff_dim c + 1" + by (metis aff_dim_convex_hull afft eq_diff_eq) + have "~ affine_dependent c" + using \c \ S\ affine_dependent_subset assms by blast + with affs have "card (S - c) = 1" + apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) + by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \c \ S\ add.commute + add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff) + then obtain u where u: "u \ S - c" + by (metis DiffI \c \ S\ aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel + card_Diff_subset subsetI subset_antisym zero_neq_one) + then have u: "S = insert u c" + by (metis Diff_subset \c \ S\ \card (S - c) = 1\ card_1_singletonE double_diff insert_Diff insert_subset singletonD) + have "T = convex hull (c - {u})" + by (metis Diff_empty Diff_insert0 \T facet_of convex hull S\ c facet_of_irrefl insert_absorb u) + with \T \ {}\ show ?rhs + using c u by auto +next + assume ?rhs + then obtain u where "T \ {}" "u \ S" and u: "T = convex hull (S - {u})" + 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" + using assms \u \ S\ + apply (simp add: aff_dim_convex_hull affine_dependent_def) + apply (drule bspec, assumption) + 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 +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 + +lemma segment_face_of: + assumes "(closed_segment a b) face_of S" + shows "a extreme_point_of S" "b extreme_point_of S" +proof - + have as: "{a} face_of S" + by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull) + moreover have "{b} face_of S" + proof - + have "b \ convex hull {a} \ b extreme_point_of convex hull {b, a}" + by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI) + moreover have "closed_segment a b = convex hull {b, a}" + using closed_segment_commute segment_convex_hull by blast + ultimately show ?thesis + by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE) + qed + ultimately show "a extreme_point_of S" "b extreme_point_of S" + using face_of_singleton by blast+ +qed + + +lemma Krein_Milman_frontier: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "compact S" + shows "S = convex hull (frontier S)" + (is "?lhs = ?rhs") +proof + 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 + finally show "?lhs \ ?rhs" . +next + have "?rhs \ convex hull S" + by (metis Diff_subset \compact S\ closure_closed compact_eq_bounded_closed frontier_def hull_mono) + also have "... \ ?lhs" + by (simp add: \convex S\ hull_same) + finally show "?rhs \ ?lhs" . +qed + +subsection\Polytopes\ + +definition polytope where + "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) + +lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" + unfolding polytope_def using convex_hull_linear_image by blast + +lemma polytope_empty: "polytope {}" + using convex_hull_empty polytope_def by blast + +lemma polytope_convex_hull: "finite S \ polytope(convex hull S)" + using polytope_def by auto + +lemma polytope_Times: "\polytope S; polytope T\ \ polytope(S \ T)" + unfolding polytope_def + by (metis finite_cartesian_product convex_hull_Times) + +lemma face_of_polytope_polytope: + fixes S :: "'a::euclidean_space set" + shows "\polytope S; F face_of S\ \ polytope F" +unfolding polytope_def +by (meson face_of_convex_hull_subset finite_imp_compact finite_subset) + +lemma finite_polytope_faces: + fixes S :: "'a::euclidean_space set" + assumes "polytope S" + shows "finite {F. F face_of S}" +proof - + obtain v where "finite v" "S = convex hull v" + using assms polytope_def by auto + have "finite (op hull convex ` {T. T \ v})" + by (simp add: \finite v\) + moreover have "{F. F face_of S} \ (op hull convex ` {T. T \ v})" + by (metis (no_types, lifting) \finite v\ \S = convex hull v\ face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI) + ultimately show ?thesis + by (blast intro: finite_subset) +qed + +lemma finite_polytope_facets: + assumes "polytope S" + shows "finite {T. T facet_of S}" +by (simp add: assms facet_of_def finite_polytope_faces) + +lemma polytope_scaling: + assumes "polytope S" shows "polytope (image (\x. c *\<^sub>R x) S)" +by (simp add: assms polytope_linear_image) + +lemma polytope_imp_compact: + fixes S :: "'a::real_normed_vector set" + shows "polytope S \ compact S" +by (metis finite_imp_compact_convex_hull polytope_def) + +lemma polytope_imp_convex: "polytope S \ convex S" + by (metis convex_convex_hull polytope_def) + +lemma polytope_imp_closed: + fixes S :: "'a::real_normed_vector set" + shows "polytope S \ closed S" +by (simp add: compact_imp_closed polytope_imp_compact) + +lemma polytope_imp_bounded: + fixes S :: "'a::real_normed_vector set" + shows "polytope S \ bounded S" +by (simp add: compact_imp_bounded polytope_imp_compact) + +lemma polytope_interval: "polytope(cbox a b)" + unfolding polytope_def by (meson closed_interval_as_convex_hull) + +lemma polytope_sing: "polytope {a}" + using polytope_def by force + + +subsection\Polyhedra\ + +definition polyhedron where + "polyhedron S \ + \F. finite F \ + S = \ F \ + (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b})" + +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) + done + +lemma polyhedron_UNIV [iff]: "polyhedron UNIV" + unfolding polyhedron_def + by (rule_tac x="{}" in exI) auto + +lemma polyhedron_Inter [intro,simp]: + "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" +by (induction F rule: finite_induct) auto + + +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) + apply (rule_tac x="-1" in exI) + apply (simp add: 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 +qed + +lemma polyhedron_halfspace_le: + fixes a :: "'a :: euclidean_space" + shows "polyhedron {x. a \ x \ b}" +proof (cases "a = 0") + case True then show ?thesis by auto +next + case False + then show ?thesis + unfolding polyhedron_def + by (rule_tac x="{{x. a \ x \ b}}" in exI) auto +qed + +lemma polyhedron_halfspace_ge: + fixes a :: "'a :: euclidean_space" + shows "polyhedron {x. a \ x \ b}" +using polyhedron_halfspace_le [of "-a" "-b"] by simp + +lemma polyhedron_hyperplane: + fixes a :: "'a :: euclidean_space" + shows "polyhedron {x. a \ x = b}" +proof - + have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" + by force + then show ?thesis + by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le) +qed + +lemma affine_imp_polyhedron: + fixes S :: "'a :: euclidean_space set" + shows "affine S \ polyhedron S" +by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S]) + +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 + +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 + +lemma polyhedron_affine_hull: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron(affine hull S)" +by (simp add: affine_imp_polyhedron) + + +subsection\Canonical polyhedron representation making facial structure explicit\ + +lemma polyhedron_Int_affine: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ + (\F. finite F \ S = (affine hull S) \ \F \ + (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}))" + (is "?lhs = ?rhs") +proof + assume ?lhs then show ?rhs + apply (simp add: polyhedron_def) + apply (erule ex_forward) + using hull_subset apply force + done +next + assume ?rhs then show ?lhs + apply clarify + apply (erule ssubst) + apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le) + done +qed + +proposition rel_interior_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 "rel_interior S = {x \ S. \h \ F. a h \ x < b h}" +proof - + have rels: "\x. x \ rel_interior S \ x \ S" + by (meson IntE mem_rel_interior) + moreover have "a i \ x < b i" if x: "x \ rel_interior S" and "i \ F" for x i + proof - + have fif: "F - {i} \ F" + using \i \ F\ Diff_insert_absorb Diff_subset set_insert psubsetI by blast + then have "S \ affine hull S \ \(F - {i})" + by (rule psub) + then obtain z where ssub: "S \ \(F - {i})" and zint: "z \ \(F - {i})" + and "z \ S" and zaff: "z \ affine hull S" + by auto + have "z \ x" + using \z \ S\ rels x by blast + have "z \ affine hull S \ \F" + using \z \ S\ seq by auto + then have aiz: "a i \ z > b i" + using faceq zint zaff by fastforce + obtain e where "e > 0" "x \ S" and e: "ball x e \ affine hull S \ S" + using x by (auto simp: mem_rel_interior_ball) + then have ins: "\y. \norm (x - y) < e; y \ affine hull S\ \ y \ S" + by (metis IntI subsetD dist_norm mem_ball) + def \ \ "min (1/2) (e / 2 / norm(z - x))" + have "norm (\ *\<^sub>R x - \ *\<^sub>R z) = norm (\ *\<^sub>R (x - z))" + by (simp add: \_def algebra_simps norm_mult) + also have "... = \ * norm (x - z)" + using \e > 0\ by (simp add: \_def) + also have "... < e" + using \z \ x\ \e > 0\ by (simp add: \_def min_def divide_simps norm_minus_commute) + finally have lte: "norm (\ *\<^sub>R x - \ *\<^sub>R z) < e" . + 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 + 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 + 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)" + using l by (simp add: algebra_simps aiz) + also have "\ \ b i" using i l + using faceq mem_Collect_eq \i \ F\ by blast + finally have "(a i \ x) * (1 - l) < b i * (1 - l)" + by (simp add: algebra_simps) + with l show ?thesis + by simp + qed + moreover have "x \ rel_interior S" + if "x \ S" and less: "\h. h \ F \ a h \ x < b h" for x + proof - + 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) + show ?thesis + apply (simp add: rel_interior \x \ S\) + apply (rule_tac x="\h\F. interior h" in exI) + apply (auto simp: \finite F\ open_INT 1 2) + done + qed + ultimately show ?thesis by blast +qed + + +lemma polyhedron_Int_affine_parallel: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ + (\F. finite F \ + S = (affine hull S) \ (\F) \ + (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ + (\x \ affine hull S. (x + a) \ affine hull S)))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then 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}" + by (fastforce simp add: polyhedron_Int_affine) + then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" + by metis + show ?rhs + proof - + have "\a' b'. a' \ 0 \ + affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ + (\w \ affine hull S. (w + a') \ affine hull S)" + if "h \ F" "~(affine hull S \ h)" for h + proof - + have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" + using \h \ F\ ab by auto + then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" + by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) + moreover have "~ (affine hull S \ {x. a h \ x \ b h})" + using \h = {x. a h \ x \ b h}\ that(2) by blast + ultimately show ?thesis + using affine_parallel_slice [of "affine hull S"] + by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) + qed + then obtain a b + where ab: "\h. \h \ F; ~ (affine hull S \ h)\ + \ a h \ 0 \ + affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ + (\w \ affine hull S. (w + a h) \ affine hull S)" + by metis + have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" + by (subst seq) (auto simp: ab INT_extend_simps) + show ?thesis + apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ ~(affine hull S \ h)}" in exI) + apply (intro conjI seq2) + using \finite F\ apply force + using ab apply blast + done + qed +next + assume ?rhs then show ?lhs + apply (simp add: polyhedron_Int_affine) + by metis +qed + + +proposition polyhedron_Int_affine_parallel_minimal: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ + (\F. finite F \ + S = (affine hull S) \ (\F) \ + (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ + (\x \ affine hull S. (x + a) \ affine hull S)) \ + (\F'. F' \ F \ S \ (affine hull S) \ (\F')))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain f0 + where f0: "finite f0" + "S = (affine hull S) \ (\f0)" + (is "?P f0") + "\h \ f0. \a b. a \ 0 \ h = {x. a \ x \ b} \ + (\x \ affine hull S. (x + a) \ affine hull S)" + (is "?Q f0") + by (force simp: polyhedron_Int_affine_parallel) + def n \ "LEAST n. \F. card F = n \ finite F \ ?P F \ ?Q F" + have nf: "\F. card F = n \ finite F \ ?P F \ ?Q F" + apply (simp add: n_def) + apply (rule LeastI [where k = "card f0"]) + using f0 apply auto + done + then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" + by blast + then have "~ (finite g \ ?P g \ ?Q g)" if "card g < n" for g + using that by (auto simp: n_def dest!: not_less_Least) + then have *: "~ (?P g \ ?Q g)" if "g \ F" for g + using that \finite F\ psubset_card_mono \card F = n\ + 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) + show ?rhs + by (metis \finite F\ seq aff psubsetI 1 2) +next + assume ?rhs then show ?lhs + by (auto simp: polyhedron_Int_affine_parallel) +qed + + +lemma polyhedron_Int_affine_minimal: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ + (\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 + +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})" +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 + then have "convex S" + by (rule polyhedron_imp_convex) + with False rel_interior_eq_empty have "rel_interior S \ {}" by blast + then obtain x where "x \ rel_interior S" by auto + then obtain T where "open T" "x \ T" "x \ S" "T \ affine hull S \ S" + by (force simp: mem_rel_interior) + then have xaff: "x \ affine hull S" and xint: "x \ \F" + using seq hull_inc by auto + have "rel_interior S = {x \ S. \h\F. a h \ x < b h}" + by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) + with \x \ rel_interior S\ + have [simp]: "\h. h\F \ a h \ x < b h" by blast + have *: "(S \ {x. a h \ x = b h}) facet_of S" if "h \ F" for h + proof - + have "S \ affine hull S \ \(F - {h})" + using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI) + then obtain z where zaff: "z \ affine hull S" and zint: "z \ \(F - {h})" and "z \ S" + by force + then have "z \ x" "z \ h" using seq \x \ S\ by auto + have "x \ h" using that xint by auto + then have able: "a h \ x \ b h" + using faceq that by blast + also have "... < a h \ z" using \z \ h\ faceq [OF that] xint by auto + finally have xltz: "a h \ x < a h \ z" . + def l \ "(b h - a h \ x) / (a h \ z - a h \ x)" + def w \ "(1 - l) *\<^sub>R x + l *\<^sub>R z" + have "0 < l" "l < 1" + using able xltz \b h < a h \ z\ \h \ F\ + by (auto simp: l_def divide_simps) + have awlt: "a i \ w < b i" if "i \ F" "i \ h" for i + proof - + 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 + 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 "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}" + 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) + apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2) + done + next + show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" + proof + fix y + assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" + obtain T where "0 < T" + and T: "\j. \j \ F; j \ h\ \ T * (a j \ y - a j \ w) \ b j - a j \ w" + proof (cases "F - {h} = {}") + case True then show ?thesis + by (rule_tac T=1 in that) auto + next + case False + then obtain h' where h': "h' \ F - {h}" by auto + def inff \ "INF j:F - {h}. if 0 < a j \ y - a j \ w + then (b j - a j \ w) / (a j \ y - a j \ w) + else 1" + have "0 < inff" + apply (simp add: inff_def) + apply (rule finite_imp_less_Inf) + using \finite F\ apply blast + using h' apply blast + apply simp + using awlt apply (force simp: divide_simps) + done + moreover have "inff * (a j \ y - a j \ w) \ b j - a j \ w" + if "j \ F" "j \ h" for j + 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 + then show ?thesis + using \0 < inff\ awlt [OF that] mult_strict_left_mono + by (fastforce simp add: algebra_simps divide_simps split: if_split_asm) + next + case False + with \0 < inff\ have "inff * (a j \ y - a j \ w) \ 0" + by (simp add: mult_le_0_iff) + also have "... < b j - a j \ w" + by (simp add: awlt that) + finally show ?thesis by simp + qed + ultimately show ?thesis + by (blast intro: that) + qed + def 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 + have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a h \ x \ b h}" + using weq yaff by (auto simp: algebra_simps) + with True faceq [OF that] show ?thesis by metis + next + case False + with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a j \ x \ b j}" + by (simp add: algebra_simps) + 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 [unfolded 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})" + 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) + show "y \ affine hull (S \ {y. a h \ y = b h})" + by (metis yeq affine_affine_hull [unfolded 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 + qed + show ?thesis + proof + 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" + by (auto simp: facet_of_def face_of_imp_convex) + then obtain x where x: "x \ rel_interior c" + by (force simp: rel_interior_eq_empty) + 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 + 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 + then have "x \ rel_interior S" + 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 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}" + 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 + + +lemma face_of_polyhedron_subset_explicit: + fixes S :: "'a :: euclidean_space set" + 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'" + 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\ + by (simp add: face_of_imp_subset) + have "polyhedron S" + apply (simp add: polyhedron_Int_affine) + by (metis \finite F\ faceq 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 + 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) + then have "x \ S" + 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 + by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) + then obtain i where "i \ F" "~ (a i \ x < b i)" + using \x \ S\ rels xnot by auto + with xint have "a i \ x = b i" + by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) + 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 +qed + +text\Initial part of proof duplicates that above\ +proposition face_of_polyhedron_explicit: + fixes S :: "'a :: euclidean_space set" + 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'" + 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\ + by (simp add: face_of_imp_subset) + have "polyhedron S" + apply (simp add: polyhedron_Int_affine) + by (metis \finite F\ faceq 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 + 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) + then have "z \ S" + 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) + then obtain e where "0 < e" + "ball z e \ (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" + 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 + 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 + next + 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} = + {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} + \ affine hull S \ \F \ \{?ab h |h. h \ F \ a h \ z = b h}" + if "i \ F" and i: "a i \ z = b i" for i + proof - + have sub: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ j" + if "j \ F" for j + proof - + have "a j \ z \ b j" using faceq that xint by auto + then consider "a j \ z < b j" | "a j \ z = b j" by linarith + then have "\G. G \ {?ab h |h. h \ F \ a h \ z = b h} \ ball z e \ G \ j" + proof cases + assume "a j \ z < b j" + then have "ball z e \ {x. a i \ x = b i} \ j" + using e [OF \j \ F\] faceq that + by (fastforce simp: ball_def) + then show ?thesis + by (rule_tac x="{x. a i \ x = b i}" in exI) (force simp: \i \ F\ i) + next + assume eq: "a j \ z = b j" + with faceq that show ?thesis + by (rule_tac x="{x. a j \ x = b j}" in exI) (fastforce simp add: \j \ F\) + qed + 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 + 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) + have 3: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ \F" + by (iprover intro: sub Inter_greatest) + have *: "\A \ (B :: 'a set); A \ C; E \ C \ D\ \ E \ A \ (B \ D) \ C" + 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 "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 + 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 +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}" +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}" + and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" + 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}" + 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 + 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}" +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}" + and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" + 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 + show ?thesis + apply (subst face_of_polyhedron_explicit [OF \finite F\ seq ab min]) + apply (auto simp: assms facet_of_polyhedron_explicit [OF \finite F\ seq ab min] cong: Collect_cong) + done +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) + + +lemma exposed_face_of_polyhedron: + assumes "polyhedron S" + shows "F exposed_face_of S \ F face_of S" +proof + show "F exposed_face_of S \ F face_of S" + by (simp add: exposed_face_of_def) +next + assume "F face_of S" + show "F exposed_face_of S" + proof (cases "F = {} \ F = S") + case True then show ?thesis + using \F face_of S\ exposed_face_of by blast + next + case False + then have "{g. g facet_of S \ F \ g} \ {}" + 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" + 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 + qed +qed + +lemma face_of_polyhedron_polyhedron: + fixes S :: "'a :: euclidean_space set" + assumes "polyhedron S" "c face_of S" + shows "polyhedron c" +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}" + and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" + 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 + show ?thesis + proof (cases "c = {} \ c = S") + case True with assms show ?thesis + by auto + next + case False + let ?ab = "\h. {x. a h \ x = b h}" + have "{S \ ?ab h |h. h \ F \ c \ S \ ?ab h} \ {S \ ?ab h |h. h \ F}" + by blast + then have fin: "finite ({S \ ?ab h |h. h \ F \ c \ S \ ?ab h})" + by (rule finite_subset) (simp add: \finite F\) + then have "polyhedron (\{S \ ?ab h |h. h \ F \ c \ S \ ?ab h})" + by (auto simp: \polyhedron S\ polyhedron_hyperplane) + with False show ?thesis + using face_of_polyhedron_explicit [OF \finite F\ seq ab min] assms + by auto + qed +qed + + +lemma finite_polyhedron_faces: + fixes S :: "'a :: euclidean_space set" + assumes "polyhedron S" + shows "finite {F. F face_of S}" +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}" + and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" + 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 + have "finite {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" + by (simp add: \finite F\) + moreover have "{F. F face_of S} - {{}, S} \ {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" + 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 + by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset) +qed + +lemma finite_polyhedron_exposed_faces: + "polyhedron S \ finite {F. F exposed_face_of S}" +using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce + +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 + +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]) + + +proposition rel_interior_of_polyhedron: + fixes S :: "'a :: euclidean_space set" + assumes "polyhedron S" + shows "rel_interior S = S - \{F. F facet_of S}" +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}" + and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" + 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 + have facet: "(c facet_of S) \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" for c + by (rule facet_of_polyhedron_explicit [OF \finite F\ seq ab min]) + have rel: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" + by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq ab min]) + have "a h \ x < b h" if "x \ S" "h \ F" and xnot: "x \ \{F. F facet_of S}" for x h + proof - + have "x \ \F" using seq that by force + with \h \ F\ ab have "a h \ x \ b h" by auto + then consider "a h \ x < b h" | "a h \ x = b h" by linarith + then show ?thesis + proof cases + case 1 then show ?thesis . + next + case 2 + have "Collect (op \ x) \ Collect (op \ (\{A. A facet_of S}))" + using xnot by fastforce + then have "F \ Collect (op \ 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 + qed + qed + moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x + using that by (force simp: facet) + ultimately show ?thesis + by (force simp: rel) +qed + +lemma rel_boundary_of_polyhedron: + fixes S :: "'a :: euclidean_space set" + assumes "polyhedron S" + shows "S - rel_interior S = \ {F. F facet_of S}" +using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms) + +lemma rel_frontier_of_polyhedron: + fixes S :: "'a :: euclidean_space set" + assumes "polyhedron S" + shows "rel_frontier S = \ {F. F facet_of S}" +by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron) + +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 + + +text\A characterization of polyhedra as having finitely many faces\ + +proposition polyhedron_eq_finite_exposed_faces: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ closed S \ convex S \ finite {F. F exposed_face_of S}" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces) +next + assume ?rhs + then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto + show ?lhs + proof (cases "S = {}") + case True then show ?thesis by auto + next + case False + def F \ "{h. h exposed_face_of S \ h \ {} \ h \ S}" + have "finite F" by (simp add: fin F_def) + have hface: "h face_of S" + and "\a b. a \ 0 \ S \ {x. a \ x \ b} \ h = S \ {x. a \ x = b}" + if "h \ F" for h + using exposed_face_of F_def that by simp_all auto + then obtain a b where ab: + "\h. h \ F \ a h \ 0 \ S \ {x. a h \ x \ b h} \ h = S \ {x. a h \ x = b h}" + by metis + have *: "False" + if paff: "p \ affine hull S" and "p \ S" and + pint: "p \ \{{x. a h \ x \ b h} |h. h \ F}" for p + proof - + have "rel_interior S \ {}" + by (simp add: \S \ {}\ \convex S\ rel_interior_eq_empty) + then obtain c where c: "c \ rel_interior S" by auto + 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) + 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_tac x="closed_segment c p \ (- S)" in spec) + using rel_interior_subset \closed S\ c \p \ S\ apply blast + done + then obtain \ where "0 \ \" "\ \ 1" and xeq: "x = (1 - \) *\<^sub>R c + \ *\<^sub>R p" + by (auto simp: in_segment) + show False + proof (cases "\=0 \ \=1") + case True with xeq c xnot \x \ S\ \p \ S\ + show False by auto + next + case False + then have xos: "x \ open_segment c p" + using \x \ S\ c open_segment_def that(2) xcl xnot by auto + have xclo: "x \ closure S" + using \x \ S\ closure_subset by blast + obtain d where "d \ 0" + and dle: "\y. y \ closure S \ d \ x \ d \ y" + and dless: "\y. y \ rel_interior S \ d \ x < d \ y" + by (metis supporting_hyperplane_relative_frontier [OF \convex S\ xclo xnot]) + have sex: "S \ {y. d \ y = d \ x} exposed_face_of S" + by (simp add: \closed S\ dle exposed_face_of_Int_supporting_hyperplane_ge [OF \convex S\]) + have sne: "S \ {y. d \ y = d \ x} \ {}" + using \x \ S\ by blast + 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 + 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" + using face_ofD [OF abface xos] \c \ S\ \h \ F\ ab pint \x \ h\ by blast + with c have "h \ rel_interior S \ {}" by blast + then show False + using \h \ F\ F_def face_of_disjoint_rel_interior hface by auto + qed + qed + have "S \ affine hull S \ \{{x. a h \ x \ b h} |h. h \ F}" + using ab by (auto simp: hull_subset) + moreover have "affine hull S \ \{{x. a h \ x \ b h} |h. h \ F} \ S" + using * by blast + ultimately have "S = affine hull S \ \ {{x. a h \ x \ b h} |h. h \ F}" .. + then show ?thesis + apply (rule ssubst) + apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \finite F\) + done + qed +qed + +corollary polyhedron_eq_finite_faces: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ closed S \ convex S \ finite {F. F face_of S}" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex) +next + assume ?rhs + then show ?lhs + by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset) +qed + +lemma polyhedron_linear_image_eq: + fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" + assumes "linear h" "bij h" + shows "polyhedron (h ` S) \ polyhedron S" +proof - + have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P + apply safe + apply (rule_tac x="inv h ` x" in image_eqI) + apply (auto simp: \bij h\ bij_is_surj image_surj_f_inv_f) + done + have "inj h" using bij_is_inj assms by blast + then have injim: "inj_on (op ` h) A" for A + by (simp add: inj_on_def inj_image_eq_iff) + show ?thesis + using \linear h\ \inj h\ + apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq) + apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim) + done +qed + +lemma polyhedron_negations: + fixes S :: "'a :: euclidean_space set" + shows "polyhedron S \ polyhedron(image uminus S)" +by (auto simp: polyhedron_linear_image_eq linear_uminus bij_uminus) + +subsection\Relation between polytopes and polyhedra\ + +lemma polytope_eq_bounded_polyhedron: + fixes S :: "'a :: euclidean_space set" + shows "polytope S \ polyhedron S \ bounded S" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + 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 +qed + +lemma polytope_Int: + fixes S :: "'a :: euclidean_space set" + shows "\polytope S; polytope T\ \ polytope(S \ T)" +by (simp add: polytope_eq_bounded_polyhedron bounded_Int) + + +lemma polytope_Int_polyhedron: + fixes S :: "'a :: euclidean_space set" + shows "\polytope S; polyhedron T\ \ polytope(S \ T)" +by (simp add: bounded_Int polytope_eq_bounded_polyhedron) + +lemma polyhedron_Int_polytope: + fixes S :: "'a :: euclidean_space set" + shows "\polyhedron S; polytope T\ \ polytope(S \ T)" +by (simp add: bounded_Int polytope_eq_bounded_polyhedron) + +lemma polytope_imp_polyhedron: + fixes S :: "'a :: euclidean_space set" + shows "polytope S \ polyhedron S" +by (simp add: polytope_eq_bounded_polyhedron) + +lemma polytope_facet_exists: + fixes p :: "'a :: euclidean_space set" + assumes "polytope p" "0 < aff_dim p" + obtains F where "F facet_of p" +proof (cases "p = {}") + case True with assms show ?thesis by auto +next + case False + then obtain v where "v extreme_point_of p" + using extreme_point_exists_convex + by (blast intro: \polytope p\ polytope_imp_compact polytope_imp_convex) + then + show ?thesis + by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing + all_not_in_conv assms face_of_singleton less_irrefl singletonI that) +qed + +lemma polyhedron_interval [iff]: "polyhedron(cbox a b)" +by (metis polytope_imp_polyhedron polytope_interval) + +lemma polyhedron_convex_hull: + fixes S :: "'a :: euclidean_space set" + shows "finite S \ polyhedron(convex hull S)" +by (simp add: polytope_convex_hull polytope_imp_polyhedron) + + +subsection\Relative and absolute frontier of a polytope\ + +lemma rel_boundary_of_convex_hull: + fixes S :: "'a::euclidean_space set" + assumes "~ affine_dependent S" + shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" +proof - + have "finite S" by (metis assms aff_independent_finite) + then consider "card S = 0" | "card S = 1" | "2 \ card S" by arith + then show ?thesis + proof cases + case 1 then have "S = {}" by (simp add: `finite S`) + then show ?thesis by simp + next + case 2 show ?thesis + by (auto intro: card_1_singletonE [OF `card S = 1`]) + next + case 3 + with assms show ?thesis + by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \finite S\) + qed +qed + +proposition frontier_of_convex_hull: + fixes S :: "'a::euclidean_space set" + assumes "card S = Suc (DIM('a))" + shows "frontier(convex hull S) = \ {convex hull (S - {a}) | a. a \ S}" +proof (cases "affine_dependent S") + case True + have [iff]: "finite S" + using assms using card_infinite by force + 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" + 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}))" + 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 + } note * = this + have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" + apply (subst caratheodory_aff_dim) + apply (blast intro: *) + done + 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 +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}" + 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) + then show ?thesis + by (simp add: False rel_frontier_convex_hull_cases) + qed + finally show ?thesis . +qed + +subsection\Special case of a triangle\ + +proposition frontier_of_triangle: + fixes a :: "'a::euclidean_space" + assumes "DIM('a) = 2" + shows "frontier(convex hull {a,b,c}) = closed_segment a b \ closed_segment b c \ closed_segment c a" + (is "?lhs = ?rhs") +proof (cases "b = a \ c = a \ c = b") + case True then show ?thesis + by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb) +next + case False then have [simp]: "card {a, b, c} = Suc (DIM('a))" + by (simp add: card_insert Set.insert_Diff_if assms) + show ?thesis + proof + show "?lhs \ ?rhs" + using False + by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm) + show "?rhs \ ?lhs" + using False + apply (simp add: frontier_of_convex_hull segment_convex_hull) + apply (intro conjI subsetI) + apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if) + apply (rule_tac X="convex hull {b,c}" in UnionI; force) + apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if) + done + qed +qed + +corollary inside_of_triangle: + fixes a :: "'a::euclidean_space" + assumes "DIM('a) = 2" + shows "inside (closed_segment a b \ closed_segment b c \ closed_segment c a) = interior(convex hull {a,b,c})" +by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull) + +corollary interior_of_triangle: + fixes a :: "'a::euclidean_space" + assumes "DIM('a) = 2" + shows "interior(convex hull {a,b,c}) = + convex hull {a,b,c} - (closed_segment a b \ closed_segment b c \ closed_segment c a)" + using interior_subset + by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int) + +end diff -r 844725394a37 -r e49dc94eb624 src/HOL/ROOT --- a/src/HOL/ROOT Tue May 10 11:56:23 2016 +0100 +++ b/src/HOL/ROOT Tue May 10 14:04:44 2016 +0100 @@ -720,6 +720,7 @@ Multivariate_Analysis Determinants PolyRoots + Polytope Complex_Analysis_Basics Complex_Transcendental Cauchy_Integral_Thm