src/HOL/Analysis/Polytope.thy
changeset 72324 7bb074cceefe
parent 72302 d7d90ed4c74e
child 72569 d56e4eeae967
equal deleted inserted replaced
72309:564012e31db1 72324:7bb074cceefe
    19 
    19 
    20 lemma face_of_translation_eq [simp]:
    20 lemma face_of_translation_eq [simp]:
    21     "((+) a ` T face_of (+) a ` S) \<longleftrightarrow> T face_of S"
    21     "((+) a ` T face_of (+) a ` S) \<longleftrightarrow> T face_of S"
    22 proof -
    22 proof -
    23   have *: "\<And>a T S. T face_of S \<Longrightarrow> ((+) a ` T face_of (+) a ` S)"
    23   have *: "\<And>a T S. T face_of S \<Longrightarrow> ((+) a ` T face_of (+) a ` S)"
    24     apply (simp add: face_of_def Ball_def, clarify)
    24     by (simp add: face_of_def)
    25     by (meson imageI open_segment_translation_eq)
       
    26   show ?thesis
    25   show ?thesis
    27     apply (rule iffI)
    26     by (force simp: image_comp o_def dest: * [where a = "-a"] intro: *)
    28     apply (force simp: image_comp o_def dest: * [where a = "-a"])
       
    29     apply (blast intro: *)
       
    30     done
       
    31 qed
    27 qed
    32 
    28 
    33 lemma face_of_linear_image:
    29 lemma face_of_linear_image:
    34   assumes "linear f" "inj f"
    30   assumes "linear f" "inj f"
    35     shows "(f ` c face_of f ` S) \<longleftrightarrow> c face_of S"
    31     shows "(f ` c face_of f ` S) \<longleftrightarrow> c face_of S"
    81   proof -
    77   proof -
    82     obtain e where "e>0" and e: "cball y e \<inter> affine hull T \<subseteq> T"
    78     obtain e where "e>0" and e: "cball y e \<inter> affine hull T \<subseteq> T"
    83       using y by (auto simp: rel_interior_cball)
    79       using y by (auto simp: rel_interior_cball)
    84     have "y \<noteq> x" "y \<in> S" "y \<in> T"
    80     have "y \<noteq> x" "y \<in> S" "y \<in> T"
    85       using face_of_imp_subset rel_interior_subset T that by blast+
    81       using face_of_imp_subset rel_interior_subset T that by blast+
    86     then have zne: "\<And>u. \<lbrakk>u \<in> {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \<in> T\<rbrakk> \<Longrightarrow>  False"
    82     then have zne: "\<And>u. \<lbrakk>u \<in> {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \<in> T\<rbrakk> \<Longrightarrow> False"
    87       using \<open>x \<in> S\<close> \<open>x \<notin> T\<close> \<open>T face_of S\<close> unfolding face_of_def
    83       using \<open>x \<in> S\<close> \<open>x \<notin> T\<close> \<open>T face_of S\<close> unfolding face_of_def
    88       apply clarify
    84       by (meson greaterThanLessThan_iff in_segment(2))
    89       apply (drule_tac x=x in bspec, assumption)
       
    90       apply (drule_tac x=y in bspec, assumption)
       
    91       apply (subst (asm) open_segment_commute)
       
    92       apply (force simp: open_segment_image_interval image_def)
       
    93       done
       
    94     have in01: "min (1/2) (e / norm (x - y)) \<in> {0<..<1}"
    85     have in01: "min (1/2) (e / norm (x - y)) \<in> {0<..<1}"
    95       using \<open>y \<noteq> x\<close> \<open>e > 0\<close> by simp
    86       using \<open>y \<noteq> x\<close> \<open>e > 0\<close> by simp
    96     show ?thesis
    87     have \<section>: "norm (min (1/2) (e / norm (x - y)) *\<^sub>R y - min (1/2) (e / norm (x - y)) *\<^sub>R x) \<le> e"
    97       apply (rule zne [OF in01])
    88       using \<open>e > 0\<close>
    98       apply (rule e [THEN subsetD])
    89         by (simp add: scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right)
    99       apply (rule IntI)
    90     show False
   100         using \<open>y \<noteq> x\<close> \<open>e > 0\<close>
    91       apply (rule zne [OF in01 e [THEN subsetD]])
   101         apply (simp add: cball_def dist_norm algebra_simps)
    92       using \<open>y \<in> T\<close>
   102         apply (simp add: Real_Vector_Spaces.scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right)
    93         apply (simp add: hull_inc mem_affine x)
   103       apply (rule mem_affine [OF affine_affine_hull _ x])
    94         by (simp add: dist_norm algebra_simps \<section>)
   104       using \<open>y \<in> T\<close>  apply (auto simp: hull_inc)
       
   105       done
       
   106   qed
    95   qed
   107   show ?thesis
    96   show ?thesis
   108     apply (rule subset_antisym)
    97   proof (rule subset_antisym)
   109     using assms apply (simp add: hull_subset face_of_imp_subset)
    98     show "T \<subseteq> affine hull T \<inter> S"
   110     apply (cases "T={}", simp)
    99       using assms by (simp add: hull_subset face_of_imp_subset)
   111     apply (force simp: rel_interior_eq_empty [symmetric] \<open>convex T\<close> intro: *)
   100     show "affine hull T \<inter> S \<subseteq> T"
   112     done
   101       using "*" \<open>convex T\<close> rel_interior_eq_empty by fastforce
       
   102   qed
   113 qed
   103 qed
   114 
   104 
   115 lemma face_of_imp_closed:
   105 lemma face_of_imp_closed:
   116      fixes S :: "'a::euclidean_space set"
   106      fixes S :: "'a::euclidean_space set"
   117      assumes "convex S" "closed S" "T face_of S" shows "closed T"
   107      assumes "convex S" "closed S" "T face_of S" shows "closed T"
   137     then have "b \<le> a \<bullet> u"
   127     then have "b \<le> a \<bullet> u"
   138       using \<open>\<xi> < 1\<close> by auto
   128       using \<open>\<xi> < 1\<close> by auto
   139     with b show "a \<bullet> x \<le> a \<bullet> u" by simp
   129     with b show "a \<bullet> x \<le> a \<bullet> u" by simp
   140   qed
   130   qed
   141   show ?thesis
   131   show ?thesis
   142     apply (simp add: face_of_def assms)
   132     using "*" open_segment_commute by (fastforce simp add: face_of_def assms)
   143     using "*" open_segment_commute by blast
       
   144 qed
   133 qed
   145 
   134 
   146 lemma face_of_Int_supporting_hyperplane_ge_strong:
   135 lemma face_of_Int_supporting_hyperplane_ge_strong:
   147    "\<lbrakk>convex(S \<inter> {x. a \<bullet> x = b}); \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk>
   136    "\<lbrakk>convex(S \<inter> {x. a \<bullet> x = b}); \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk>
   148     \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) face_of S"
   137     \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) face_of S"
   198       by (simp add: divide_simps) (simp add: algebra_simps)
   187       by (simp add: divide_simps) (simp add: algebra_simps)
   199     have "b \<in> open_segment d c"
   188     have "b \<in> open_segment d c"
   200       apply (simp add: open_segment_image_interval)
   189       apply (simp add: open_segment_image_interval)
   201       apply (simp add: d_def algebra_simps image_def)
   190       apply (simp add: d_def algebra_simps image_def)
   202       apply (rule_tac x="e / (e + norm (b - c))" in bexI)
   191       apply (rule_tac x="e / (e + norm (b - c))" in bexI)
   203       using False nbc \<open>0 < e\<close>
   192       using False nbc \<open>0 < e\<close> by (auto simp: algebra_simps)
   204       apply (auto simp: algebra_simps)
       
   205       done
       
   206     then have "d \<in> T \<and> c \<in> T"
   193     then have "d \<in> T \<and> c \<in> T"
   207       apply (rule face_ofD [OF \<open>T face_of S\<close>])
   194       by (meson \<open>b \<in> T\<close> \<open>c \<in> u\<close> \<open>d \<in> u\<close> assms face_ofD subset_iff)
   208       using \<open>d \<in> u\<close>  \<open>c \<in> u\<close> \<open>u \<subseteq> S\<close>  \<open>b \<in> T\<close>  apply auto
       
   209       done
       
   210     then show ?thesis ..
   195     then show ?thesis ..
   211   qed
   196   qed
   212 qed
   197 qed
   213 
   198 
   214 lemma face_of_eq:
   199 lemma face_of_eq:
   215     fixes S :: "'a::real_normed_vector set"
   200     fixes S :: "'a::real_normed_vector set"
   216     assumes "T face_of S" "u face_of S" "(rel_interior T) \<inter> (rel_interior u) \<noteq> {}"
   201     assumes "T face_of S" "U face_of S" "(rel_interior T) \<inter> (rel_interior U) \<noteq> {}"
   217       shows "T = u"
   202     shows "T = U"
   218   apply (rule subset_antisym)
   203   using assms
   219   apply (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subsetCE subset_of_face_of)
   204   unfolding disjoint_iff_not_equal
   220   by (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subset_iff subset_of_face_of)
   205   by (metis IntI empty_iff face_of_imp_subset mem_rel_interior_ball subset_antisym subset_of_face_of)
   221 
   206 
   222 lemma face_of_disjoint_rel_interior:
   207 lemma face_of_disjoint_rel_interior:
   223       fixes S :: "'a::real_normed_vector set"
   208       fixes S :: "'a::real_normed_vector set"
   224       assumes "T face_of S" "T \<noteq> S"
   209       assumes "T face_of S" "T \<noteq> S"
   225         shows "T \<inter> rel_interior S = {}"
   210         shows "T \<inter> rel_interior S = {}"
   268 
   253 
   269 lemma subset_of_face_of_affine_hull:
   254 lemma subset_of_face_of_affine_hull:
   270     fixes S :: "'a::euclidean_space set"
   255     fixes S :: "'a::euclidean_space set"
   271   assumes T: "T face_of S" and "convex S" "U \<subseteq> S" and dis: "\<not> disjnt (affine hull T) (rel_interior U)"
   256   assumes T: "T face_of S" and "convex S" "U \<subseteq> S" and dis: "\<not> disjnt (affine hull T) (rel_interior U)"
   272   shows "U \<subseteq> T"
   257   shows "U \<subseteq> T"
   273   apply (rule subset_of_face_of [OF T \<open>U \<subseteq> S\<close>])
   258 proof (rule subset_of_face_of [OF T \<open>U \<subseteq> S\<close>])
   274   using face_of_imp_eq_affine_Int [OF \<open>convex S\<close> T]
   259   show "T \<inter> rel_interior U \<noteq> {}"
   275   using rel_interior_subset [of U] dis
   260     using face_of_imp_eq_affine_Int [OF \<open>convex S\<close> T] rel_interior_subset [of U] dis \<open>U \<subseteq> S\<close> disjnt_def 
   276   using \<open>U \<subseteq> S\<close> disjnt_def by fastforce
   261     by fastforce
       
   262 qed
   277 
   263 
   278 lemma affine_hull_face_of_disjoint_rel_interior:
   264 lemma affine_hull_face_of_disjoint_rel_interior:
   279     fixes S :: "'a::euclidean_space set"
   265     fixes S :: "'a::euclidean_space set"
   280   assumes "convex S" "F face_of S" "F \<noteq> S"
   266   assumes "convex S" "F face_of S" "F \<noteq> S"
   281   shows "affine hull F \<inter> rel_interior S = {}"
   267   shows "affine hull F \<inter> rel_interior S = {}"
   361           using disj waff by blast
   347           using disj waff by blast
   362       next
   348       next
   363         case False
   349         case False
   364         then have sumcf: "sum c T = 1 - k"
   350         then have sumcf: "sum c T = 1 - k"
   365           by (simp add: S k_def sum_diff sumc1)
   351           by (simp add: S k_def sum_diff sumc1)
       
   352         have ge0: "\<And>x. x \<in> T \<Longrightarrow> 0 \<le> inverse (1 - k) * c x"
       
   353           by (metis \<open>T \<subseteq> S\<close> cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg subsetD sum_nonneg sumcf)
       
   354         have eq1: "(\<Sum>x\<in>T. inverse (1 - k) * c x) = 1"
       
   355           by (metis False eq_iff_diff_eq_0 mult.commute right_inverse sum_distrib_left sumcf)
   366         have "(\<Sum>i\<in>T. c i *\<^sub>R i) /\<^sub>R (1 - k) \<in> convex hull T"
   356         have "(\<Sum>i\<in>T. c i *\<^sub>R i) /\<^sub>R (1 - k) \<in> convex hull T"
   367           apply (simp add: convex_hull_finite fin)
   357           apply (simp add: convex_hull_finite fin)
   368           apply (rule_tac x="\<lambda>i. inverse (1-k) * c i" in exI)
   358           apply (rule_tac x="\<lambda>i. inverse (1-k) * c i" in exI)
   369           apply auto
   359           by (metis (mono_tags, lifting) eq1 ge0 scaleR_scaleR scale_sum_right sum.cong)
   370           apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE)
       
   371           apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf)
       
   372           by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong)
       
   373         with \<open>0 < k\<close>  have "inverse(k) *\<^sub>R (w - sum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
   360         with \<open>0 < k\<close>  have "inverse(k) *\<^sub>R (w - sum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
   374           by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD])
   361           by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD])
   375         moreover have "inverse(k) *\<^sub>R (w - sum (\<lambda>x. c x *\<^sub>R x) T) \<in> convex hull (S - T)"
   362         moreover have "inverse(k) *\<^sub>R (w - sum (\<lambda>x. c x *\<^sub>R x) T) \<in> convex hull (S - T)"
   376           apply (simp add: weq_sumsum convex_hull_finite fin)
   363           apply (simp add: weq_sumsum convex_hull_finite fin)
   377           apply (rule_tac x="\<lambda>i. inverse k * c i" in exI)
   364           apply (rule_tac x="\<lambda>i. inverse k * c i" in exI)
   388   show ?thesis
   375   show ?thesis
   389     using open_segment_commute by (auto simp: face_of_def intro: *)
   376     using open_segment_commute by (auto simp: face_of_def intro: *)
   390 qed
   377 qed
   391 
   378 
   392 proposition face_of_convex_hull_insert:
   379 proposition face_of_convex_hull_insert:
   393    "\<lbrakk>finite S; a \<notin> affine hull S; T face_of convex hull S\<rbrakk> \<Longrightarrow> T face_of convex hull insert a S"
   380   assumes "finite S" "a \<notin> affine hull S" and T: "T face_of convex hull S"
   394   apply (rule face_of_trans, blast)
   381   shows "T face_of convex hull insert a S"
   395   apply (rule face_of_convex_hulls; force simp: insert_Diff_if)
   382 proof -
   396   done
   383   have "convex hull S face_of convex hull insert a S"
       
   384     by (simp add: assms face_of_convex_hulls insert_Diff_if subset_insertI)
       
   385   then show ?thesis
       
   386     using T face_of_trans by blast
       
   387 qed
   397 
   388 
   398 proposition face_of_affine_trivial:
   389 proposition face_of_affine_trivial:
   399     assumes "affine S" "T face_of S"
   390     assumes "affine S" "T face_of S"
   400     shows "T = {} \<or> T = S"
   391     shows "T = {} \<or> T = S"
   401 proof (rule ccontr, clarsimp)
   392 proof (rule ccontr, clarsimp)
   409     show "b \<in> T"
   400     show "b \<in> T"
   410     proof (cases "a = b")
   401     proof (cases "a = b")
   411       case True with \<open>a \<in> T\<close> show ?thesis by auto
   402       case True with \<open>a \<in> T\<close> show ?thesis by auto
   412     next
   403     next
   413       case False
   404       case False
   414       then have "a \<in> open_segment (2 *\<^sub>R a - b) b"
   405       then have "a \<noteq> 2 *\<^sub>R a - b"
   415         apply (auto simp: open_segment_def closed_segment_def)
   406         by (simp add: scaleR_2)
       
   407         with False have "a \<in> open_segment (2 *\<^sub>R a - b) b"
       
   408         apply (clarsimp simp: open_segment_def closed_segment_def)
   416         apply (rule_tac x="1/2" in exI)
   409         apply (rule_tac x="1/2" in exI)
   417         apply (simp add: algebra_simps)
   410           by (simp add: algebra_simps)
   418         by (simp add: scaleR_2)
       
   419       moreover have "2 *\<^sub>R a - b \<in> S"
   411       moreover have "2 *\<^sub>R a - b \<in> S"
   420         by (rule mem_affine [OF \<open>affine S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>, of 2 "-1", simplified])
   412         by (rule mem_affine [OF \<open>affine S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>, of 2 "-1", simplified])
   421       moreover note \<open>b \<in> S\<close> \<open>a \<in> T\<close>
   413       moreover note \<open>b \<in> S\<close> \<open>a \<in> T\<close>
   422       ultimately show ?thesis
   414       ultimately show ?thesis
   423         by (rule face_ofD [OF \<open>T face_of S\<close>, THEN conjunct2])
   415         by (rule face_ofD [OF \<open>T face_of S\<close>, THEN conjunct2])
   497     by blast
   489     by blast
   498 qed
   490 qed
   499 
   491 
   500 lemma faces_of_translation:
   492 lemma faces_of_translation:
   501    "{F. F face_of image (\<lambda>x. a + x) S} = image (image (\<lambda>x. a + x)) {F. F face_of S}"
   493    "{F. F face_of image (\<lambda>x. a + x) S} = image (image (\<lambda>x. a + x)) {F. F face_of S}"
   502 apply (rule subset_antisym, clarify)
   494 proof -
   503 apply (auto simp: image_iff)
   495   have "\<And>F. F face_of (+) a ` S \<Longrightarrow> \<exists>G. G face_of S \<and> F = (+) a ` G"
   504 apply (metis face_of_imp_subset face_of_translation_eq subset_imageE)
   496     by (metis face_of_imp_subset face_of_translation_eq subset_imageE)
   505 done
   497   then show ?thesis
       
   498     by (auto simp: image_iff)
       
   499 qed
   506 
   500 
   507 proposition face_of_Times:
   501 proposition face_of_Times:
   508   assumes "F face_of S" and "F' face_of S'"
   502   assumes "F face_of S" and "F' face_of S'"
   509     shows "(F \<times> F') face_of (S \<times> S')"
   503     shows "(F \<times> F') face_of (S \<times> S')"
   510 proof -
   504 proof -
   529     unfolding face_of_def by blast
   523     unfolding face_of_def by blast
   530 qed
   524 qed
   531 
   525 
   532 corollary face_of_Times_decomp:
   526 corollary face_of_Times_decomp:
   533     fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
   527     fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
   534     shows "c face_of (S \<times> S') \<longleftrightarrow> (\<exists>F F'. F face_of S \<and> F' face_of S' \<and> c = F \<times> F')"
   528     shows "C face_of (S \<times> S') \<longleftrightarrow> (\<exists>F F'. F face_of S \<and> F' face_of S' \<and> C = F \<times> F')"
   535      (is "?lhs = ?rhs")
   529      (is "?lhs = ?rhs")
   536 proof
   530 proof
   537   assume c: ?lhs
   531   assume C: ?lhs
   538   show ?rhs
   532   show ?rhs
   539   proof (cases "c = {}")
   533   proof (cases "C = {}")
   540     case True then show ?thesis by auto
   534     case True then show ?thesis by auto
   541   next
   535   next
   542     case False
   536     case False
   543     have 1: "fst ` c \<subseteq> S" "snd ` c \<subseteq> S'"
   537     have 1: "fst ` C \<subseteq> S" "snd ` C \<subseteq> S'"
   544       using c face_of_imp_subset by fastforce+
   538       using C face_of_imp_subset by fastforce+
   545     have "convex c"
   539     have "convex C"
   546       using c by (metis face_of_imp_convex)
   540       using C by (metis face_of_imp_convex)
   547     have conv: "convex (fst ` c)" "convex (snd ` c)"
   541     have conv: "convex (fst ` C)" "convex (snd ` C)"
   548       by (simp_all add: \<open>convex c\<close> convex_linear_image linear_fst linear_snd)
   542       by (simp_all add: \<open>convex C\<close> convex_linear_image linear_fst linear_snd)
   549     have fstab: "a \<in> fst ` c \<and> b \<in> fst ` c"
   543     have fstab: "a \<in> fst ` C \<and> b \<in> fst ` C"
   550             if "a \<in> S" "b \<in> S" "x \<in> open_segment a b" "(x,x') \<in> c" for a b x x'
   544             if "a \<in> S" "b \<in> S" "x \<in> open_segment a b" "(x,x') \<in> C" for a b x x'
   551     proof -
   545     proof -
   552       have *: "(x,x') \<in> open_segment (a,x') (b,x')"
   546       have *: "(x,x') \<in> open_segment (a,x') (b,x')"
   553         using that by (auto simp: in_segment)
   547         using that by (auto simp: in_segment)
   554       show ?thesis
   548       show ?thesis
   555         using face_ofD [OF c *] that face_of_imp_subset [OF c] by force
   549         using face_ofD [OF C *] that face_of_imp_subset [OF C] by force
   556     qed
   550     qed
   557     have fst: "fst ` c face_of S"
   551     have fst: "fst ` C face_of S"
   558       by (force simp: face_of_def 1 conv fstab)
   552       by (force simp: face_of_def 1 conv fstab)
   559     have sndab: "a' \<in> snd ` c \<and> b' \<in> snd ` c"
   553     have sndab: "a' \<in> snd ` C \<and> b' \<in> snd ` C"
   560             if "a' \<in> S'" "b' \<in> S'" "x' \<in> open_segment a' b'" "(x,x') \<in> c" for a' b' x x'
   554       if "a' \<in> S'" "b' \<in> S'" "x' \<in> open_segment a' b'" "(x,x') \<in> C" for a' b' x x'
   561     proof -
   555     proof -
   562       have *: "(x,x') \<in> open_segment (x,a') (x,b')"
   556       have *: "(x,x') \<in> open_segment (x,a') (x,b')"
   563         using that by (auto simp: in_segment)
   557         using that by (auto simp: in_segment)
   564       show ?thesis
   558       show ?thesis
   565         using face_ofD [OF c *] that face_of_imp_subset [OF c] by force
   559         using face_ofD [OF C *] that face_of_imp_subset [OF C] by force
   566     qed
   560     qed
   567     have snd: "snd ` c face_of S'"
   561     have snd: "snd ` C face_of S'"
   568       by (force simp: face_of_def 1 conv sndab)
   562       by (force simp: face_of_def 1 conv sndab)
   569     have cc: "rel_interior c \<subseteq> rel_interior (fst ` c) \<times> rel_interior (snd ` c)"
   563     have cc: "rel_interior C \<subseteq> rel_interior (fst ` C) \<times> rel_interior (snd ` C)"
   570       by (force simp: face_of_Times rel_interior_Times conv fst snd \<open>convex c\<close> linear_fst linear_snd rel_interior_convex_linear_image [symmetric])
   564       by (force simp: face_of_Times rel_interior_Times conv fst snd \<open>convex C\<close> linear_fst linear_snd rel_interior_convex_linear_image [symmetric])
   571     have "c = fst ` c \<times> snd ` c"
   565     have "C = fst ` C \<times> snd ` C"
   572       apply (rule face_of_eq [OF c])
   566     proof (rule face_of_eq [OF C])
   573       apply (simp_all add: face_of_Times rel_interior_Times conv fst snd)
   567       show "fst ` C \<times> snd ` C face_of S \<times> S'"
   574       using False rel_interior_eq_empty \<open>convex c\<close> cc
   568         by (simp add: face_of_Times rel_interior_Times conv fst snd)
   575       apply blast
   569       show "rel_interior C \<inter> rel_interior (fst ` C \<times> snd ` C) \<noteq> {}"
   576       done
   570         using False rel_interior_eq_empty \<open>convex C\<close> cc
       
   571         by (auto simp: face_of_Times rel_interior_Times conv fst)
       
   572     qed
   577     with fst snd show ?thesis by metis
   573     with fst snd show ?thesis by metis
   578   qed
   574   qed
   579 next
   575 next
   580   assume ?rhs with face_of_Times show ?lhs by auto
   576   assume ?rhs with face_of_Times show ?lhs by auto
   581 qed
   577 qed
   615   then have ine: "interior {x. a \<bullet> x \<le> b} \<noteq> {}"
   611   then have ine: "interior {x. a \<bullet> x \<le> b} \<noteq> {}"
   616     using halfspace_eq_empty_lt interior_halfspace_le by blast
   612     using halfspace_eq_empty_lt interior_halfspace_le by blast
   617   show ?thesis
   613   show ?thesis
   618   proof
   614   proof
   619     assume L: ?lhs
   615     assume L: ?lhs
   620     have "F \<noteq> {x. a \<bullet> x \<le> b} \<Longrightarrow> F face_of {x. a \<bullet> x = b}"
   616     have "F face_of {x. a \<bullet> x = b}" if "F \<noteq> {x. a \<bullet> x \<le> b}"
   621       using False
   617     proof -
   622       apply (simp add: frontier_halfspace_le [symmetric] rel_frontier_nonempty_interior [OF ine, symmetric])
   618       have "F face_of rel_frontier {x. a \<bullet> x \<le> b}"
   623       apply (rule face_of_subset [OF L])
   619       proof (rule face_of_subset [OF L])
   624       apply (simp add: face_of_subset_rel_frontier [OF L])
   620         show "F \<subseteq> rel_frontier {x. a \<bullet> x \<le> b}"
   625       apply (force simp: rel_frontier_def closed_halfspace_le)
   621           by (simp add: L face_of_subset_rel_frontier that)
   626       done
   622       qed (force simp: rel_frontier_def closed_halfspace_le)
       
   623       then show ?thesis
       
   624         using False 
       
   625         by (simp add: frontier_halfspace_le rel_frontier_nonempty_interior [OF ine])
       
   626     qed
   627     with L show ?rhs
   627     with L show ?rhs
   628       using affine_hyperplane face_of_affine_eq by blast
   628       using affine_hyperplane face_of_affine_eq by blast
   629   next
   629   next
   630     assume ?rhs
   630     assume ?rhs
   631     then show ?lhs
   631     then show ?lhs
   653   apply (rule_tac x=0 in exI)
   653   apply (rule_tac x=0 in exI)
   654   apply (rule_tac x=1 in exI, force)
   654   apply (rule_tac x=1 in exI, force)
   655   done
   655   done
   656 
   656 
   657 lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \<longleftrightarrow> convex S"
   657 lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \<longleftrightarrow> convex S"
   658   apply (simp add: exposed_face_of_def face_of_refl_eq, auto)
   658 proof
   659   apply (rule_tac x=0 in exI)+
   659   assume S: "convex S"
   660   apply force
   660   have "S \<subseteq> {x. 0 \<bullet> x \<le> 0} \<and> S = S \<inter> {x. 0 \<bullet> x = 0}"
   661   done
   661     by auto
       
   662   with S show "S exposed_face_of S"
       
   663     using exposed_face_of_def face_of_refl_eq by blast
       
   664 qed (simp add: exposed_face_of_def face_of_refl_eq)
   662 
   665 
   663 lemma exposed_face_of_refl: "convex S \<Longrightarrow> S exposed_face_of S"
   666 lemma exposed_face_of_refl: "convex S \<Longrightarrow> S exposed_face_of S"
   664   by simp
   667   by simp
   665 
   668 
   666 lemma exposed_face_of:
   669 lemma exposed_face_of:
   737     case False
   740     case False
   738     have "Q \<subseteq> {T. T exposed_face_of S}"
   741     have "Q \<subseteq> {T. T exposed_face_of S}"
   739       using QsubP assms by blast
   742       using QsubP assms by blast
   740     moreover have "Q \<subseteq> {T. T exposed_face_of S} \<Longrightarrow> \<Inter>Q exposed_face_of S"
   743     moreover have "Q \<subseteq> {T. T exposed_face_of S} \<Longrightarrow> \<Inter>Q exposed_face_of S"
   741       using \<open>finite Q\<close> False
   744       using \<open>finite Q\<close> False
   742       apply (induction Q rule: finite_induct)
   745       by (induction Q rule: finite_induct; use exposed_face_of_Int in fastforce)
   743       using exposed_face_of_Int apply fastforce+
       
   744       done
       
   745     ultimately show ?thesis
   746     ultimately show ?thesis
   746       by (simp add: IntQ)
   747       by (simp add: IntQ)
   747   qed
   748   qed
   748 qed
   749 qed
   749 
   750 
   775             where p: "p = a0 + b0" and "a0 \<in> S" "b0 \<in> T" and z: "u \<bullet> p = z"
   776             where p: "p = a0 + b0" and "a0 \<in> S" "b0 \<in> T" and z: "u \<bullet> p = z"
   776       by auto
   777       by auto
   777     have lez: "u \<bullet> (x + y) \<le> z" if "x \<in> S" "y \<in> T" for x y
   778     have lez: "u \<bullet> (x + y) \<le> z" if "x \<in> S" "y \<in> T" for x y
   778       using S that by auto
   779       using S that by auto
   779     have sef: "S \<inter> {x. u \<bullet> x = u \<bullet> a0} exposed_face_of S"
   780     have sef: "S \<inter> {x. u \<bullet> x = u \<bullet> a0} exposed_face_of S"
   780       apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex S\<close>])
   781     proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex S\<close>])
   781       apply (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \<open>b0 \<in> T\<close>])
   782       show "\<And>x. x \<in> S \<Longrightarrow> u \<bullet> x \<le> u \<bullet> a0"
   782       done
   783         by (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \<open>b0 \<in> T\<close>])
       
   784     qed
   783     have tef: "T \<inter> {x. u \<bullet> x = u \<bullet> b0} exposed_face_of T"
   785     have tef: "T \<inter> {x. u \<bullet> x = u \<bullet> b0} exposed_face_of T"
   784       apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex T\<close>])
   786     proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex T\<close>])
   785       apply (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \<open>a0 \<in> S\<close>])
   787       show "\<And>x. x \<in> T \<Longrightarrow> u \<bullet> x \<le> u \<bullet> b0"
   786       done
   788         by (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \<open>a0 \<in> S\<close>])
       
   789     qed
   787     have "{x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0} \<subseteq> F"
   790     have "{x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0} \<subseteq> F"
   788       by (auto simp: feq) (metis inner_right_distrib p z)
   791       by (auto simp: feq) (metis inner_right_distrib p z)
   789     moreover have "F \<subseteq> {x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0}"
   792     moreover have "F \<subseteq> {x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0}"
   790       apply (auto simp: feq)
   793     proof -
   791       apply (rename_tac x y)
   794       have "\<And>x y. \<lbrakk>z = u \<bullet> (x + y); x \<in> S; y \<in> T\<rbrakk>
   792       apply (rule_tac x=x in exI)
   795            \<Longrightarrow> u \<bullet> x = u \<bullet> a0 \<and> u \<bullet> y = u \<bullet> b0"
   793       apply (rule_tac x=y in exI, simp)
   796         using z p \<open>a0 \<in> S\<close> \<open>b0 \<in> T\<close>
   794       using z p \<open>a0 \<in> S\<close> \<open>b0 \<in> T\<close>
   797         apply (simp add: inner_right_distrib)
   795       apply clarify
   798         apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute)
   796       apply (simp add: inner_right_distrib)
   799         done
   797       apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute)
   800       then show ?thesis
   798       done
   801         using feq by blast
       
   802     qed
   799     ultimately have "F = {x + y |x y. x \<in> S \<inter> {x. u \<bullet> x = u \<bullet> a0} \<and> y \<in> T \<inter> {x. u \<bullet> x = u \<bullet> b0}}"
   803     ultimately have "F = {x + y |x y. x \<in> S \<inter> {x. u \<bullet> x = u \<bullet> a0} \<and> y \<in> T \<inter> {x. u \<bullet> x = u \<bullet> b0}}"
   800       by blast
   804       by blast
   801     then show ?thesis
   805     then show ?thesis
   802       by (rule that [OF sef tef])
   806       by (rule that [OF sef tef])
   803   qed
   807   qed
   846     proof (intro conjI impI allI ballI exI)
   850     proof (intro conjI impI allI ballI exI)
   847       have *: "S \<subseteq> - (affine hull S \<inter> {x. P x}) \<union> affine hull S \<inter> {x. Q x} \<Longrightarrow> S \<subseteq> {x. \<not> P x \<or> Q x}" 
   851       have *: "S \<subseteq> - (affine hull S \<inter> {x. P x}) \<union> affine hull S \<inter> {x. Q x} \<Longrightarrow> S \<subseteq> {x. \<not> P x \<or> Q x}" 
   848         for P Q 
   852         for P Q 
   849         using hull_subset by fastforce  
   853         using hull_subset by fastforce  
   850       have "S \<subseteq> {x. \<not> (a' \<bullet> x \<le> b') \<or> a' \<bullet> x = b'}"
   854       have "S \<subseteq> {x. \<not> (a' \<bullet> x \<le> b') \<or> a' \<bullet> x = b'}"
   851         apply (rule *)
   855         by (rule *) (use le eq Ssub in auto)
   852         apply (simp only: le eq)
       
   853         using Ssub by auto
       
   854       then show "S \<subseteq> {x. - a' \<bullet> x \<le> - b'}"
   856       then show "S \<subseteq> {x. - a' \<bullet> x \<le> - b'}"
   855         by auto 
   857         by auto 
   856       show "S \<inter> {x. a \<bullet> x = b} = S \<inter> {x. - a' \<bullet> x = - b'}"
   858       show "S \<inter> {x. a \<bullet> x = b} = S \<inter> {x. - a' \<bullet> x = - b'}"
   857         using eq hull_subset [of S affine] by force
   859         using eq hull_subset [of S affine] by force
   858       show "\<lbrakk>S \<inter> {x. a \<bullet> x = b} \<noteq> {}; S \<inter> {x. a \<bullet> x = b} \<noteq> S\<rbrakk> \<Longrightarrow> - a' \<noteq> 0"
   860       show "\<lbrakk>S \<inter> {x. a \<bullet> x = b} \<noteq> {}; S \<inter> {x. a \<bullet> x = b} \<noteq> S\<rbrakk> \<Longrightarrow> - a' \<noteq> 0"
   888 by (fastforce simp add: extreme_point_of_def face_of_def)
   890 by (fastforce simp add: extreme_point_of_def face_of_def)
   889 
   891 
   890 lemma extreme_point_not_in_REL_INTERIOR:
   892 lemma extreme_point_not_in_REL_INTERIOR:
   891     fixes S :: "'a::real_normed_vector set"
   893     fixes S :: "'a::real_normed_vector set"
   892     shows "\<lbrakk>x extreme_point_of S; S \<noteq> {x}\<rbrakk> \<Longrightarrow> x \<notin> rel_interior S"
   894     shows "\<lbrakk>x extreme_point_of S; S \<noteq> {x}\<rbrakk> \<Longrightarrow> x \<notin> rel_interior S"
   893 apply (simp add: face_of_singleton [symmetric])
   895   by (metis disjoint_iff face_of_disjoint_rel_interior face_of_singleton insertI1)
   894 apply (blast dest: face_of_disjoint_rel_interior)
       
   895 done
       
   896 
   896 
   897 lemma extreme_point_not_in_interior:
   897 lemma extreme_point_not_in_interior:
   898     fixes S :: "'a::{real_normed_vector, perfect_space} set"
   898   fixes S :: "'a::{real_normed_vector, perfect_space} set"
   899     shows "x extreme_point_of S \<Longrightarrow> x \<notin> interior S"
   899   assumes "x extreme_point_of S" shows "x \<notin> interior S"
   900 apply (case_tac "S = {x}")
   900 proof (cases "S = {x}")
   901 apply (simp add: empty_interior_finite)
   901   case False
   902 by (meson contra_subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior)
   902   then show ?thesis
       
   903     by (meson assms subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior)
       
   904 qed (simp add: empty_interior_finite)
   903 
   905 
   904 lemma extreme_point_of_face:
   906 lemma extreme_point_of_face:
   905      "F face_of S \<Longrightarrow> v extreme_point_of F \<longleftrightarrow> v extreme_point_of S \<and> v \<in> F"
   907      "F face_of S \<Longrightarrow> v extreme_point_of F \<longleftrightarrow> v extreme_point_of S \<and> v \<in> F"
   906   by (meson empty_subsetI face_of_face face_of_singleton insert_subset)
   908   by (meson empty_subsetI face_of_face face_of_singleton insert_subset)
   907 
   909 
   908 lemma extreme_point_of_convex_hull:
   910 lemma extreme_point_of_convex_hull:
   909    "x extreme_point_of (convex hull S) \<Longrightarrow> x \<in> S"
   911   "x extreme_point_of (convex hull S) \<Longrightarrow> x \<in> S"
   910 apply (simp add: extreme_point_of_stillconvex)
   912   using hull_minimal [of S "(convex hull S) - {x}" convex]
   911 using hull_minimal [of S "(convex hull S) - {x}" convex]
   913   using hull_subset [of S convex]
   912 using hull_subset [of S convex]
   914   by (force simp add: extreme_point_of_stillconvex)
   913 apply blast
       
   914 done
       
   915 
   915 
   916 proposition extreme_points_of_convex_hull:
   916 proposition extreme_points_of_convex_hull:
   917    "{x. x extreme_point_of (convex hull S)} \<subseteq> S"
   917    "{x. x extreme_point_of (convex hull S)} \<subseteq> S"
   918   using extreme_point_of_convex_hull by auto
   918   using extreme_point_of_convex_hull by auto
   919 
   919 
   943    "\<lbrakk>x extreme_point_of S; x extreme_point_of T\<rbrakk> \<Longrightarrow> x extreme_point_of (S \<inter> T)"
   943    "\<lbrakk>x extreme_point_of S; x extreme_point_of T\<rbrakk> \<Longrightarrow> x extreme_point_of (S \<inter> T)"
   944 by (simp add: extreme_point_of_def)
   944 by (simp add: extreme_point_of_def)
   945 
   945 
   946 lemma extreme_point_of_Int_supporting_hyperplane_le:
   946 lemma extreme_point_of_Int_supporting_hyperplane_le:
   947    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
   947    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
   948 apply (simp add: face_of_singleton [symmetric])
   948   by (metis convex_singleton face_of_Int_supporting_hyperplane_le_strong face_of_singleton)
   949 by (metis face_of_Int_supporting_hyperplane_le_strong convex_singleton)
       
   950 
   949 
   951 lemma extreme_point_of_Int_supporting_hyperplane_ge:
   950 lemma extreme_point_of_Int_supporting_hyperplane_ge:
   952    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
   951    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
   953 apply (simp add: face_of_singleton [symmetric])
   952   using extreme_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c]
   954 by (metis face_of_Int_supporting_hyperplane_ge_strong convex_singleton)
   953   by simp
   955 
   954 
   956 lemma exposed_point_of_Int_supporting_hyperplane_le:
   955 lemma exposed_point_of_Int_supporting_hyperplane_le:
   957    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
   956    "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
   958 apply (simp add: exposed_face_of_def face_of_singleton)
   957   unfolding exposed_face_of_def
   959 apply (force simp: extreme_point_of_Int_supporting_hyperplane_le)
   958   by (force simp: face_of_singleton extreme_point_of_Int_supporting_hyperplane_le)
   960 done
       
   961 
   959 
   962 lemma exposed_point_of_Int_supporting_hyperplane_ge:
   960 lemma exposed_point_of_Int_supporting_hyperplane_ge:
   963     "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
   961   "\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
   964 using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c]
   962   using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c]
   965 by simp
   963   by simp
   966 
   964 
   967 lemma extreme_point_of_convex_hull_insert:
   965 lemma extreme_point_of_convex_hull_insert:
   968    "\<lbrakk>finite S; a \<notin> convex hull S\<rbrakk> \<Longrightarrow> a extreme_point_of (convex hull (insert a S))"
   966   assumes "finite S" "a \<notin> convex hull S"
   969 apply (case_tac "a \<in> S")
   967   shows "a extreme_point_of (convex hull (insert a S))"
   970 apply (simp add: hull_inc)
   968 proof (cases "a \<in> S")
   971 using face_of_convex_hulls [of "insert a S" "{a}"]
   969   case False
   972 apply (auto simp: face_of_singleton hull_same)
   970   then show ?thesis
   973 done
   971    using face_of_convex_hulls [of "insert a S" "{a}"] assms
       
   972    by (auto simp: face_of_singleton hull_same)
       
   973 qed (use assms  in \<open>simp add: hull_inc\<close>)
   974 
   974 
   975 subsection\<open>Facets\<close>
   975 subsection\<open>Facets\<close>
   976 
   976 
   977 definition\<^marker>\<open>tag important\<close> facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
   977 definition\<^marker>\<open>tag important\<close> facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
   978                     (infixr "(facet'_of)" 50)
   978                     (infixr "(facet'_of)" 50)
  1064   have False if "a \<in> S" "b \<in> S" and x: "x \<in> open_segment a b" for a b
  1064   have False if "a \<in> S" "b \<in> S" and x: "x \<in> open_segment a b" for a b
  1065   proof -
  1065   proof -
  1066     have noax: "norm a \<le> norm x" and nobx: "norm b \<le> norm x" using xsup that by auto
  1066     have noax: "norm a \<le> norm x" and nobx: "norm b \<le> norm x" using xsup that by auto
  1067     have "a \<noteq> b"
  1067     have "a \<noteq> b"
  1068       using empty_iff open_segment_idem x by auto
  1068       using empty_iff open_segment_idem x by auto
  1069     have *: "(1 - u) * na + u * nb < norm x" if "na < norm x"  "nb \<le> norm x" "0 < u" "u < 1" for na nb u
  1069     show False
  1070     proof -
  1070       by (metis dist_0_norm dist_decreases_open_segment noax nobx not_le x)
  1071       have "(1 - u) * na + u * nb < (1 - u) * norm x + u * nb"
       
  1072         by (simp add: that)
       
  1073       also have "... \<le> (1 - u) * norm x + u * norm x"
       
  1074         by (simp add: that)
       
  1075       finally have "(1 - u) * na + u * nb < (1 - u) * norm x + u * norm x" .
       
  1076       then show ?thesis
       
  1077       using scaleR_collapse [symmetric, of "norm x" u] by auto
       
  1078     qed
       
  1079     have "norm x < norm x" if "norm a < norm x"
       
  1080       using x
       
  1081       apply (clarsimp simp only: open_segment_image_interval \<open>a \<noteq> b\<close> if_False)
       
  1082       apply (rule norm_triangle_lt)
       
  1083       apply (simp add: norm_mult)
       
  1084       using * [of "norm a" "norm b"] nobx that
       
  1085         apply blast
       
  1086       done
       
  1087     moreover have "norm x < norm x" if "norm b < norm x"
       
  1088       using x
       
  1089       apply (clarsimp simp only: open_segment_image_interval \<open>a \<noteq> b\<close> if_False)
       
  1090       apply (rule norm_triangle_lt)
       
  1091       apply (simp add: norm_mult)
       
  1092       using * [of "norm b" "norm a" "1-u" for u] noax that
       
  1093         apply (simp add: add.commute)
       
  1094       done
       
  1095     ultimately have "\<not> (norm a < norm x) \<and> \<not> (norm b < norm x)"
       
  1096       by auto
       
  1097     then show ?thesis
       
  1098       using different_norm_3_collinear_points noax nobx that(3) by fastforce
       
  1099   qed
  1071   qed
  1100   then show ?thesis
  1072   then show ?thesis
  1101     apply (rule_tac x=x in that)
  1073     by (meson \<open>x \<in> S\<close> extreme_point_of_def that)
  1102     apply (force simp: extreme_point_of_def \<open>x \<in> S\<close>)
       
  1103     done
       
  1104 qed
  1074 qed
  1105 
  1075 
  1106 subsection\<open>Krein-Milman, the weaker form\<close>
  1076 subsection\<open>Krein-Milman, the weaker form\<close>
  1107 
  1077 
  1108 proposition Krein_Milman:
  1078 proposition Krein_Milman:
  1114 next
  1084 next
  1115   case False
  1085   case False
  1116   have "closed S"
  1086   have "closed S"
  1117     by (simp add: \<open>compact S\<close> compact_imp_closed)
  1087     by (simp add: \<open>compact S\<close> compact_imp_closed)
  1118   have "closure (convex hull {x. x extreme_point_of S}) \<subseteq> S"
  1088   have "closure (convex hull {x. x extreme_point_of S}) \<subseteq> S"
  1119     apply (rule closure_minimal [OF hull_minimal \<open>closed S\<close>])
  1089     by (simp add: \<open>closed S\<close> assms closure_minimal extreme_point_of_def hull_minimal)
  1120     using assms
       
  1121     apply (auto simp: extreme_point_of_def)
       
  1122     done
       
  1123   moreover have "u \<in> closure (convex hull {x. x extreme_point_of S})"
  1090   moreover have "u \<in> closure (convex hull {x. x extreme_point_of S})"
  1124                 if "u \<in> S" for u
  1091                 if "u \<in> S" for u
  1125   proof (rule ccontr)
  1092   proof (rule ccontr)
  1126     assume unot: "u \<notin> closure(convex hull {x. x extreme_point_of S})"
  1093     assume unot: "u \<notin> closure(convex hull {x. x extreme_point_of S})"
  1127     then obtain a b where "a \<bullet> u < b"
  1094     then obtain a b where "a \<bullet> u < b"
  1165     shows "0 \<in> convex hull {x. x extreme_point_of S}"
  1132     shows "0 \<in> convex hull {x. x extreme_point_of S}"
  1166 using n S
  1133 using n S
  1167 proof (induction n arbitrary: S rule: less_induct)
  1134 proof (induction n arbitrary: S rule: less_induct)
  1168   case (less n S) show ?case
  1135   case (less n S) show ?case
  1169   proof (cases "0 \<in> rel_interior S")
  1136   proof (cases "0 \<in> rel_interior S")
  1170     case True with Krein_Milman show ?thesis
  1137     case True with Krein_Milman less.prems 
  1171       by (metis subsetD convex_convex_hull convex_rel_interior_closure less.prems(2) less.prems(3) rel_interior_subset)
  1138     show ?thesis
       
  1139       by (metis subsetD convex_convex_hull convex_rel_interior_closure rel_interior_subset)
  1172   next
  1140   next
  1173     case False
  1141     case False
  1174     have "rel_interior S \<noteq> {}"
  1142     have "rel_interior S \<noteq> {}"
  1175       by (simp add: rel_interior_convex_nonempty_aux less)
  1143       by (simp add: rel_interior_convex_nonempty_aux less)
  1176     then obtain c where c: "c \<in> rel_interior S" by blast
  1144     then obtain c where c: "c \<in> rel_interior S" by blast
  1177     obtain a where "a \<noteq> 0"
  1145     obtain a where "a \<noteq> 0"
  1178               and le_ay: "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> 0 \<le> a \<bullet> y"
  1146               and le_ay: "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> 0 \<le> a \<bullet> y"
  1179               and less_ay: "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> 0 < a \<bullet> y"
  1147               and less_ay: "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> 0 < a \<bullet> y"
  1180       by (blast intro: supporting_hyperplane_rel_boundary intro!: less False)
  1148       by (blast intro: supporting_hyperplane_rel_boundary intro!: less False)
  1181     have face: "S \<inter> {x. a \<bullet> x = 0} face_of S"
  1149     have face: "S \<inter> {x. a \<bullet> x = 0} face_of S"
  1182       apply (rule face_of_Int_supporting_hyperplane_ge [OF \<open>convex S\<close>])
  1150       using face_of_Int_supporting_hyperplane_ge le_ay \<open>convex S\<close> by auto
  1183       using le_ay by auto
       
  1184     then have co: "compact (S \<inter> {x. a \<bullet> x = 0})" "convex (S \<inter> {x. a \<bullet> x = 0})"
  1151     then have co: "compact (S \<inter> {x. a \<bullet> x = 0})" "convex (S \<inter> {x. a \<bullet> x = 0})"
  1185       using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+
  1152       using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+
  1186     have "a \<bullet> y = 0" if "y \<in> span (S \<inter> {x. a \<bullet> x = 0})" for y
  1153     have "a \<bullet> y = 0" if "y \<in> span (S \<inter> {x. a \<bullet> x = 0})" for y
  1187     proof -
  1154     proof -
  1188       have "y \<in> span {x. a \<bullet> x = 0}"
  1155       have "y \<in> span {x. a \<bullet> x = 0}"
  1194       by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff
  1161       by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff
  1195            inf_le1 \<open>dim S = n\<close> not_le rel_interior_subset span_0 span_base)
  1162            inf_le1 \<open>dim S = n\<close> not_le rel_interior_subset span_0 span_base)
  1196     then have "0 \<in> convex hull {x. x extreme_point_of (S \<inter> {x. a \<bullet> x = 0})}"
  1163     then have "0 \<in> convex hull {x. x extreme_point_of (S \<inter> {x. a \<bullet> x = 0})}"
  1197       by (rule less.IH) (auto simp: co less.prems)
  1164       by (rule less.IH) (auto simp: co less.prems)
  1198     then show ?thesis
  1165     then show ?thesis
  1199       by (metis (mono_tags, lifting) Collect_mono_iff \<open>S \<inter> {x. a \<bullet> x = 0} face_of S\<close> extreme_point_of_face hull_mono subset_iff)
  1166       by (metis (mono_tags, lifting) Collect_mono_iff face extreme_point_of_face hull_mono subset_iff)
  1200   qed
  1167   qed
  1201 qed
  1168 qed
  1202 
  1169 
  1203 
  1170 
  1204 theorem Krein_Milman_Minkowski:
  1171 theorem Krein_Milman_Minkowski:
  1295 
  1262 
  1296 lemma face_of_convex_hull_subset:
  1263 lemma face_of_convex_hull_subset:
  1297   fixes S :: "'a::euclidean_space set"
  1264   fixes S :: "'a::euclidean_space set"
  1298   assumes "compact S" and T: "T face_of (convex hull S)"
  1265   assumes "compact S" and T: "T face_of (convex hull S)"
  1299   obtains s' where "s' \<subseteq> S" "T = convex hull s'"
  1266   obtains s' where "s' \<subseteq> S" "T = convex hull s'"
  1300 apply (rule_tac s' = "{x. x extreme_point_of T}" in that)
  1267 proof
  1301 using T extreme_point_of_convex_hull extreme_point_of_face apply blast
  1268   show "{x. x extreme_point_of T} \<subseteq> S"
  1302 by (metis (no_types) Krein_Milman_Minkowski assms compact_convex_hull convex_convex_hull face_of_imp_compact face_of_imp_convex)
  1269     using T extreme_point_of_convex_hull extreme_point_of_face by blast
       
  1270   show "T = convex hull {x. x extreme_point_of T}"
       
  1271   proof (rule Krein_Milman_Minkowski)
       
  1272     show "compact T"
       
  1273       using T assms compact_convex_hull face_of_imp_compact by auto
       
  1274     show "convex T"
       
  1275       using T face_of_imp_convex by blast
       
  1276   qed
       
  1277 qed
  1303 
  1278 
  1304 
  1279 
  1305 lemma face_of_convex_hull_aux:
  1280 lemma face_of_convex_hull_aux:
  1306   assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c"
  1281   assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c"
  1307     and x: "u + v + w = x" "x \<noteq> 0" and S: "affine S" "a \<in> S" "b \<in> S" "c \<in> S"
  1282     and x: "u + v + w = x" "x \<noteq> 0" and S: "affine S" "a \<in> S" "b \<in> S" "c \<in> S"
  1341       have "T \<subseteq> insert a (convex hull T \<inter> convex hull S)"
  1316       have "T \<subseteq> insert a (convex hull T \<inter> convex hull S)"
  1342         using T hull_subset by fastforce
  1317         using T hull_subset by fastforce
  1343       then show "F \<subseteq> convex hull insert a (convex hull T \<inter> convex hull S)"
  1318       then show "F \<subseteq> convex hull insert a (convex hull T \<inter> convex hull S)"
  1344         by (simp add: FeqT hull_mono)
  1319         by (simp add: FeqT hull_mono)
  1345       show "convex hull insert a (convex hull T \<inter> convex hull S) \<subseteq> F"
  1320       show "convex hull insert a (convex hull T \<inter> convex hull S) \<subseteq> F"
  1346         apply (rule hull_minimal)
  1321         by (simp add: FeqT True hull_inc hull_minimal)
  1347         using True by (auto simp: \<open>F = convex hull T\<close> hull_inc)
       
  1348     qed
  1322     qed
  1349     moreover have "convex hull T \<inter> convex hull S face_of convex hull S"
  1323     moreover have "convex hull T \<inter> convex hull S face_of convex hull S"
  1350       by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI)
  1324       by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI)
  1351     ultimately show ?thesis
  1325     ultimately show ?thesis
  1352       using * by force
  1326       using * by force
  1388                \<in> open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
  1362                \<in> open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
  1389           and "0 \<le> ub" "ub \<le> 1" "0 \<le> uc" "uc \<le> 1" "0 \<le> ux" "ux \<le> 1"
  1363           and "0 \<le> ub" "ub \<le> 1" "0 \<le> uc" "uc \<le> 1" "0 \<le> ux" "ux \<le> 1"
  1390           and b: "b \<in> convex hull S" and c: "c \<in> convex hull S" and "x \<in> F"
  1364           and b: "b \<in> convex hull S" and c: "c \<in> convex hull S" and "x \<in> F"
  1391         for b c ub uc ux x
  1365         for b c ub uc ux x
  1392       proof -
  1366       proof -
       
  1367         have xah: "x \<in> affine hull S"
       
  1368           using F convex_hull_subset_affine_hull face_of_imp_subset \<open>x \<in> F\<close> by blast
       
  1369         have ah: "b \<in> affine hull S" "c \<in> affine hull S" 
       
  1370           using b c convex_hull_subset_affine_hull by blast+
  1393         obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \<noteq> (1 - uc) *\<^sub>R a + uc *\<^sub>R c"
  1371         obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \<noteq> (1 - uc) *\<^sub>R a + uc *\<^sub>R c"
  1394           and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x =
  1372           and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x =
  1395                     (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
  1373                     (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
  1396           and "0 < v" "v < 1"
  1374           and "0 < v" "v < 1"
  1397           using * by (auto simp: in_segment)
  1375           using * by (auto simp: in_segment)
  1400           by (auto simp: algebra_simps)
  1378           by (auto simp: algebra_simps)
  1401         then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a =
  1379         then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a =
  1402                    ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x"
  1380                    ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x"
  1403           by (auto simp: algebra_simps)
  1381           by (auto simp: algebra_simps)
  1404         then have "a \<in> affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \<noteq> 0"
  1382         then have "a \<in> affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \<noteq> 0"
  1405           apply (rule face_of_convex_hull_aux)
  1383           by (rule face_of_convex_hull_aux) (use b c xah ah that in \<open>auto simp: algebra_simps\<close>)
  1406           using b c that apply (auto simp: algebra_simps)
       
  1407           using F convex_hull_subset_affine_hull face_of_imp_subset \<open>x \<in> F\<close> apply blast+
       
  1408           done
       
  1409         then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0"
  1384         then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0"
  1410           using a by blast
  1385           using a by blast
  1411         with 0 have equx: "(1 - v) * ub + v * uc = ux"
  1386         with 0 have equx: "(1 - v) * ub + v * uc = ux"
  1412           and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)"
  1387           and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)"
  1413           by auto (auto simp: algebra_simps)
  1388           by auto (auto simp: algebra_simps)
  1414         show ?thesis
  1389         show ?thesis
  1415         proof (cases "uc = 0")
  1390         proof (cases "uc = 0")
  1416           case True
  1391           case True
  1417           then show ?thesis
  1392           then show ?thesis
  1418             using equx 0 \<open>0 \<le> ub\<close> \<open>ub \<le> 1\<close> \<open>v < 1\<close> \<open>x \<in> F\<close>
  1393             using equx \<open>0 \<le> ub\<close> \<open>ub \<le> 1\<close> \<open>v < 1\<close> uxx \<open>x \<in> F\<close> by force
  1419             apply (auto simp: algebra_simps)
       
  1420              apply (rule_tac x=x in exI, simp)
       
  1421              apply (rule_tac x=ub in exI, auto)
       
  1422              apply (metis add.left_neutral diff_eq_eq less_irrefl mult.commute mult_cancel_right1 real_vector.scale_cancel_left real_vector.scale_left_diff_distrib)
       
  1423             using \<open>x \<in> F\<close> \<open>uc \<le> 1\<close> apply blast
       
  1424             done
       
  1425         next
  1394         next
  1426           case False
  1395           case False
  1427           show ?thesis
  1396           show ?thesis
  1428           proof (cases "ub = 0")
  1397           proof (cases "ub = 0")
  1429             case True
  1398             case True
  1430             then show ?thesis
  1399             then show ?thesis
  1431               using equx 0 \<open>0 \<le> uc\<close> \<open>uc \<le> 1\<close> \<open>0 < v\<close> \<open>x \<in> F\<close> \<open>uc \<noteq> 0\<close> by (force simp: algebra_simps)
  1400               using equx \<open>0 \<le> uc\<close> \<open>uc \<le> 1\<close> \<open>0 < v\<close> uxx \<open>x \<in> F\<close> by force
  1432           next
  1401           next
  1433             case False
  1402             case False
  1434             then have "0 < ub" "0 < uc"
  1403             then have "0 < ub" "0 < uc"
  1435               using \<open>uc \<noteq> 0\<close> \<open>0 \<le> ub\<close> \<open>0 \<le> uc\<close> by auto
  1404               using \<open>uc \<noteq> 0\<close> \<open>0 \<le> ub\<close> \<open>0 \<le> uc\<close> by auto
       
  1405             then have "(1 - v) * ub > 0" "v * uc > 0"
       
  1406               by (simp_all add: \<open>0 < uc\<close> \<open>0 < v\<close> \<open>v < 1\<close>)
  1436             then have "ux \<noteq> 0"
  1407             then have "ux \<noteq> 0"
  1437               by (metis \<open>0 < v\<close> \<open>v < 1\<close> diff_ge_0_iff_ge dual_order.strict_implies_order equx leD le_add_same_cancel2 zero_le_mult_iff zero_less_mult_iff)
  1408               using equx \<open>0 < v\<close> by auto
  1438             have "b \<in> F \<and> c \<in> F"
  1409             have "b \<in> F \<and> c \<in> F"
  1439             proof (cases "b = c")
  1410             proof (cases "b = c")
  1440               case True
  1411               case True
  1441               then show ?thesis
  1412               then show ?thesis
  1442                 by (metis \<open>ux \<noteq> 0\<close> equx real_vector.scale_cancel_left scaleR_add_left uxx \<open>x \<in> F\<close>)
  1413                 by (metis \<open>ux \<noteq> 0\<close> equx real_vector.scale_cancel_left scaleR_add_left uxx \<open>x \<in> F\<close>)
  1487 next
  1458 next
  1488   assume ?rhs
  1459   assume ?rhs
  1489   then obtain c where "c \<subseteq> S" and T: "T = convex hull c"
  1460   then obtain c where "c \<subseteq> S" and T: "T = convex hull c"
  1490     by blast
  1461     by blast
  1491   have "affine hull c \<inter> affine hull (S - c) = {}"
  1462   have "affine hull c \<inter> affine hull (S - c) = {}"
  1492     apply (rule disjoint_affine_hull [OF assms \<open>c \<subseteq> S\<close>], auto)
  1463     by (intro disjoint_affine_hull [OF assms \<open>c \<subseteq> S\<close>], auto)
  1493     done
       
  1494   then have "affine hull c \<inter> convex hull (S - c) = {}"
  1464   then have "affine hull c \<inter> convex hull (S - c) = {}"
  1495     using convex_hull_subset_affine_hull by fastforce
  1465     using convex_hull_subset_affine_hull by fastforce
  1496   then show ?lhs
  1466   then show ?lhs
  1497     by (metis face_of_convex_hulls \<open>c \<subseteq> S\<close> aff_independent_finite assms T)
  1467     by (metis face_of_convex_hulls \<open>c \<subseteq> S\<close> aff_independent_finite assms T)
  1498 qed
  1468 qed
  1531   assume ?rhs
  1501   assume ?rhs
  1532   then obtain u where "T \<noteq> {}" "u \<in> S" and u: "T = convex hull (S - {u})"
  1502   then obtain u where "T \<noteq> {}" "u \<in> S" and u: "T = convex hull (S - {u})"
  1533     by (force simp: facet_of_def)
  1503     by (force simp: facet_of_def)
  1534   then have "\<not> S \<subseteq> {u}"
  1504   then have "\<not> S \<subseteq> {u}"
  1535     using \<open>T \<noteq> {}\<close> u by auto
  1505     using \<open>T \<noteq> {}\<close> u by auto
  1536   have [simp]: "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1"
  1506   have "aff_dim (S - {u}) = aff_dim S - 1"
  1537     using assms \<open>u \<in> S\<close>
  1507     using assms \<open>u \<in> S\<close>
  1538     apply (simp add: aff_dim_convex_hull affine_dependent_def)
  1508     unfolding affine_dependent_def
  1539     apply (drule bspec, assumption)
       
  1540     by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S])
  1509     by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S])
  1541   show ?lhs
  1510   then have "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1"
  1542     apply (subst u)
  1511     by (simp add: aff_dim_convex_hull)
  1543     apply (simp add: \<open>\<not> S \<subseteq> {u}\<close> facet_of_def face_of_convex_hull_affine_independent [OF assms], blast)
  1512   then show ?lhs
  1544     done
  1513     by (metis Diff_subset \<open>T \<noteq> {}\<close> assms face_of_convex_hull_affine_independent facet_of_def u)
  1545 qed
  1514 qed
  1546 
  1515 
  1547 lemma facet_of_convex_hull_affine_independent_alt:
  1516 lemma facet_of_convex_hull_affine_independent_alt:
  1548   fixes S :: "'a::euclidean_space set"
  1517   fixes S :: "'a::euclidean_space set"
  1549   shows
  1518   assumes "\<not> affine_dependent S"
  1550    "\<not>affine_dependent S
  1519   shows "(T facet_of (convex hull S) \<longleftrightarrow> 2 \<le> card S \<and> (\<exists>u. u \<in> S \<and> T = convex hull (S - {u})))"
  1551         \<Longrightarrow> (T facet_of (convex hull S) \<longleftrightarrow>
  1520         (is "?lhs = ?rhs")
  1552              2 \<le> card S \<and> (\<exists>u. u \<in> S \<and> T = convex hull (S - {u})))"
  1521 proof
  1553 apply (simp add: facet_of_convex_hull_affine_independent)
  1522   assume L: ?lhs 
  1554 apply (auto simp: Set.subset_singleton_iff)
  1523   then obtain x where
  1555 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)
  1524     "x \<in> S" and x: "T = convex hull (S - {x})" and "finite S"
  1556 done
  1525     using assms facet_of_convex_hull_affine_independent aff_independent_finite by blast
       
  1526   moreover have "Suc (Suc 0) \<le> card S"
       
  1527     using L  x \<open>x \<in> S\<close> \<open>finite S\<close>
       
  1528     by (metis Suc_leI assms card.remove convex_hull_eq_empty card_gt_0_iff facet_of_convex_hull_affine_independent finite_Diff not_less_eq_eq)
       
  1529   ultimately show ?rhs
       
  1530     by auto
       
  1531 next
       
  1532   assume ?rhs then show ?lhs
       
  1533     using assms
       
  1534     by (auto simp: facet_of_convex_hull_affine_independent Set.subset_singleton_iff)
       
  1535 qed
  1557 
  1536 
  1558 lemma segment_face_of:
  1537 lemma segment_face_of:
  1559   assumes "(closed_segment a b) face_of S"
  1538   assumes "(closed_segment a b) face_of S"
  1560   shows "a extreme_point_of S" "b extreme_point_of S"
  1539   shows "a extreme_point_of S" "b extreme_point_of S"
  1561 proof -
  1540 proof -
  1582           (is "?lhs = ?rhs")
  1561           (is "?lhs = ?rhs")
  1583 proof
  1562 proof
  1584   have "?lhs \<subseteq> convex hull {x. x extreme_point_of S}"
  1563   have "?lhs \<subseteq> convex hull {x. x extreme_point_of S}"
  1585     using Krein_Milman_Minkowski assms by blast
  1564     using Krein_Milman_Minkowski assms by blast
  1586   also have "... \<subseteq> ?rhs"
  1565   also have "... \<subseteq> ?rhs"
  1587     apply (rule hull_mono)
  1566   proof (rule hull_mono)
  1588     apply (auto simp: frontier_def extreme_point_not_in_interior)
  1567     show "{x. x extreme_point_of S} \<subseteq> frontier S"
  1589     using closure_subset apply (force simp: extreme_point_of_def)
  1568       using closure_subset
  1590     done
  1569       by (auto simp: frontier_def extreme_point_not_in_interior extreme_point_of_def)
       
  1570   qed
  1591   finally show "?lhs \<subseteq> ?rhs" .
  1571   finally show "?lhs \<subseteq> ?rhs" .
  1592 next
  1572 next
  1593   have "?rhs \<subseteq> convex hull S"
  1573   have "?rhs \<subseteq> convex hull S"
  1594     by (metis Diff_subset \<open>compact S\<close> closure_closed compact_eq_bounded_closed frontier_def hull_mono)
  1574     by (metis Diff_subset \<open>compact S\<close> closure_closed compact_eq_bounded_closed frontier_def hull_mono)
  1595   also have "... \<subseteq> ?lhs"
  1575   also have "... \<subseteq> ?lhs"
  1601 
  1581 
  1602 definition\<^marker>\<open>tag important\<close> polytope where
  1582 definition\<^marker>\<open>tag important\<close> polytope where
  1603  "polytope S \<equiv> \<exists>v. finite v \<and> S = convex hull v"
  1583  "polytope S \<equiv> \<exists>v. finite v \<and> S = convex hull v"
  1604 
  1584 
  1605 lemma polytope_translation_eq: "polytope (image (\<lambda>x. a + x) S) \<longleftrightarrow> polytope S"
  1585 lemma polytope_translation_eq: "polytope (image (\<lambda>x. a + x) S) \<longleftrightarrow> polytope S"
  1606 apply (simp add: polytope_def, safe)
  1586 proof -
  1607 apply (metis convex_hull_translation finite_imageI translation_galois)
  1587   have "\<And>a A. polytope A \<Longrightarrow> polytope ((+) a ` A)"
  1608 by (metis convex_hull_translation finite_imageI)
  1588     by (metis (no_types) convex_hull_translation finite_imageI polytope_def)
       
  1589   then show ?thesis
       
  1590     by (metis (no_types) add.left_inverse image_add_0 translation_assoc)
       
  1591 qed
  1609 
  1592 
  1610 lemma polytope_linear_image: "\<lbrakk>linear f; polytope p\<rbrakk> \<Longrightarrow> polytope(image f p)"
  1593 lemma polytope_linear_image: "\<lbrakk>linear f; polytope p\<rbrakk> \<Longrightarrow> polytope(image f p)"
  1611   unfolding polytope_def using convex_hull_linear_image by blast
  1594   unfolding polytope_def using convex_hull_linear_image by blast
  1612 
  1595 
  1613 lemma polytope_empty: "polytope {}"
  1596 lemma polytope_empty: "polytope {}"
  1700             S = \<Inter> F \<and>
  1683             S = \<Inter> F \<and>
  1701             (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b})"
  1684             (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b})"
  1702 
  1685 
  1703 lemma polyhedron_Int [intro,simp]:
  1686 lemma polyhedron_Int [intro,simp]:
  1704    "\<lbrakk>polyhedron S; polyhedron T\<rbrakk> \<Longrightarrow> polyhedron (S \<inter> T)"
  1687    "\<lbrakk>polyhedron S; polyhedron T\<rbrakk> \<Longrightarrow> polyhedron (S \<inter> T)"
  1705   apply (simp add: polyhedron_def, clarify)
  1688   apply (clarsimp simp add: polyhedron_def)
  1706   apply (rename_tac F G)
  1689   subgoal for F G
  1707   apply (rule_tac x="F \<union> G" in exI, auto)
  1690     by (rule_tac x="F \<union> G" in exI, auto)
  1708   done
  1691   done
  1709 
  1692 
  1710 lemma polyhedron_UNIV [iff]: "polyhedron UNIV"
  1693 lemma polyhedron_UNIV [iff]: "polyhedron UNIV"
  1711   unfolding polyhedron_def
  1694   unfolding polyhedron_def
  1712   by (rule_tac x="{}" in exI) auto
  1695   by (rule_tac x="{}" in exI) auto
  1716 by (induction F rule: finite_induct) auto
  1699 by (induction F rule: finite_induct) auto
  1717 
  1700 
  1718 
  1701 
  1719 lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)"
  1702 lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)"
  1720 proof -
  1703 proof -
  1721   have "\<exists>a. a \<noteq> 0 \<and>
  1704   define i::'a where "(i \<equiv> SOME i. i \<in> Basis)"
  1722              (\<exists>b. {x. (SOME i. i \<in> Basis) \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b})"
  1705   have "\<exists>a. a \<noteq> 0 \<and> (\<exists>b. {x. i \<bullet> x \<le> -1} = {x. a \<bullet> x \<le> b})"
  1723     by (rule_tac x="(SOME i. i \<in> Basis)" in exI) (force simp: SOME_Basis nonzero_Basis)
  1706     by (rule_tac x="i" in exI) (force simp: i_def SOME_Basis nonzero_Basis)
  1724   moreover have "\<exists>a b. a \<noteq> 0 \<and>
  1707   moreover have "\<exists>a b. a \<noteq> 0 \<and> {x. -i \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b}"
  1725                        {x. - (SOME i. i \<in> Basis) \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b}"
  1708       apply (rule_tac x="-i" in exI)
  1726       apply (rule_tac x="-(SOME i. i \<in> Basis)" in exI)
       
  1727       apply (rule_tac x="-1" in exI)
  1709       apply (rule_tac x="-1" in exI)
  1728       apply (simp add: SOME_Basis nonzero_Basis)
  1710       apply (simp add: i_def SOME_Basis nonzero_Basis)
  1729       done
  1711       done
  1730   ultimately show ?thesis
  1712   ultimately show ?thesis
  1731     unfolding polyhedron_def
  1713     unfolding polyhedron_def
  1732     apply (rule_tac x="{{x. (SOME i. i \<in> Basis) \<bullet> x \<le> -1},
  1714     by (rule_tac x="{{x. i \<bullet> x \<le> -1}, {x. -i \<bullet> x \<le> -1}}" in exI) force
  1733                         {x. -(SOME i. i \<in> Basis) \<bullet> x \<le> -1}}" in exI)
       
  1734     apply force
       
  1735     done
       
  1736 qed
  1715 qed
  1737 
  1716 
  1738 lemma polyhedron_halfspace_le:
  1717 lemma polyhedron_halfspace_le:
  1739   fixes a :: "'a :: euclidean_space"
  1718   fixes a :: "'a :: euclidean_space"
  1740   shows "polyhedron {x. a \<bullet> x \<le> b}"
  1719   shows "polyhedron {x. a \<bullet> x \<le> b}"
  1768 by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S])
  1747 by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S])
  1769 
  1748 
  1770 lemma polyhedron_imp_closed:
  1749 lemma polyhedron_imp_closed:
  1771   fixes S :: "'a :: euclidean_space set"
  1750   fixes S :: "'a :: euclidean_space set"
  1772   shows "polyhedron S \<Longrightarrow> closed S"
  1751   shows "polyhedron S \<Longrightarrow> closed S"
  1773 apply (simp add: polyhedron_def)
  1752   by (metis closed_Inter closed_halfspace_le polyhedron_def)
  1774 using closed_halfspace_le by fastforce
       
  1775 
  1753 
  1776 lemma polyhedron_imp_convex:
  1754 lemma polyhedron_imp_convex:
  1777   fixes S :: "'a :: euclidean_space set"
  1755   fixes S :: "'a :: euclidean_space set"
  1778   shows "polyhedron S \<Longrightarrow> convex S"
  1756   shows "polyhedron S \<Longrightarrow> convex S"
  1779 apply (simp add: polyhedron_def)
  1757   by (metis convex_Inter convex_halfspace_le polyhedron_def)
  1780 using convex_Inter convex_halfspace_le by fastforce
       
  1781 
  1758 
  1782 lemma polyhedron_affine_hull:
  1759 lemma polyhedron_affine_hull:
  1783   fixes S :: "'a :: euclidean_space set"
  1760   fixes S :: "'a :: euclidean_space set"
  1784   shows "polyhedron(affine hull S)"
  1761   shows "polyhedron(affine hull S)"
  1785 by (simp add: affine_imp_polyhedron)
  1762 by (simp add: affine_imp_polyhedron)
  1793            (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
  1770            (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
  1794                 (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}))"
  1771                 (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}))"
  1795         (is "?lhs = ?rhs")
  1772         (is "?lhs = ?rhs")
  1796 proof
  1773 proof
  1797   assume ?lhs then show ?rhs
  1774   assume ?lhs then show ?rhs
  1798     apply (simp add: polyhedron_def)
  1775     using hull_subset polyhedron_def by fastforce
  1799     apply (erule ex_forward)
       
  1800     using hull_subset apply force
       
  1801     done
       
  1802 next
  1776 next
  1803   assume ?rhs then show ?lhs
  1777   assume ?rhs then show ?lhs
  1804     apply clarify
  1778     by (metis polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le)
  1805     apply (erule ssubst)
       
  1806     apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le)
       
  1807     done
       
  1808 qed
  1779 qed
  1809 
  1780 
  1810 proposition rel_interior_polyhedron_explicit:
  1781 proposition rel_interior_polyhedron_explicit:
  1811   assumes "finite F"
  1782   assumes "finite F"
  1812       and seq: "S = affine hull S \<inter> \<Inter>F"
  1783       and seq: "S = affine hull S \<inter> \<Inter>F"
  1844       using \<open>z \<noteq> x\<close> \<open>e > 0\<close> by (simp add: \<xi>_def min_def field_split_simps norm_minus_commute)
  1815       using \<open>z \<noteq> x\<close> \<open>e > 0\<close> by (simp add: \<xi>_def min_def field_split_simps norm_minus_commute)
  1845     finally have lte: "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) < e" .
  1816     finally have lte: "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) < e" .
  1846     have \<xi>_aff: "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> affine hull S"
  1817     have \<xi>_aff: "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> affine hull S"
  1847       by (metis \<open>x \<in> S\<close> add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff)
  1818       by (metis \<open>x \<in> S\<close> add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff)
  1848     have "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> S"
  1819     have "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> S"
  1849       apply (rule ins [OF _ \<xi>_aff])
  1820       using ins [OF _ \<xi>_aff] by (simp add: algebra_simps lte)
  1850       apply (simp add: algebra_simps lte)
       
  1851       done
       
  1852     then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \<in> S"
  1821     then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \<in> S"
  1853       apply (rule_tac l = \<xi> in that)
  1822       using \<open>e > 0\<close> \<open>z \<noteq> x\<close>  
  1854       using \<open>e > 0\<close> \<open>z \<noteq> x\<close>  apply (auto simp: \<xi>_def)
  1823       by (rule_tac l = \<xi> in that) (auto simp: \<xi>_def)
  1855       done
       
  1856     then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \<in> i"
  1824     then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \<in> i"
  1857       using seq \<open>i \<in> F\<close> by auto
  1825       using seq \<open>i \<in> F\<close> by auto
  1858     have "b i * l + (a i \<bullet> x) * (1 - l) < a i \<bullet> (l *\<^sub>R z + (1 - l) *\<^sub>R x)"
  1826     have "b i * l + (a i \<bullet> x) * (1 - l) < a i \<bullet> (l *\<^sub>R z + (1 - l) *\<^sub>R x)"
  1859       using l by (simp add: algebra_simps aiz)
  1827       using l by (simp add: algebra_simps aiz)
  1860     also have "\<dots> \<le> b i" using i l
  1828     also have "\<dots> \<le> b i" using i l
  1868            if "x \<in> S" and less: "\<And>h. h \<in> F \<Longrightarrow> a h \<bullet> x < b h" for x
  1836            if "x \<in> S" and less: "\<And>h. h \<in> F \<Longrightarrow> a h \<bullet> x < b h" for x
  1869   proof -
  1837   proof -
  1870     have 1: "\<And>h. h \<in> F \<Longrightarrow> x \<in> interior h"
  1838     have 1: "\<And>h. h \<in> F \<Longrightarrow> x \<in> interior h"
  1871       by (metis interior_halfspace_le mem_Collect_eq less faceq)
  1839       by (metis interior_halfspace_le mem_Collect_eq less faceq)
  1872     have 2: "\<And>y. \<lbrakk>\<forall>h\<in>F. y \<in> interior h; y \<in> affine hull S\<rbrakk> \<Longrightarrow> y \<in> S"
  1840     have 2: "\<And>y. \<lbrakk>\<forall>h\<in>F. y \<in> interior h; y \<in> affine hull S\<rbrakk> \<Longrightarrow> y \<in> S"
  1873       by (metis IntI Inter_iff contra_subsetD interior_subset seq)
  1841       by (metis IntI Inter_iff subsetD interior_subset seq)
  1874     show ?thesis
  1842     show ?thesis
  1875       apply (simp add: rel_interior \<open>x \<in> S\<close>)
  1843       apply (simp add: rel_interior \<open>x \<in> S\<close>)
  1876       apply (rule_tac x="\<Inter>h\<in>F. interior h" in exI)
  1844       apply (rule_tac x="\<Inter>h\<in>F. interior h" in exI)
  1877       apply (auto simp: \<open>finite F\<close> open_INT 1 2)
  1845       apply (auto simp: \<open>finite F\<close> open_INT 1 2)
  1878       done
  1846       done
  1928        using ab apply blast
  1896        using ab apply blast
  1929        done
  1897        done
  1930   qed
  1898   qed
  1931 next
  1899 next
  1932   assume ?rhs then show ?lhs
  1900   assume ?rhs then show ?lhs
  1933     apply (simp add: polyhedron_Int_affine)
  1901     by (metis polyhedron_Int_affine)
  1934     by metis
       
  1935 qed
  1902 qed
  1936 
  1903 
  1937 
  1904 
  1938 proposition polyhedron_Int_affine_parallel_minimal:
  1905 proposition polyhedron_Int_affine_parallel_minimal:
  1939   fixes S :: "'a :: euclidean_space set"
  1906   fixes S :: "'a :: euclidean_space set"
  1967   then have *: "\<not> (?P g \<and> ?Q g)" if "g \<subset> F" for g
  1934   then have *: "\<not> (?P g \<and> ?Q g)" if "g \<subset> F" for g
  1968     using that \<open>finite F\<close> psubset_card_mono \<open>card F = n\<close>
  1935     using that \<open>finite F\<close> psubset_card_mono \<open>card F = n\<close>
  1969     by (metis finite_Int inf.strict_order_iff)
  1936     by (metis finite_Int inf.strict_order_iff)
  1970   have 1: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subseteq> affine hull S \<inter> \<Inter>F'"
  1937   have 1: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subseteq> affine hull S \<inter> \<Inter>F'"
  1971     by (subst seq) blast
  1938     by (subst seq) blast
  1972   have 2: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<noteq> affine hull S \<inter> \<Inter>F'"
  1939   have 2: "S \<noteq> affine hull S \<inter> \<Inter>F'" if "F' \<subset> F" for F'
  1973     apply (frule *)
  1940     using * [OF that] by (metis IntE aff inf.strict_order_iff that)
  1974     by (metis aff subsetCE subset_iff_psubset_eq)
       
  1975   show ?rhs
  1941   show ?rhs
  1976     by (metis \<open>finite F\<close> seq aff psubsetI 1 2)
  1942     by (metis \<open>finite F\<close> seq aff psubsetI 1 2)
  1977 next
  1943 next
  1978   assume ?rhs then show ?lhs
  1944   assume ?rhs then show ?lhs
  1979     by (auto simp: polyhedron_Int_affine_parallel)
  1945     by (auto simp: polyhedron_Int_affine_parallel)
  1984   fixes S :: "'a :: euclidean_space set"
  1950   fixes S :: "'a :: euclidean_space set"
  1985   shows "polyhedron S \<longleftrightarrow>
  1951   shows "polyhedron S \<longleftrightarrow>
  1986          (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
  1952          (\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
  1987               (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}) \<and>
  1953               (\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}) \<and>
  1988               (\<forall>F'. F' \<subset> F \<longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'))"
  1954               (\<forall>F'. F' \<subset> F \<longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'))"
  1989 apply (rule iffI)
  1955      (is "?lhs = ?rhs")
  1990  apply (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward)
  1956 proof
  1991 apply (auto simp: polyhedron_Int_affine elim!: ex_forward)
  1957   assume ?lhs
  1992 done
  1958   then show ?rhs
       
  1959     by (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward)
       
  1960 qed (auto simp: polyhedron_Int_affine elim!: ex_forward)
  1993 
  1961 
  1994 proposition facet_of_polyhedron_explicit:
  1962 proposition facet_of_polyhedron_explicit:
  1995   assumes "finite F"
  1963   assumes "finite F"
  1996       and seq: "S = affine hull S \<inter> \<Inter>F"
  1964       and seq: "S = affine hull S \<inter> \<Inter>F"
  1997       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  1965       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  1998       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  1966       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  1999     shows "c facet_of S \<longleftrightarrow> (\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h})"
  1967     shows "C facet_of S \<longleftrightarrow> (\<exists>h. h \<in> F \<and> C = S \<inter> {x. a h \<bullet> x = b h})"
  2000 proof (cases "S = {}")
  1968 proof (cases "S = {}")
  2001   case True with psub show ?thesis by force
  1969   case True with psub show ?thesis by force
  2002 next
  1970 next
  2003   case False
  1971   case False
  2004   have "polyhedron S"
  1972   have "polyhedron S"
  2005     apply (simp add: polyhedron_Int_affine)
  1973     unfolding polyhedron_Int_affine by (metis \<open>finite F\<close> faceq seq)
  2006     apply (rule_tac x=F in exI)
       
  2007     using assms  apply force
       
  2008     done
       
  2009   then have "convex S"
  1974   then have "convex S"
  2010     by (rule polyhedron_imp_convex)
  1975     by (rule polyhedron_imp_convex)
  2011   with False rel_interior_eq_empty have "rel_interior S \<noteq> {}" by blast
  1976   with False rel_interior_eq_empty have "rel_interior S \<noteq> {}" by blast
  2012   then obtain x where "x \<in> rel_interior S" by auto
  1977   then obtain x where "x \<in> rel_interior S" by auto
  2013   then obtain T where "open T" "x \<in> T" "x \<in> S" "T \<inter> affine hull S \<subseteq> S"
  1978   then obtain T where "open T" "x \<in> T" "x \<in> S" "T \<inter> affine hull S \<subseteq> S"
  2038     have awlt: "a i \<bullet> w < b i" if "i \<in> F" "i \<noteq> h" for i
  2003     have awlt: "a i \<bullet> w < b i" if "i \<in> F" "i \<noteq> h" for i
  2039     proof -
  2004     proof -
  2040       have "(1 - l) * (a i \<bullet> x) < (1 - l) * b i"
  2005       have "(1 - l) * (a i \<bullet> x) < (1 - l) * b i"
  2041         by (simp add: \<open>l < 1\<close> \<open>i \<in> F\<close>)
  2006         by (simp add: \<open>l < 1\<close> \<open>i \<in> F\<close>)
  2042       moreover have "l * (a i \<bullet> z) \<le> l * b i"
  2007       moreover have "l * (a i \<bullet> z) \<le> l * b i"
  2043         apply (rule mult_left_mono)
  2008       proof (rule mult_left_mono)
  2044         apply (metis Diff_insert_absorb Inter_iff Set.set_insert \<open>h \<in> F\<close> faceq insertE mem_Collect_eq that zint)
  2009         show "a i \<bullet> z \<le> b i"
  2045         using \<open>0 < l\<close>
  2010           by (metis Diff_insert_absorb Inter_iff Set.set_insert \<open>h \<in> F\<close> faceq insertE mem_Collect_eq that zint)
  2046         apply simp
  2011       qed (use \<open>0 < l\<close> in auto)
  2047         done
       
  2048       ultimately show ?thesis by (simp add: w_def algebra_simps)
  2012       ultimately show ?thesis by (simp add: w_def algebra_simps)
  2049     qed
  2013     qed
  2050     have weq: "a h \<bullet> w = b h"
  2014     have weq: "a h \<bullet> w = b h"
  2051       using xltz unfolding w_def l_def
  2015       using xltz unfolding w_def l_def
  2052       by (simp add: algebra_simps) (simp add: field_simps)
  2016       by (simp add: algebra_simps) (simp add: field_simps)
       
  2017     have faceS: "S \<inter> {x. a h \<bullet> x = b h} face_of S"
       
  2018     proof (rule face_of_Int_supporting_hyperplane_le)
       
  2019       show "\<And>x. x \<in> S \<Longrightarrow> a h \<bullet> x \<le> b h"
       
  2020         using faceq seq that by fastforce
       
  2021     qed fact
  2053     have "w \<in> affine hull S"
  2022     have "w \<in> affine hull S"
  2054       by (simp add: w_def mem_affine xaff zaff)
  2023       by (simp add: w_def mem_affine xaff zaff)
  2055     moreover have "w \<in> \<Inter>F"
  2024     moreover have "w \<in> \<Inter>F"
  2056       using \<open>a h \<bullet> w = b h\<close> awlt faceq less_eq_real_def by blast
  2025       using \<open>a h \<bullet> w = b h\<close> awlt faceq less_eq_real_def by blast
  2057     ultimately have "w \<in> S"
  2026     ultimately have "w \<in> S"
  2058       using seq by blast
  2027       using seq by blast
  2059     with weq have "S \<inter> {x. a h \<bullet> x = b h} \<noteq> {}" by blast
  2028     with weq have ne: "S \<inter> {x. a h \<bullet> x = b h} \<noteq> {}" by blast
  2060     moreover have "S \<inter> {x. a h \<bullet> x = b h} face_of S"
  2029     moreover have "affine hull (S \<inter> {x. a h \<bullet> x = b h}) = (affine hull S) \<inter> {x. a h \<bullet> x = b h}"
  2061       apply (rule face_of_Int_supporting_hyperplane_le)
       
  2062       apply (rule \<open>convex S\<close>)
       
  2063       apply (subst (asm) seq)
       
  2064       using faceq that apply fastforce
       
  2065       done
       
  2066     moreover have "affine hull (S \<inter> {x. a h \<bullet> x = b h}) =
       
  2067                    (affine hull S) \<inter> {x. a h \<bullet> x = b h}"
       
  2068     proof
  2030     proof
  2069       show "affine hull (S \<inter> {x. a h \<bullet> x = b h}) \<subseteq> affine hull S \<inter> {x. a h \<bullet> x = b h}"
  2031       show "affine hull (S \<inter> {x. a h \<bullet> x = b h}) \<subseteq> affine hull S \<inter> {x. a h \<bullet> x = b h}"
  2070         apply (intro Int_greatest hull_mono Int_lower1)
  2032         apply (intro Int_greatest hull_mono Int_lower1)
  2071         apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2)
  2033         apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2)
  2072         done
  2034         done
  2082             by (rule_tac T=1 in that) auto
  2044             by (rule_tac T=1 in that) auto
  2083         next
  2045         next
  2084           case False
  2046           case False
  2085           then obtain h' where h': "h' \<in> F - {h}" by auto
  2047           then obtain h' where h': "h' \<in> F - {h}" by auto
  2086           let ?body = "(\<lambda>j. if 0 < a j \<bullet> y - a j \<bullet> w
  2048           let ?body = "(\<lambda>j. if 0 < a j \<bullet> y - a j \<bullet> w
  2087               then (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w)
  2049               then (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w) else 1) ` (F - {h})"
  2088               else 1) ` (F - {h})"
       
  2089           define inff where "inff = Inf ?body"
  2050           define inff where "inff = Inf ?body"
  2090           from \<open>finite F\<close> have "finite ?body"
  2051           from \<open>finite F\<close> have "finite ?body"
  2091             by blast
  2052             by blast
  2092           moreover from h' have "?body \<noteq> {}"
  2053           moreover from h' have "?body \<noteq> {}"
  2093             by blast
  2054             by blast
  2107           moreover have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> b j - a j \<bullet> w"
  2068           moreover have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> b j - a j \<bullet> w"
  2108                         if "j \<in> F" "j \<noteq> h" for j
  2069                         if "j \<in> F" "j \<noteq> h" for j
  2109           proof (cases "a j \<bullet> w < a j \<bullet> y")
  2070           proof (cases "a j \<bullet> w < a j \<bullet> y")
  2110             case True
  2071             case True
  2111             then have "inff \<le> (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w)"
  2072             then have "inff \<le> (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w)"
  2112               apply (simp add: inff_def)
  2073               unfolding inff_def
  2113               apply (rule cInf_le_finite)
  2074               using \<open>finite F\<close> by (auto intro: cInf_le_finite simp add: that split: if_split_asm)
  2114               using \<open>finite F\<close> apply blast
       
  2115               apply (simp add: that split: if_split_asm)
       
  2116               done
       
  2117             then show ?thesis
  2075             then show ?thesis
  2118               using \<open>0 < inff\<close> awlt [OF that] mult_strict_left_mono
  2076               using \<open>0 < inff\<close> awlt [OF that] mult_strict_left_mono
  2119               by (fastforce simp add: field_split_simps split: if_split_asm)
  2077               by (fastforce simp add: field_split_simps split: if_split_asm)
  2120           next
  2078           next
  2121             case False
  2079             case False
  2126             finally show ?thesis by simp
  2084             finally show ?thesis by simp
  2127           qed
  2085           qed
  2128           ultimately show ?thesis
  2086           ultimately show ?thesis
  2129             by (blast intro: that)
  2087             by (blast intro: that)
  2130         qed
  2088         qed
  2131         define c where "c = (1 - T) *\<^sub>R w + T *\<^sub>R y"
  2089         define C where "C = (1 - T) *\<^sub>R w + T *\<^sub>R y"
  2132         have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> j" if "j \<in> F" for j
  2090         have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> j" if "j \<in> F" for j
  2133         proof (cases "j = h")
  2091         proof (cases "j = h")
  2134           case True
  2092           case True
  2135           have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a h \<bullet> x \<le> b h}"
  2093           have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a h \<bullet> x \<le> b h}"
  2136             using weq yaff by (auto simp: algebra_simps)
  2094             using weq yaff by (auto simp: algebra_simps)
  2140           with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a j \<bullet> x \<le> b j}"
  2098           with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a j \<bullet> x \<le> b j}"
  2141             by (simp add: algebra_simps)
  2099             by (simp add: algebra_simps)
  2142           with faceq [OF that] show ?thesis by simp
  2100           with faceq [OF that] show ?thesis by simp
  2143         qed
  2101         qed
  2144         moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> affine hull S"
  2102         moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> affine hull S"
  2145           apply (rule affine_affine_hull [simplified affine_alt, rule_format])
  2103           using yaff \<open>w \<in> affine hull S\<close> affine_affine_hull affine_alt by blast
  2146           apply (simp add: \<open>w \<in> affine hull S\<close>)
  2104         ultimately have "C \<in> S"
  2147           using yaff apply blast
  2105           using seq by (force simp: C_def)
  2148           done
  2106         moreover have "a h \<bullet> C = b h"
  2149         ultimately have "c \<in> S"
  2107           using yaff by (force simp: C_def algebra_simps weq)
  2150           using seq by (force simp: c_def)
  2108         ultimately have caff: "C \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
  2151         moreover have "a h \<bullet> c = b h"
       
  2152           using yaff by (force simp: c_def algebra_simps weq)
       
  2153         ultimately have caff: "c \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
       
  2154           by (simp add: hull_inc)
  2109           by (simp add: hull_inc)
  2155         have waff: "w \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
  2110         have waff: "w \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
  2156           using \<open>w \<in> S\<close> weq by (blast intro: hull_inc)
  2111           using \<open>w \<in> S\<close> weq by (blast intro: hull_inc)
  2157         have yeq: "y = (1 - inverse T) *\<^sub>R w + c /\<^sub>R T"
  2112         have yeq: "y = (1 - inverse T) *\<^sub>R w + C /\<^sub>R T"
  2158           using \<open>0 < T\<close> by (simp add: c_def algebra_simps)
  2113           using \<open>0 < T\<close> by (simp add: C_def algebra_simps)
  2159         show "y \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
  2114         show "y \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
  2160           by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff])
  2115           by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff])
  2161       qed
  2116       qed
  2162     qed
  2117     qed
  2163     ultimately show ?thesis
  2118     ultimately have "aff_dim (affine hull (S \<inter> {x. a h \<bullet> x = b h})) = aff_dim S - 1"
  2164       apply (simp add: facet_of_def)
  2119       using \<open>b h < a h \<bullet> z\<close> zaff by (force simp: aff_dim_affine_Int_hyperplane)
  2165       apply (subst aff_dim_affine_hull [symmetric])
  2120     then show ?thesis
  2166       using  \<open>b h < a h \<bullet> z\<close> zaff
  2121       by (simp add: ne faceS facet_of_def)
  2167       apply (force simp: aff_dim_affine_Int_hyperplane)
       
  2168       done
       
  2169   qed
  2122   qed
  2170   show ?thesis
  2123   show ?thesis
  2171   proof
  2124   proof
  2172     show "\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h} \<Longrightarrow> c facet_of S"
  2125     show "\<exists>h. h \<in> F \<and> C = S \<inter> {x. a h \<bullet> x = b h} \<Longrightarrow> C facet_of S"
  2173       using * by blast
  2126       using * by blast
  2174   next
  2127   next
  2175     assume "c facet_of S"
  2128     assume "C facet_of S"
  2176     then have "c face_of S" "convex c" "c \<noteq> {}" and affc: "aff_dim c = aff_dim S - 1"
  2129     then have "C face_of S" "convex C" "C \<noteq> {}" and affc: "aff_dim C = aff_dim S - 1"
  2177       by (auto simp: facet_of_def face_of_imp_convex)
  2130       by (auto simp: facet_of_def face_of_imp_convex)
  2178     then obtain x where x: "x \<in> rel_interior c"
  2131     then obtain x where x: "x \<in> rel_interior C"
  2179       by (force simp: rel_interior_eq_empty)
  2132       by (force simp: rel_interior_eq_empty)
  2180     then have "x \<in> c"
  2133     then have "x \<in> C"
  2181       by (meson subsetD rel_interior_subset)
  2134       by (meson subsetD rel_interior_subset)
  2182     then have "x \<in> S"
  2135     then have "x \<in> S"
  2183       using \<open>c facet_of S\<close> facet_of_imp_subset by blast
  2136       using \<open>C facet_of S\<close> facet_of_imp_subset by blast
  2184     have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
  2137     have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
  2185       by (rule rel_interior_polyhedron_explicit [OF assms])
  2138       by (rule rel_interior_polyhedron_explicit [OF assms])
  2186     have "c \<noteq> S"
  2139     have "C \<noteq> S"
  2187       using \<open>c facet_of S\<close> facet_of_irrefl by blast
  2140       using \<open>C facet_of S\<close> facet_of_irrefl by blast
  2188     then have "x \<notin> rel_interior S"
  2141     then have "x \<notin> rel_interior S"
  2189       by (metis IntI empty_iff \<open>x \<in> c\<close> \<open>c \<noteq> S\<close> \<open>c face_of S\<close> face_of_disjoint_rel_interior)
  2142       by (metis IntI empty_iff \<open>x \<in> C\<close> \<open>C \<noteq> S\<close> \<open>C face_of S\<close> face_of_disjoint_rel_interior)
  2190     with rels \<open>x \<in> S\<close> obtain i where "i \<in> F" and i: "a i \<bullet> x \<ge> b i"
  2143     with rels \<open>x \<in> S\<close> obtain i where "i \<in> F" and i: "a i \<bullet> x \<ge> b i"
  2191       by force
  2144       by force
  2192     have "x \<in> {u. a i \<bullet> u \<le> b i}"
  2145     have "x \<in> {u. a i \<bullet> u \<le> b i}"
  2193       by (metis IntD2 InterE \<open>i \<in> F\<close> \<open>x \<in> S\<close> faceq seq)
  2146       by (metis IntD2 InterE \<open>i \<in> F\<close> \<open>x \<in> S\<close> faceq seq)
  2194     then have "a i \<bullet> x \<le> b i" by simp
  2147     then have "a i \<bullet> x \<le> b i" by simp
  2195     then have "a i \<bullet> x = b i" using i by auto
  2148     then have "a i \<bullet> x = b i" using i by auto
  2196     have "c \<subseteq> S \<inter> {x. a i \<bullet> x = b i}"
  2149     have "C \<subseteq> S \<inter> {x. a i \<bullet> x = b i}"
  2197       apply (rule subset_of_face_of [of _ S])
  2150     proof (rule subset_of_face_of [of _ S])
  2198         apply (simp add: "*" \<open>i \<in> F\<close> facet_of_imp_face_of)
  2151       show "S \<inter> {x. a i \<bullet> x = b i} face_of S"
  2199        apply (simp add: \<open>c face_of S\<close> face_of_imp_subset)
  2152         by (simp add: "*" \<open>i \<in> F\<close> facet_of_imp_face_of)
  2200       using \<open>a i \<bullet> x = b i\<close> \<open>x \<in> S\<close> x by blast
  2153       show "C \<subseteq> S"
  2201     then have cface: "c face_of (S \<inter> {x. a i \<bullet> x = b i})"
  2154         by (simp add: \<open>C face_of S\<close> face_of_imp_subset)
  2202       by (meson \<open>c face_of S\<close> face_of_subset inf_le1)
  2155       show "S \<inter> {x. a i \<bullet> x = b i} \<inter> rel_interior C \<noteq> {}"
       
  2156         using \<open>a i \<bullet> x = b i\<close> \<open>x \<in> S\<close> x by blast
       
  2157     qed
       
  2158     then have cface: "C face_of (S \<inter> {x. a i \<bullet> x = b i})"
       
  2159       by (meson \<open>C face_of S\<close> face_of_subset inf_le1)
  2203     have con: "convex (S \<inter> {x. a i \<bullet> x = b i})"
  2160     have con: "convex (S \<inter> {x. a i \<bullet> x = b i})"
  2204       by (simp add: \<open>convex S\<close> convex_Int convex_hyperplane)
  2161       by (simp add: \<open>convex S\<close> convex_Int convex_hyperplane)
  2205     show "\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h}"
  2162     show "\<exists>h. h \<in> F \<and> C = S \<inter> {x. a h \<bullet> x = b h}"
  2206       apply (rule_tac x=i in exI)
  2163       apply (rule_tac x=i in exI)
  2207       apply (simp add: \<open>i \<in> F\<close>)
       
  2208       by (metis (no_types) * \<open>i \<in> F\<close> affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface])
  2164       by (metis (no_types) * \<open>i \<in> F\<close> affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface])
  2209   qed
  2165   qed
  2210 qed
  2166 qed
  2211 
  2167 
  2212 
  2168 
  2214   fixes S :: "'a :: euclidean_space set"
  2170   fixes S :: "'a :: euclidean_space set"
  2215   assumes "finite F"
  2171   assumes "finite F"
  2216       and seq: "S = affine hull S \<inter> \<Inter>F"
  2172       and seq: "S = affine hull S \<inter> \<Inter>F"
  2217       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2173       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2218       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  2174       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  2219       and c: "c face_of S" and "c \<noteq> {}" "c \<noteq> S"
  2175       and C: "C face_of S" and "C \<noteq> {}" "C \<noteq> S"
  2220    obtains h where "h \<in> F" "c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}"
  2176    obtains h where "h \<in> F" "C \<subseteq> S \<inter> {x. a h \<bullet> x = b h}"
  2221 proof -
  2177 proof -
  2222   have "c \<subseteq> S" using \<open>c face_of S\<close>
  2178   have "C \<subseteq> S" using \<open>C face_of S\<close>
  2223     by (simp add: face_of_imp_subset)
  2179     by (simp add: face_of_imp_subset)
  2224   have "polyhedron S"
  2180   have "polyhedron S"
  2225     apply (simp add: polyhedron_Int_affine)
  2181     by (metis \<open>finite F\<close> faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq)
  2226     by (metis \<open>finite F\<close> faceq seq)
       
  2227   then have "convex S"
  2182   then have "convex S"
  2228     by (simp add: polyhedron_imp_convex)
  2183     by (simp add: polyhedron_imp_convex)
  2229   then have *: "(S \<inter> {x. a h \<bullet> x = b h}) face_of S" if "h \<in> F" for h
  2184   then have *: "(S \<inter> {x. a h \<bullet> x = b h}) face_of S" if "h \<in> F" for h
  2230     apply (rule face_of_Int_supporting_hyperplane_le)
  2185     using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce
  2231     using faceq seq that by fastforce
  2186   have "rel_interior C \<noteq> {}"
  2232   have "rel_interior c \<noteq> {}"
  2187     using C \<open>C \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
  2233     using c \<open>c \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
  2188   then obtain x where "x \<in> rel_interior C" by auto
  2234   then obtain x where "x \<in> rel_interior c" by auto
       
  2235   have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
  2189   have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
  2236     by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
  2190     by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
  2237   then have xnot: "x \<notin> rel_interior S"
  2191   then have xnot: "x \<notin> rel_interior S"
  2238     by (metis IntI \<open>x \<in> rel_interior c\<close> c \<open>c \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
  2192     by (metis IntI \<open>x \<in> rel_interior C\<close> C \<open>C \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
  2239   then have "x \<in> S"
  2193   then have "x \<in> S"
  2240     using \<open>c \<subseteq> S\<close> \<open>x \<in> rel_interior c\<close> rel_interior_subset by auto
  2194     using \<open>C \<subseteq> S\<close> \<open>x \<in> rel_interior C\<close> rel_interior_subset by auto
  2241   then have xint: "x \<in> \<Inter>F"
  2195   then have xint: "x \<in> \<Inter>F"
  2242     using seq by blast
  2196     using seq by blast
  2243   have "F \<noteq> {}" using assms
  2197   have "F \<noteq> {}" using assms
  2244     by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial)
  2198     by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial)
  2245   then obtain i where "i \<in> F" "\<not> (a i \<bullet> x < b i)"
  2199   then obtain i where "i \<in> F" "\<not> (a i \<bullet> x < b i)"
  2247   with xint have "a i \<bullet> x = b i"
  2201   with xint have "a i \<bullet> x = b i"
  2248     by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq)
  2202     by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq)
  2249   have face: "S \<inter> {x. a i \<bullet> x = b i} face_of S"
  2203   have face: "S \<inter> {x. a i \<bullet> x = b i} face_of S"
  2250     by (simp add: "*" \<open>i \<in> F\<close>)
  2204     by (simp add: "*" \<open>i \<in> F\<close>)
  2251   show ?thesis
  2205   show ?thesis
  2252     apply (rule_tac h = i in that)
  2206   proof
  2253      apply (rule \<open>i \<in> F\<close>)
  2207     show "C \<subseteq> S \<inter> {x. a i \<bullet> x = b i}"
  2254     apply (rule subset_of_face_of [OF face \<open>c \<subseteq> S\<close>])
  2208       using subset_of_face_of [OF face \<open>C \<subseteq> S\<close>] \<open>a i \<bullet> x = b i\<close> \<open>x \<in> rel_interior C\<close> \<open>x \<in> S\<close> by blast
  2255     using \<open>a i \<bullet> x = b i\<close> \<open>x \<in> rel_interior c\<close> \<open>x \<in> S\<close> apply blast
  2209   qed fact
  2256     done
       
  2257 qed
  2210 qed
  2258 
  2211 
  2259 text\<open>Initial part of proof duplicates that above\<close>
  2212 text\<open>Initial part of proof duplicates that above\<close>
  2260 proposition face_of_polyhedron_explicit:
  2213 proposition face_of_polyhedron_explicit:
  2261   fixes S :: "'a :: euclidean_space set"
  2214   fixes S :: "'a :: euclidean_space set"
  2262   assumes "finite F"
  2215   assumes "finite F"
  2263       and seq: "S = affine hull S \<inter> \<Inter>F"
  2216       and seq: "S = affine hull S \<inter> \<Inter>F"
  2264       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2217       and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2265       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  2218       and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
  2266       and c: "c face_of S" and "c \<noteq> {}" "c \<noteq> S"
  2219       and C: "C face_of S" and "C \<noteq> {}" "C \<noteq> S"
  2267     shows "c = \<Inter>{S \<inter> {x. a h \<bullet> x = b h} | h. h \<in> F \<and> c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}"
  2220     shows "C = \<Inter>{S \<inter> {x. a h \<bullet> x = b h} | h. h \<in> F \<and> C \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}"
  2268 proof -
  2221 proof -
  2269   let ?ab = "\<lambda>h. {x. a h \<bullet> x = b h}"
  2222   let ?ab = "\<lambda>h. {x. a h \<bullet> x = b h}"
  2270   have "c \<subseteq> S" using \<open>c face_of S\<close>
  2223   have "C \<subseteq> S" using \<open>C face_of S\<close>
  2271     by (simp add: face_of_imp_subset)
  2224     by (simp add: face_of_imp_subset)
  2272   have "polyhedron S"
  2225   have "polyhedron S"
  2273     apply (simp add: polyhedron_Int_affine)
  2226     by (metis \<open>finite F\<close> faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq)
  2274     by (metis \<open>finite F\<close> faceq seq)
       
  2275   then have "convex S"
  2227   then have "convex S"
  2276     by (simp add: polyhedron_imp_convex)
  2228     by (simp add: polyhedron_imp_convex)
  2277   then have *: "(S \<inter> ?ab h) face_of S" if "h \<in> F" for h
  2229   then have *: "(S \<inter> ?ab h) face_of S" if "h \<in> F" for h
  2278     apply (rule face_of_Int_supporting_hyperplane_le)
  2230     using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce
  2279     using faceq seq that by fastforce
  2231   have "rel_interior C \<noteq> {}"
  2280   have "rel_interior c \<noteq> {}"
  2232     using C \<open>C \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
  2281     using c \<open>c \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
  2233   then obtain z where z: "z \<in> rel_interior C" by auto
  2282   then obtain z where z: "z \<in> rel_interior c" by auto
       
  2283   have rels: "rel_interior S = {z \<in> S. \<forall>h\<in>F. a h \<bullet> z < b h}"
  2234   have rels: "rel_interior S = {z \<in> S. \<forall>h\<in>F. a h \<bullet> z < b h}"
  2284     by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
  2235     by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
  2285   then have xnot: "z \<notin> rel_interior S"
  2236   then have xnot: "z \<notin> rel_interior S"
  2286     by (metis IntI \<open>z \<in> rel_interior c\<close> c \<open>c \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
  2237     by (metis IntI \<open>z \<in> rel_interior C\<close> C \<open>C \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
  2287   then have "z \<in> S"
  2238   then have "z \<in> S"
  2288     using \<open>c \<subseteq> S\<close> \<open>z \<in> rel_interior c\<close> rel_interior_subset by auto
  2239     using \<open>C \<subseteq> S\<close> \<open>z \<in> rel_interior C\<close> rel_interior_subset by auto
  2289   with seq have xint: "z \<in> \<Inter>F" by blast
  2240   with seq have xint: "z \<in> \<Inter>F" by blast
  2290   have "open (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
  2241   have "open (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
  2291     by (auto simp: \<open>finite F\<close> open_halfspace_lt open_INT)
  2242     by (auto simp: \<open>finite F\<close> open_halfspace_lt open_INT)
  2292   then obtain e where "0 < e"
  2243   then obtain e where "0 < e"
  2293                  "ball z e \<subseteq> (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
  2244                  "ball z e \<subseteq> (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
  2294     by (auto intro: openE [of _ z])
  2245     by (auto intro: openE [of _ z])
  2295   then have e: "\<And>h. \<lbrakk>h \<in> F; a h \<bullet> z < b h\<rbrakk> \<Longrightarrow> ball z e \<subseteq> {w. a h \<bullet> w < b h}"
  2246   then have e: "\<And>h. \<lbrakk>h \<in> F; a h \<bullet> z < b h\<rbrakk> \<Longrightarrow> ball z e \<subseteq> {w. a h \<bullet> w < b h}"
  2296     by blast
  2247     by blast
  2297   have "c \<subseteq> (S \<inter> ?ab h) \<longleftrightarrow> z \<in> S \<inter> ?ab h" if "h \<in> F" for h
  2248   have "C \<subseteq> (S \<inter> ?ab h) \<longleftrightarrow> z \<in> S \<inter> ?ab h" if "h \<in> F" for h
  2298   proof
  2249   proof
  2299     show "z \<in> S \<inter> ?ab h \<Longrightarrow> c \<subseteq> S \<inter> ?ab h"
  2250     show "z \<in> S \<inter> ?ab h \<Longrightarrow> C \<subseteq> S \<inter> ?ab h"
  2300       apply (rule subset_of_face_of [of _ S])
  2251       by (metis "*" Collect_cong IntI \<open>C \<subseteq> S\<close> empty_iff subset_of_face_of that z)
  2301       using that \<open>c \<subseteq> S\<close> \<open>z \<in> rel_interior c\<close>
       
  2302       using facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub]
       
  2303             unfolding facet_of_def
       
  2304       apply auto
       
  2305       done
       
  2306   next
  2252   next
  2307     show "c \<subseteq> S \<inter> ?ab h \<Longrightarrow> z \<in> S \<inter> ?ab h"
  2253     show "C \<subseteq> S \<inter> ?ab h \<Longrightarrow> z \<in> S \<inter> ?ab h"
  2308       using \<open>z \<in> rel_interior c\<close> rel_interior_subset by force
  2254       using \<open>z \<in> rel_interior C\<close> rel_interior_subset by force
  2309   qed
  2255   qed
  2310   then have **: "{S \<inter> ?ab h | h. h \<in> F \<and> c \<subseteq> S \<and> c \<subseteq> ?ab h} =
  2256   then have **: "{S \<inter> ?ab h | h. h \<in> F \<and> C \<subseteq> S \<and> C \<subseteq> ?ab h} =
  2311                  {S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<inter> ?ab h}"
  2257                  {S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<inter> ?ab h}"
  2312     by blast
  2258     by blast
  2313   have bsub: "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2259   have bsub: "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2314              \<subseteq> affine hull S \<inter> \<Inter>F \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2260              \<subseteq> affine hull S \<inter> \<Inter>F \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2315             if "i \<in> F" and i: "a i \<bullet> z = b i" for i
  2261             if "i \<in> F" and i: "a i \<bullet> z = b i" for i
  2333           by (rule_tac x="{x. a j \<bullet> x = b j}" in exI) (fastforce simp add: \<open>j \<in> F\<close>)
  2279           by (rule_tac x="{x. a j \<bullet> x = b j}" in exI) (fastforce simp add: \<open>j \<in> F\<close>)
  2334       qed
  2280       qed
  2335       then show ?thesis  by blast
  2281       then show ?thesis  by blast
  2336     qed
  2282     qed
  2337     have 1: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> affine hull S"
  2283     have 1: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> affine hull S"
  2338       apply (rule hull_mono)
  2284       using that \<open>z \<in> S\<close> by (intro hull_mono) auto
  2339       using that \<open>z \<in> S\<close> by auto
       
  2340     have 2: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2285     have 2: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2341           \<subseteq> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2286           \<subseteq> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2342       by (rule hull_minimal) (auto intro: affine_hyperplane)
  2287       by (rule hull_minimal) (auto intro: affine_hyperplane)
  2343     have 3: "ball z e \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> \<Inter>F"
  2288     have 3: "ball z e \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> \<Inter>F"
  2344       by (iprover intro: sub Inter_greatest)
  2289       by (iprover intro: sub Inter_greatest)
  2345     have *: "\<lbrakk>A \<subseteq> (B :: 'a set); A \<subseteq> C; E \<inter> C \<subseteq> D\<rbrakk> \<Longrightarrow> E \<inter> A \<subseteq> (B \<inter> D) \<inter> C"
  2290     have *: "\<lbrakk>A \<subseteq> (B :: 'a set); A \<subseteq> C; E \<inter> C \<subseteq> D\<rbrakk> \<Longrightarrow> E \<inter> A \<subseteq> (B \<inter> D) \<inter> C"
  2346              for A B C D E  by blast
  2291              for A B C D E  by blast
  2347     show ?thesis by (intro * 1 2 3)
  2292     show ?thesis by (intro * 1 2 3)
  2348   qed
  2293   qed
  2349   have "\<exists>h. h \<in> F \<and> c \<subseteq> ?ab h"
  2294   have "\<exists>h. h \<in> F \<and> C \<subseteq> ?ab h"
  2350     apply (rule face_of_polyhedron_subset_explicit [OF \<open>finite F\<close> seq faceq psub])
  2295     using assms
  2351     using assms by auto
  2296     by (metis face_of_polyhedron_subset_explicit [OF \<open>finite F\<close> seq faceq psub] le_inf_iff)
  2352   then have fac: "\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> c \<subseteq> S \<inter> ?ab h} face_of S"
  2297   then have fac: "\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> C \<subseteq> S \<inter> ?ab h} face_of S"
  2353     using * by (force simp: \<open>c \<subseteq> S\<close> intro: face_of_Inter)
  2298     using * by (force simp: \<open>C \<subseteq> S\<close> intro: face_of_Inter)
  2354   have red:
  2299   have red: "(\<And>a. P a \<Longrightarrow> T \<subseteq> S \<inter> \<Inter>{F X |X. P X}) \<Longrightarrow> T \<subseteq> \<Inter>{S \<inter> F X |X::'a set. P X}" for P T F   
  2355      "(\<And>a. P a \<Longrightarrow> T \<subseteq> S \<inter> \<Inter>{F x |x. P x}) \<Longrightarrow> T \<subseteq> \<Inter>{S \<inter> F x |x. P x}"
  2300     by blast
  2356      for P T F   by blast
       
  2357   have "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2301   have "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
  2358         \<subseteq> \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2302         \<subseteq> \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
  2359     apply (rule red)
  2303     by (rule red) (metis seq bsub)
  2360     apply (metis seq bsub)
       
  2361     done
       
  2362   with \<open>0 < e\<close> have zinrel: "z \<in> rel_interior
  2304   with \<open>0 < e\<close> have zinrel: "z \<in> rel_interior
  2363                     (\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<and> a h \<bullet> z = b h})"
  2305                     (\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<and> a h \<bullet> z = b h})"
  2364     by (auto simp: mem_rel_interior_ball \<open>z \<in> S\<close>)
  2306     by (auto simp: mem_rel_interior_ball \<open>z \<in> S\<close>)
  2365   show ?thesis
  2307   show ?thesis
  2366     apply (rule face_of_eq [OF c fac])
  2308     using z zinrel
  2367     using z zinrel apply (force simp: **)
  2309     by (intro face_of_eq [OF C fac]) (force simp: **)
  2368     done
       
  2369 qed
  2310 qed
  2370 
  2311 
  2371 
  2312 
  2372 subsection\<open>More general corollaries from the explicit representation\<close>
  2313 subsection\<open>More general corollaries from the explicit representation\<close>
  2373 
  2314 
  2374 corollary facet_of_polyhedron:
  2315 corollary facet_of_polyhedron:
  2375   assumes "polyhedron S" and "c facet_of S"
  2316   assumes "polyhedron S" and "C facet_of S"
  2376   obtains a b where "a \<noteq> 0" "S \<subseteq> {x. a \<bullet> x \<le> b}" "c = S \<inter> {x. a \<bullet> x = b}"
  2317   obtains a b where "a \<noteq> 0" "S \<subseteq> {x. a \<bullet> x \<le> b}" "C = S \<inter> {x. a \<bullet> x = b}"
  2377 proof -
  2318 proof -
  2378   obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
  2319   obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
  2379              and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
  2320              and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
  2380              and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
  2321              and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
  2381     using assms by (simp add: polyhedron_Int_affine_minimal) meson
  2322     using assms by (simp add: polyhedron_Int_affine_minimal) meson
  2382   then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2323   then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
  2383     by metis
  2324     by metis
  2384   obtain i where "i \<in> F" and c: "c = S \<inter> {x. a i \<bullet> x = b i}"
  2325   obtain i where "i \<in> F" and C: "C = S \<inter> {x. a i \<bullet> x = b i}"
  2385     using facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] assms
  2326     using facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] assms
  2386     by force
  2327     by force
  2387   moreover have ssub: "S \<subseteq> {x. a i \<bullet> x \<le> b i}"
  2328   moreover have ssub: "S \<subseteq> {x. a i \<bullet> x \<le> b i}"
  2388      apply (subst seq)
  2329      using \<open>i \<in> F\<close> ab by (subst seq) auto
  2389      using \<open>i \<in> F\<close> ab by auto
       
  2390   ultimately show ?thesis
  2330   ultimately show ?thesis
  2391     by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab)
  2331     by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab)
  2392 qed
  2332 qed
  2393 
  2333 
  2394 corollary face_of_polyhedron:
  2334 corollary face_of_polyhedron:
  2395   assumes "polyhedron S" and "c face_of S" and "c \<noteq> {}" and "c \<noteq> S"
  2335   assumes "polyhedron S" and "C face_of S" and "C \<noteq> {}" and "C \<noteq> S"
  2396     shows "c = \<Inter>{F. F facet_of S \<and> c \<subseteq> F}"
  2336     shows "C = \<Inter>{F. F facet_of S \<and> C \<subseteq> F}"
  2397 proof -
  2337 proof -
  2398   obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
  2338   obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
  2399              and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
  2339              and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
  2400              and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
  2340              and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
  2401     using assms by (simp add: polyhedron_Int_affine_minimal) meson
  2341     using assms by (simp add: polyhedron_Int_affine_minimal) meson
  2406     apply (auto simp: assms facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] cong: Collect_cong)
  2346     apply (auto simp: assms facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] cong: Collect_cong)
  2407     done
  2347     done
  2408 qed
  2348 qed
  2409 
  2349 
  2410 lemma face_of_polyhedron_subset_facet:
  2350 lemma face_of_polyhedron_subset_facet:
  2411   assumes "polyhedron S" and "c face_of S" and "c \<noteq> {}" and "c \<noteq> S"
  2351   assumes "polyhedron S" and "C face_of S" and "C \<noteq> {}" and "C \<noteq> S"
  2412   obtains F where "F facet_of S" "c \<subseteq> F"
  2352   obtains F where "F facet_of S" "C \<subseteq> F"
  2413 using face_of_polyhedron assms
  2353   using face_of_polyhedron assms
  2414 by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq)
  2354   by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq)
  2415 
  2355 
  2416 
  2356 
  2417 lemma exposed_face_of_polyhedron:
  2357 lemma exposed_face_of_polyhedron:
  2418   assumes "polyhedron S"
  2358   assumes "polyhedron S"
  2419     shows "F exposed_face_of S \<longleftrightarrow> F face_of S"
  2359     shows "F exposed_face_of S \<longleftrightarrow> F face_of S"
  2430     case False
  2370     case False
  2431     then have "{g. g facet_of S \<and> F \<subseteq> g} \<noteq> {}"
  2371     then have "{g. g facet_of S \<and> F \<subseteq> g} \<noteq> {}"
  2432       by (metis Collect_empty_eq_bot \<open>F face_of S\<close> assms empty_def face_of_polyhedron_subset_facet)
  2372       by (metis Collect_empty_eq_bot \<open>F face_of S\<close> assms empty_def face_of_polyhedron_subset_facet)
  2433     moreover have "\<And>T. \<lbrakk>T facet_of S; F \<subseteq> T\<rbrakk> \<Longrightarrow> T exposed_face_of S"
  2373     moreover have "\<And>T. \<lbrakk>T facet_of S; F \<subseteq> T\<rbrakk> \<Longrightarrow> T exposed_face_of S"
  2434       by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron)
  2374       by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron)
  2435     ultimately have "\<Inter>{fa.
  2375     ultimately have "\<Inter>{G. G facet_of S \<and> F \<subseteq> G} exposed_face_of S"
  2436        fa facet_of S \<and> F \<subseteq> fa} exposed_face_of S"
       
  2437       by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter)
  2376       by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter)
  2438     then show ?thesis
  2377     then show ?thesis
  2439       using False
  2378       using False \<open>F face_of S\<close> assms face_of_polyhedron by fastforce
  2440       apply (subst face_of_polyhedron [OF assms \<open>F face_of S\<close>], auto)
       
  2441       done
       
  2442   qed
  2379   qed
  2443 qed
  2380 qed
  2444 
  2381 
  2445 lemma face_of_polyhedron_polyhedron:
  2382 lemma face_of_polyhedron_polyhedron:
  2446   fixes S :: "'a :: euclidean_space set"
  2383   fixes S :: "'a :: euclidean_space set"
  2462     by (simp add: \<open>finite F\<close>)
  2399     by (simp add: \<open>finite F\<close>)
  2463   moreover have "{F. F face_of S} - {{}, S} \<subseteq> {\<Inter>{S \<inter> {x. a h \<bullet> x = b h} |h. h \<in> F'}| F'. F' \<in> Pow F}"
  2400   moreover have "{F. F face_of S} - {{}, S} \<subseteq> {\<Inter>{S \<inter> {x. a h \<bullet> x = b h} |h. h \<in> F'}| F'. F' \<in> Pow F}"
  2464     apply clarify
  2401     apply clarify
  2465     apply (rename_tac c)
  2402     apply (rename_tac c)
  2466     apply (drule face_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min, simplified], simp_all)
  2403     apply (drule face_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min, simplified], simp_all)
  2467     apply (erule ssubst)
       
  2468     apply (rule_tac x="{h \<in> F. c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}" in exI, auto)
  2404     apply (rule_tac x="{h \<in> F. c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}" in exI, auto)
  2469     done
  2405     done
  2470   ultimately show ?thesis
  2406   ultimately show ?thesis
  2471     by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset)
  2407     by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset)
  2472 qed
  2408 qed
  2475    "polyhedron S \<Longrightarrow> finite {F. F exposed_face_of S}"
  2411    "polyhedron S \<Longrightarrow> finite {F. F exposed_face_of S}"
  2476 using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce
  2412 using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce
  2477 
  2413 
  2478 lemma finite_polyhedron_extreme_points:
  2414 lemma finite_polyhedron_extreme_points:
  2479   fixes S :: "'a :: euclidean_space set"
  2415   fixes S :: "'a :: euclidean_space set"
  2480   shows "polyhedron S \<Longrightarrow> finite {v. v extreme_point_of S}"
  2416   assumes "polyhedron S" shows "finite {v. v extreme_point_of S}"
  2481 apply (simp add: face_of_singleton [symmetric])
  2417 proof -
  2482 apply (rule finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto)
  2418   have "finite {v. {v} face_of S}"
  2483 done
  2419     using assms by (intro finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto)
       
  2420   then show ?thesis
       
  2421     by (simp add: face_of_singleton)
       
  2422 qed
  2484 
  2423 
  2485 lemma finite_polyhedron_facets:
  2424 lemma finite_polyhedron_facets:
  2486   fixes S :: "'a :: euclidean_space set"
  2425   fixes S :: "'a :: euclidean_space set"
  2487   shows "polyhedron S \<Longrightarrow> finite {F. F facet_of S}"
  2426   shows "polyhedron S \<Longrightarrow> finite {F. F facet_of S}"
  2488 unfolding facet_of_def
  2427   unfolding facet_of_def
  2489 by (blast intro: finite_subset [OF _ finite_polyhedron_faces])
  2428   by (blast intro: finite_subset [OF _ finite_polyhedron_faces])
  2490 
  2429 
  2491 
  2430 
  2492 proposition rel_interior_of_polyhedron:
  2431 proposition rel_interior_of_polyhedron:
  2493   fixes S :: "'a :: euclidean_space set"
  2432   fixes S :: "'a :: euclidean_space set"
  2494   assumes "polyhedron S"
  2433   assumes "polyhedron S"
  2516       case 2
  2455       case 2
  2517       have "Collect ((\<in>) x) \<notin> Collect ((\<in>) (\<Union>{A. A facet_of S}))"
  2456       have "Collect ((\<in>) x) \<notin> Collect ((\<in>) (\<Union>{A. A facet_of S}))"
  2518         using xnot by fastforce
  2457         using xnot by fastforce
  2519       then have "F \<notin> Collect ((\<in>) h)"
  2458       then have "F \<notin> Collect ((\<in>) h)"
  2520         using 2 \<open>x \<in> S\<close> facet by blast
  2459         using 2 \<open>x \<in> S\<close> facet by blast
  2521       with \<open>h \<in> F\<close> have "\<Inter>F \<subseteq> S \<inter> {x. a h \<bullet> x = b h}" by blast
       
  2522       with 2 that \<open>x \<in> \<Inter>F\<close> show ?thesis
  2460       with 2 that \<open>x \<in> \<Inter>F\<close> show ?thesis
  2523         apply simp
  2461         by blast
  2524         apply (drule_tac x="\<Inter>F" in spec)
       
  2525         apply (simp add: facet)
       
  2526         apply (drule_tac x=h in spec)
       
  2527         using seq by auto
       
  2528       qed
  2462       qed
  2529   qed
  2463   qed
  2530   moreover have "\<exists>h\<in>F. a h \<bullet> x \<ge> b h" if "x \<in> \<Union>{F. F facet_of S}" for x
  2464   moreover have "\<exists>h\<in>F. a h \<bullet> x \<ge> b h" if "x \<in> \<Union>{F. F facet_of S}" for x
  2531     using that by (force simp: facet)
  2465     using that by (force simp: facet)
  2532   ultimately show ?thesis
  2466   ultimately show ?thesis
  2546 by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron)
  2480 by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron)
  2547 
  2481 
  2548 lemma rel_frontier_of_polyhedron_alt:
  2482 lemma rel_frontier_of_polyhedron_alt:
  2549   fixes S :: "'a :: euclidean_space set"
  2483   fixes S :: "'a :: euclidean_space set"
  2550   assumes "polyhedron S"
  2484   assumes "polyhedron S"
  2551     shows "rel_frontier S = \<Union> {F. F face_of S \<and> (F \<noteq> S)}"
  2485   shows "rel_frontier S = \<Union> {F. F face_of S \<and> F \<noteq> S}"
  2552 apply (rule subset_antisym)
  2486 proof
  2553   apply (force simp: rel_frontier_of_polyhedron facet_of_def assms)
  2487   show "rel_frontier S \<subseteq> \<Union> {F. F face_of S \<and> F \<noteq> S}"
  2554 using face_of_subset_rel_frontier by fastforce
  2488     by (force simp: rel_frontier_of_polyhedron facet_of_def assms)
       
  2489 qed (use face_of_subset_rel_frontier in fastforce)
  2555 
  2490 
  2556 
  2491 
  2557 text\<open>A characterization of polyhedra as having finitely many faces\<close>
  2492 text\<open>A characterization of polyhedra as having finitely many faces\<close>
  2558 
  2493 
  2559 proposition polyhedron_eq_finite_exposed_faces:
  2494 proposition polyhedron_eq_finite_exposed_faces:
  2589         by (simp add: \<open>S \<noteq> {}\<close> \<open>convex S\<close> rel_interior_eq_empty)
  2524         by (simp add: \<open>S \<noteq> {}\<close> \<open>convex S\<close> rel_interior_eq_empty)
  2590       then obtain c where c: "c \<in> rel_interior S" by auto
  2525       then obtain c where c: "c \<in> rel_interior S" by auto
  2591       with rel_interior_subset have "c \<in> S"  by blast
  2526       with rel_interior_subset have "c \<in> S"  by blast
  2592       have ccp: "closed_segment c p \<subseteq> affine hull S"
  2527       have ccp: "closed_segment c p \<subseteq> affine hull S"
  2593         by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE)
  2528         by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE)
       
  2529       have oS: "openin (top_of_set (closed_segment c p)) (closed_segment c p \<inter> rel_interior S)"
       
  2530         by (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp])
  2594       obtain x where xcl: "x \<in> closed_segment c p" and "x \<in> S" and xnot: "x \<notin> rel_interior S"
  2531       obtain x where xcl: "x \<in> closed_segment c p" and "x \<in> S" and xnot: "x \<notin> rel_interior S"
  2595         using connected_openin [of "closed_segment c p"]
  2532         using connected_openin [of "closed_segment c p"]
  2596         apply simp
  2533         apply simp
  2597         apply (drule_tac x="closed_segment c p \<inter> rel_interior S" in spec)
  2534         apply (drule_tac x="closed_segment c p \<inter> rel_interior S" in spec)
  2598         apply (erule impE)
  2535         apply (drule mp [OF _ oS])
  2599          apply (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp])
       
  2600         apply (drule_tac x="closed_segment c p \<inter> (- S)" in spec)
  2536         apply (drule_tac x="closed_segment c p \<inter> (- S)" in spec)
  2601         using rel_interior_subset \<open>closed S\<close> c \<open>p \<notin> S\<close> apply blast
  2537         using rel_interior_subset \<open>closed S\<close> c \<open>p \<notin> S\<close> apply blast
  2602         done
  2538         done
  2603       then obtain \<mu> where "0 \<le> \<mu>" "\<mu> \<le> 1" and xeq: "x = (1 - \<mu>) *\<^sub>R c + \<mu> *\<^sub>R p"
  2539       then obtain \<mu> where "0 \<le> \<mu>" "\<mu> \<le> 1" and xeq: "x = (1 - \<mu>) *\<^sub>R c + \<mu> *\<^sub>R p"
  2604         by (auto simp: in_segment)
  2540         by (auto simp: in_segment)
  2621         have sne: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> {}"
  2557         have sne: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> {}"
  2622           using \<open>x \<in> S\<close> by blast
  2558           using \<open>x \<in> S\<close> by blast
  2623         have sns: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> S"
  2559         have sns: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> S"
  2624           by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset)
  2560           by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset)
  2625         obtain h where "h \<in> F" "x \<in> h"
  2561         obtain h where "h \<in> F" "x \<in> h"
  2626           apply (rule_tac h="S \<inter> {y. d \<bullet> y = d \<bullet> x}" in that)
  2562           using F_def \<open>x \<in> S\<close> sex sns by blast
  2627           apply (simp_all add: F_def sex sne sns \<open>x \<in> S\<close>)
       
  2628           done
       
  2629         have abface: "{y. a h \<bullet> y = b h} face_of {y. a h \<bullet> y \<le> b h}"
  2563         have abface: "{y. a h \<bullet> y = b h} face_of {y. a h \<bullet> y \<le> b h}"
  2630           using hyperplane_face_of_halfspace_le by blast
  2564           using hyperplane_face_of_halfspace_le by blast
  2631         then have "c \<in> h"
  2565         then have "c \<in> h"
  2632           using face_ofD [OF abface xos] \<open>c \<in> S\<close> \<open>h \<in> F\<close> ab pint \<open>x \<in> h\<close> by blast
  2566           using face_ofD [OF abface xos] \<open>c \<in> S\<close> \<open>h \<in> F\<close> ab pint \<open>x \<in> h\<close> by blast
  2633         with c have "h \<inter> rel_interior S \<noteq> {}" by blast
  2567         with c have "h \<inter> rel_interior S \<noteq> {}" by blast
  2682 qed
  2616 qed
  2683 
  2617 
  2684 lemma polyhedron_negations:
  2618 lemma polyhedron_negations:
  2685   fixes S :: "'a :: euclidean_space set"
  2619   fixes S :: "'a :: euclidean_space set"
  2686   shows   "polyhedron S \<Longrightarrow> polyhedron(image uminus S)"
  2620   shows   "polyhedron S \<Longrightarrow> polyhedron(image uminus S)"
  2687   by (subst polyhedron_linear_image_eq)
  2621   by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus)
  2688     (auto simp: bij_uminus intro!: linear_uminus)
       
  2689 
  2622 
  2690 subsection\<open>Relation between polytopes and polyhedra\<close>
  2623 subsection\<open>Relation between polytopes and polyhedra\<close>
  2691 
  2624 
  2692 proposition polytope_eq_bounded_polyhedron:
  2625 proposition polytope_eq_bounded_polyhedron:
  2693   fixes S :: "'a :: euclidean_space set"
  2626   fixes S :: "'a :: euclidean_space set"
  2697   assume ?lhs
  2630   assume ?lhs
  2698   then show ?rhs
  2631   then show ?rhs
  2699     by (simp add: finite_polytope_faces polyhedron_eq_finite_faces
  2632     by (simp add: finite_polytope_faces polyhedron_eq_finite_faces
  2700                   polytope_imp_closed polytope_imp_convex polytope_imp_bounded)
  2633                   polytope_imp_closed polytope_imp_convex polytope_imp_bounded)
  2701 next
  2634 next
  2702   assume ?rhs then show ?lhs
  2635   assume R: ?rhs 
  2703     unfolding polytope_def
  2636   then have "finite {v. v extreme_point_of S}"
  2704     apply (rule_tac x="{v. v extreme_point_of S}" in exI)
  2637     by (simp add: finite_polyhedron_extreme_points)
  2705     apply (simp add: finite_polyhedron_extreme_points Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex)
  2638   moreover have "S = convex hull {v. v extreme_point_of S}"
  2706     done
  2639     using R by (simp add: Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex)
       
  2640   ultimately show ?lhs
       
  2641     unfolding polytope_def by blast
  2707 qed
  2642 qed
  2708 
  2643 
  2709 lemma polytope_Int:
  2644 lemma polytope_Int:
  2710   fixes S :: "'a :: euclidean_space set"
  2645   fixes S :: "'a :: euclidean_space set"
  2711   shows "\<lbrakk>polytope S; polytope T\<rbrakk> \<Longrightarrow> polytope(S \<inter> T)"
  2646   shows "\<lbrakk>polytope S; polytope T\<rbrakk> \<Longrightarrow> polytope(S \<inter> T)"
  2785     have [iff]: "finite S"
  2720     have [iff]: "finite S"
  2786       using assms using card.infinite by force
  2721       using assms using card.infinite by force
  2787     then have ccs: "closed (convex hull S)"
  2722     then have ccs: "closed (convex hull S)"
  2788       by (simp add: compact_imp_closed finite_imp_compact_convex_hull)
  2723       by (simp add: compact_imp_closed finite_imp_compact_convex_hull)
  2789     { fix x T
  2724     { fix x T
  2790       assume "finite T" "T \<subseteq> S" "int (card T) \<le> aff_dim S + 1" "x \<in> convex hull T"
  2725       assume "int (card T) \<le> aff_dim S + 1"  "finite T" "T \<subseteq> S""x \<in> convex hull T"
  2791       then have "S \<noteq> T"
  2726       then have "S \<noteq> T"
  2792         using True \<open>finite S\<close> aff_dim_le_card affine_independent_iff_card by fastforce
  2727         using True \<open>finite S\<close> aff_dim_le_card affine_independent_iff_card by fastforce
  2793       then obtain a where "a \<in> S" "a \<notin> T"
  2728       then obtain a where "a \<in> S" "a \<notin> T"
  2794         using \<open>T \<subseteq> S\<close> by blast
  2729         using \<open>T \<subseteq> S\<close> by blast
  2795       then have "x \<in> (\<Union>a\<in>S. convex hull (S - {a}))"
  2730       then have "\<exists>y\<in>S. x \<in> convex hull (S - {y})"
  2796         using True affine_independent_iff_card [of S]
  2731         using True affine_independent_iff_card [of S]
  2797         apply simp
  2732         by (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \<open>a \<notin> T\<close> \<open>T \<subseteq> S\<close> \<open>x \<in> convex hull T\<close> hull_mono insert_Diff_single subsetCE)
  2798         apply (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \<open>a \<notin> T\<close> \<open>T \<subseteq> S\<close> \<open>x \<in> convex hull T\<close>  hull_mono insert_Diff_single   subsetCE)
       
  2799         done
       
  2800     } note * = this
  2733     } note * = this
  2801     have 1: "convex hull S \<subseteq> (\<Union> a\<in>S. convex hull (S - {a}))"
  2734     have 1: "convex hull S \<subseteq> (\<Union> a\<in>S. convex hull (S - {a}))"
  2802       apply (subst caratheodory_aff_dim)
  2735       by (subst caratheodory_aff_dim) (blast dest: *)
  2803       apply (blast intro: *)
       
  2804       done
       
  2805     have 2: "\<Union>((\<lambda>a. convex hull (S - {a})) ` S) \<subseteq> convex hull S"
  2736     have 2: "\<Union>((\<lambda>a. convex hull (S - {a})) ` S) \<subseteq> convex hull S"
  2806       by (rule Union_least) (metis (no_types, lifting)  Diff_subset hull_mono imageE)
  2737       by (rule Union_least) (metis (no_types, lifting)  Diff_subset hull_mono imageE)
  2807     show ?thesis using True
  2738     show ?thesis using True
  2808       apply (simp add: segment_convex_hull frontier_def)
  2739       apply (simp add: segment_convex_hull frontier_def)
  2809       using interior_convex_hull_eq_empty [OF assms]
  2740       using interior_convex_hull_eq_empty [OF assms]
  2810       apply (simp add: closure_closed [OF ccs])
  2741       apply (simp add: closure_closed [OF ccs])
  2811       apply (rule subset_antisym)
  2742       using "1" "2" by auto
  2812       using 1 apply blast
       
  2813       using 2 apply blast
       
  2814       done
       
  2815 next
  2743 next
  2816   case False
  2744   case False
  2817   then have "frontier (convex hull S) = (convex hull S) - rel_interior(convex hull S)"
  2745   then have "frontier (convex hull S) = closure (convex hull S) - interior (convex hull S)"
  2818     apply (simp add: rel_boundary_of_convex_hull [symmetric] frontier_def)
  2746     by (simp add: rel_boundary_of_convex_hull frontier_def)
  2819     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)
  2747   also have "\<dots> = (convex hull S) - rel_interior(convex hull S)"
  2820   also have "... = \<Union>{convex hull (S - {a}) |a. a \<in> S}"
  2748     by (metis False aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior)
       
  2749   also have "\<dots> = \<Union>{convex hull (S - {a}) |a. a \<in> S}"
  2821   proof -
  2750   proof -
  2822     have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)"
  2751     have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)"
  2823       by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron)
  2752       by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron)
  2824     then show ?thesis
  2753     then show ?thesis
  2825       by (simp add: False rel_frontier_convex_hull_cases)
  2754       by (simp add: False rel_frontier_convex_hull_cases)
  2995     by (simp add: \<open>finite \<F>\<close> poly bounded_Union polytope_imp_bounded)
  2924     by (simp add: \<open>finite \<F>\<close> poly bounded_Union polytope_imp_bounded)
  2996   then obtain B where "B > 0" and B: "\<And>x. x \<in> \<Union>\<F> \<Longrightarrow> norm x < B"
  2925   then obtain B where "B > 0" and B: "\<And>x. x \<in> \<Union>\<F> \<Longrightarrow> norm x < B"
  2997     by (meson bounded_pos_less)
  2926     by (meson bounded_pos_less)
  2998   define C where "C \<equiv> {z \<in> \<int>. \<bar>z * e / 2 / real DIM('a)\<bar> \<le> B}"
  2927   define C where "C \<equiv> {z \<in> \<int>. \<bar>z * e / 2 / real DIM('a)\<bar> \<le> B}"
  2999   define I where "I \<equiv> \<Union>i \<in> Basis. \<Union>j \<in> C. { (i::'a, j * e / 2 / DIM('a)) }"
  2928   define I where "I \<equiv> \<Union>i \<in> Basis. \<Union>j \<in> C. { (i::'a, j * e / 2 / DIM('a)) }"
  3000   have "finite C"
  2929   have "C \<subseteq> {x \<in> \<int>. - B / (e / 2 / real DIM('a)) \<le> x \<and> x \<le> B / (e / 2 / real DIM('a))}"
  3001     using finite_int_segment [of "-B / (e / 2 / DIM('a))" "B / (e / 2 / DIM('a))"]
  2930     using \<open>0 < e\<close> by (auto simp: field_split_simps C_def)
  3002     apply (simp add: C_def)
  2931   then have "finite C"
  3003     apply (erule rev_finite_subset)
  2932     using finite_int_segment finite_subset by blast
  3004     using \<open>0 < e\<close>
       
  3005     apply (auto simp: field_split_simps)
       
  3006     done
       
  3007   then have "finite I"
  2933   then have "finite I"
  3008     by (simp add: I_def)
  2934     by (simp add: I_def)
  3009   obtain \<F>' where eq: "\<Union>\<F>' = \<Union>\<F>" and "finite \<F>'"
  2935   obtain \<F>' where eq: "\<Union>\<F>' = \<Union>\<F>" and "finite \<F>'"
  3010               and poly: "\<And>X. X \<in> \<F>' \<Longrightarrow> polytope X"
  2936               and poly: "\<And>X. X \<in> \<F>' \<Longrightarrow> polytope X"
  3011               and aff: "\<And>X. X \<in> \<F>' \<Longrightarrow> aff_dim X \<le> d"
  2937               and aff: "\<And>X. X \<in> \<F>' \<Longrightarrow> aff_dim X \<le> d"
  3013               and I: "\<And>X x y a b.  \<lbrakk>X \<in> \<F>'; x \<in> X; y \<in> X; (a,b) \<in> I\<rbrakk> \<Longrightarrow>
  2939               and I: "\<And>X x y a b.  \<lbrakk>X \<in> \<F>'; x \<in> X; y \<in> X; (a,b) \<in> I\<rbrakk> \<Longrightarrow>
  3014                                      a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or> a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b"
  2940                                      a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or> a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b"
  3015               and sub1: "\<And>C. C \<in> \<F>' \<Longrightarrow> \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
  2941               and sub1: "\<And>C. C \<in> \<F>' \<Longrightarrow> \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
  3016               and sub2: "\<And>C x. C \<in> \<F> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<F>' \<and> x \<in> D \<and> D \<subseteq> C"
  2942               and sub2: "\<And>C x. C \<in> \<F> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<F>' \<and> x \<in> D \<and> D \<subseteq> C"
  3017     apply (rule exE [OF cell_subdivision_lemma])
  2943     apply (rule exE [OF cell_subdivision_lemma])
  3018     using assms \<open>finite I\<close> apply auto
  2944     using assms \<open>finite I\<close> by auto
  3019     done
       
  3020   show ?thesis
  2945   show ?thesis
  3021   proof (rule_tac \<F>'="\<F>'" in that)
  2946   proof (rule_tac \<F>'="\<F>'" in that)
  3022     show "diameter X < e" if "X \<in> \<F>'" for X
  2947     show "diameter X < e" if "X \<in> \<F>'" for X
  3023     proof -
  2948     proof -
  3024       have "diameter X \<le> e/2"
  2949       have "diameter X \<le> e/2"
  3139    "n simplex S \<Longrightarrow> aff_dim S = n"
  3064    "n simplex S \<Longrightarrow> aff_dim S = n"
  3140   by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card)
  3065   by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card)
  3141 
  3066 
  3142 lemma zero_simplex_sing: "0 simplex {a}"
  3067 lemma zero_simplex_sing: "0 simplex {a}"
  3143   apply (simp add: simplex_def)
  3068   apply (simp add: simplex_def)
  3144   by (metis affine_independent_1 card.empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI)
  3069   using affine_independent_1 card_1_singleton_iff convex_hull_singleton by blast
  3145 
  3070 
  3146 lemma simplex_sing [simp]: "n simplex {a} \<longleftrightarrow> n = 0"
  3071 lemma simplex_sing [simp]: "n simplex {a} \<longleftrightarrow> n = 0"
  3147   using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast
  3072   using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast
  3148 
  3073 
  3149 lemma simplex_zero: "0 simplex S \<longleftrightarrow> (\<exists>a. S = {a})"
  3074 lemma simplex_zero: "0 simplex S \<longleftrightarrow> (\<exists>a. S = {a})"
  3150 apply (auto simp: )
  3075   by (metis aff_dim_eq_0 aff_dim_simplex simplex_sing)
  3151   using aff_dim_eq_0 aff_dim_simplex by blast
       
  3152 
  3076 
  3153 lemma one_simplex_segment: "a \<noteq> b \<Longrightarrow> 1 simplex closed_segment a b"
  3077 lemma one_simplex_segment: "a \<noteq> b \<Longrightarrow> 1 simplex closed_segment a b"
  3154   apply (simp add: simplex_def)
  3078   unfolding simplex_def
  3155   apply (rule_tac x="{a,b}" in exI)
  3079   by (rule_tac x="{a,b}" in exI) (auto simp: segment_convex_hull)
  3156   apply (auto simp: segment_convex_hull)
       
  3157   done
       
  3158 
  3080 
  3159 lemma simplex_segment_cases:
  3081 lemma simplex_segment_cases:
  3160    "(if a = b then 0 else 1) simplex closed_segment a b"
  3082    "(if a = b then 0 else 1) simplex closed_segment a b"
  3161   by (auto simp: one_simplex_segment)
  3083   by (auto simp: one_simplex_segment)
  3162 
  3084 
  3280               U [unfolded rel_frontier_def] tnot
  3202               U [unfolded rel_frontier_def] tnot
  3281             by (auto simp: closed_segment_eq_open)
  3203             by (auto simp: closed_segment_eq_open)
  3282           ultimately
  3204           ultimately
  3283           have "\<not>(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \<noteq> z"
  3205           have "\<not>(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \<noteq> z"
  3284             using that xt xu
  3206             using that xt xu
  3285             apply (simp add: between_mem_segment [symmetric])
  3207             by (meson between_antisym between_mem_segment between_trans_2 ends_in_segment(2))
  3286             by (metis between_commute between_trans_2 between_antisym)
       
  3287           then have "\<not> collinear {t, z, u}" if "x \<noteq> z"
  3208           then have "\<not> collinear {t, z, u}" if "x \<noteq> z"
  3288             by (auto simp: that collinear_between_cases between_commute)
  3209             by (auto simp: that collinear_between_cases between_commute)
  3289           moreover have "collinear {t, z, x}"
  3210           moreover have "collinear {t, z, x}"
  3290             by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt)
  3211             by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt)
  3291           moreover have "collinear {z, x, u}"
  3212           moreover have "collinear {z, x, u}"
  3327   proof (cases "n \<le> 1")
  3248   proof (cases "n \<le> 1")
  3328     case True
  3249     case True
  3329     have "\<And>s. \<lbrakk>n \<le> 1; s \<in> \<M>\<rbrakk> \<Longrightarrow> \<exists>m. m simplex s"
  3250     have "\<And>s. \<lbrakk>n \<le> 1; s \<in> \<M>\<rbrakk> \<Longrightarrow> \<exists>m. m simplex s"
  3330       using poly\<M> aff\<M> by (force intro: polytope_lowdim_imp_simplex)
  3251       using poly\<M> aff\<M> by (force intro: polytope_lowdim_imp_simplex)
  3331     then show ?thesis
  3252     then show ?thesis
  3332       unfolding simplicial_complex_def
  3253       unfolding simplicial_complex_def using True
  3333       apply (rule_tac x="\<M>" in exI)
  3254       by (rule_tac x="\<M>" in exI) (auto simp: less.prems)
  3334       using True by (auto simp: less.prems)
       
  3335   next
  3255   next
  3336     case False
  3256     case False
  3337     define \<S> where "\<S> \<equiv> {C \<in> \<M>. aff_dim C < n}"
  3257     define \<S> where "\<S> \<equiv> {C \<in> \<M>. aff_dim C < n}"
  3338     have "finite \<S>" "\<And>C. C \<in> \<S> \<Longrightarrow> polytope C" "\<And>C. C \<in> \<S> \<Longrightarrow> aff_dim C \<le> int (n - 1)"
  3258     have "finite \<S>" "\<And>C. C \<in> \<S> \<Longrightarrow> polytope C" "\<And>C. C \<in> \<S> \<Longrightarrow> aff_dim C \<le> int (n - 1)"
  3339       "\<And>C F. \<lbrakk>C \<in> \<S>; F face_of C\<rbrakk> \<Longrightarrow> F \<in> \<S>"
       
  3340       "\<And>C1 C2. \<lbrakk>C1 \<in> \<S>; C2 \<in> \<S>\<rbrakk>  \<Longrightarrow> C1 \<inter> C2 face_of C1"
  3259       "\<And>C1 C2. \<lbrakk>C1 \<in> \<S>; C2 \<in> \<S>\<rbrakk>  \<Longrightarrow> C1 \<inter> C2 face_of C1"
  3341       using less.prems
  3260       using less.prems by (auto simp: \<S>_def)
  3342           apply (auto simp: \<S>_def)
  3261     moreover have \<section>: "\<And>C F. \<lbrakk>C \<in> \<S>; F face_of C\<rbrakk> \<Longrightarrow> F \<in> \<S>"
  3343       by (metis aff_dim_subset face_of_imp_subset less_le not_le)
  3262       using less.prems unfolding \<S>_def 
  3344     with less.IH [of "n-1" \<S>] False
  3263       by (metis (no_types, lifting) mem_Collect_eq aff_dim_subset face_of_imp_subset less_le not_le)
  3345     obtain \<U> where "simplicial_complex \<U>"
  3264     ultimately obtain \<U> where "simplicial_complex \<U>"
  3346       and aff_dim\<U>: "\<And>K. K \<in> \<U> \<Longrightarrow> aff_dim K \<le> int (n - 1)"
  3265       and aff_dim\<U>: "\<And>K. K \<in> \<U> \<Longrightarrow> aff_dim K \<le> int (n - 1)"
  3347       and        "\<Union>\<U> = \<Union>\<S>"
  3266       and        "\<Union>\<U> = \<Union>\<S>"
  3348       and fin\<U>:  "\<And>C. C \<in> \<S> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<U> \<and> C = \<Union>F"
  3267       and fin\<U>:  "\<And>C. C \<in> \<S> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<U> \<and> C = \<Union>F"
  3349       and C\<U>:    "\<And>K. K \<in> \<U> \<Longrightarrow> \<exists>C. C \<in> \<S> \<and> K \<subseteq> C"
  3268       and C\<U>:    "\<And>K. K \<in> \<U> \<Longrightarrow> \<exists>C. C \<in> \<S> \<and> K \<subseteq> C"
  3350       by auto
  3269       using less.IH [of "n-1" \<S>] False by auto
  3351     then have "finite \<U>"
  3270     then have "finite \<U>"
  3352       and simpl\<U>: "\<And>S. S \<in> \<U> \<Longrightarrow> \<exists>n. n simplex S"
  3271       and simpl\<U>: "\<And>S. S \<in> \<U> \<Longrightarrow> \<exists>n. n simplex S"
  3353       and face\<U>:  "\<And>F S. \<lbrakk>S \<in> \<U>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<U>"
  3272       and face\<U>:  "\<And>F S. \<lbrakk>S \<in> \<U>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<U>"
  3354       and faceI\<U>: "\<And>S S'. \<lbrakk>S \<in> \<U>; S' \<in> \<U>\<rbrakk> \<Longrightarrow> (S \<inter> S') face_of S"
  3273       and faceI\<U>: "\<And>S S'. \<lbrakk>S \<in> \<U>; S' \<in> \<U>\<rbrakk> \<Longrightarrow> (S \<inter> S') face_of S"
  3355       by (auto simp: simplicial_complex_def)
  3274       by (auto simp: simplicial_complex_def)
  3390         finally have "K \<subseteq> C" .
  3309         finally have "K \<subseteq> C" .
  3391         have "L \<inter> C face_of C"
  3310         have "L \<inter> C face_of C"
  3392           using \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> intface\<M> by (simp add: inf_commute)
  3311           using \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> intface\<M> by (simp add: inf_commute)
  3393         moreover have "L \<inter> C \<noteq> C"
  3312         moreover have "L \<inter> C \<noteq> C"
  3394           using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close>
  3313           using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close>
  3395           apply (clarsimp simp: \<N>_def \<S>_def)
  3314           by (metis (mono_tags, lifting) \<N>_def \<S>_def intface\<M> mem_Collect_eq not_le order_refl \<section>)
  3396           by (metis aff_dim_subset inf_le1 not_le)
       
  3397         moreover have "K \<subseteq> L \<inter> C"
  3315         moreover have "K \<subseteq> L \<inter> C"
  3398           using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> \<open>K \<subseteq> C\<close> \<open>K \<subseteq> L\<close>
  3316           using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> \<open>K \<subseteq> C\<close> \<open>K \<subseteq> L\<close> by (auto simp: \<N>_def \<S>_def)
  3399           by (auto simp: \<N>_def \<S>_def)
       
  3400         ultimately show ?thesis using that by metis
  3317         ultimately show ?thesis using that by metis
  3401       qed
  3318       qed
  3402       have "affine hull F \<inter> rel_interior C = {}"
  3319       have "affine hull F \<inter> rel_interior C = {}"
  3403         by (rule affine_hull_face_of_disjoint_rel_interior [OF \<open>convex C\<close> F \<open>F \<noteq> C\<close>])
  3320         by (rule affine_hull_face_of_disjoint_rel_interior [OF \<open>convex C\<close> F \<open>F \<noteq> C\<close>])
  3404       with hull_mono [OF \<open>K \<subseteq> F\<close>]
  3321       with hull_mono [OF \<open>K \<subseteq> F\<close>]
  3502             show "C \<inter> D \<noteq> C"
  3419             show "C \<inter> D \<noteq> C"
  3503               by (metis (mono_tags, lifting) Int_lower2 \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<S>\<close> aff_dim_subset mem_Collect_eq not_le)
  3420               by (metis (mono_tags, lifting) Int_lower2 \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<S>\<close> aff_dim_subset mem_Collect_eq not_le)
  3504           qed
  3421           qed
  3505           finally have DC: "D \<inter> rel_interior C = {}" .
  3422           finally have DC: "D \<inter> rel_interior C = {}" .
  3506           have eq: "X \<inter> convex hull (insert ?z K) = X \<inter> convex hull K"
  3423           have eq: "X \<inter> convex hull (insert ?z K) = X \<inter> convex hull K"
  3507             apply (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex C\<close> \<open>K \<subseteq> C\<close> z])
  3424           proof (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex C\<close> \<open>K \<subseteq> C\<close> z])
  3508             using DC by (meson \<open>X \<subseteq> D\<close> disjnt_def disjnt_subset1)
  3425             show "disjnt X (rel_interior C)"
       
  3426               using DC by (meson \<open>X \<subseteq> D\<close> disjnt_def disjnt_subset1)
       
  3427           qed
  3509           obtain I where I: "\<not> affine_dependent I"
  3428           obtain I where I: "\<not> affine_dependent I"
  3510             and Keq: "K = convex hull I" and [simp]: "convex hull K = K"
  3429             and Keq: "K = convex hull I" and [simp]: "convex hull K = K"
  3511             using "*" \<open>K \<in> \<U>\<close> by force
  3430             using "*" \<open>K \<in> \<U>\<close> by force
  3512           then have "?z \<notin> affine hull I"
  3431           then have "?z \<notin> affine hull I"
  3513             using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull z by blast
  3432             using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull z by blast
  3556             show ?thesis
  3475             show ?thesis
  3557             proof (cases "D=C")
  3476             proof (cases "D=C")
  3558               case True
  3477               case True
  3559               then have "L \<subseteq> rel_frontier C"
  3478               then have "L \<subseteq> rel_frontier C"
  3560                 using \<open>L \<subseteq> rel_frontier D\<close> by auto
  3479                 using \<open>L \<subseteq> rel_frontier D\<close> by auto
  3561               show ?thesis
  3480               have "convex hull insert (SOME z. z \<in> rel_interior C) (K \<inter> L) face_of
  3562                 apply (simp add: X Y True)
  3481                     convex hull insert (SOME z. z \<in> rel_interior C) K"
  3563                 apply (simp add: convex_hull_insert_Int_eq [OF z] \<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> \<open>convex C\<close> \<open>convex K\<close> \<open>convex L\<close>)
  3482                 by (metis face_of_polytope_insert2 "*" IntI \<open>C \<in> \<N>\<close> aff_independent_finite ahK_C_disjoint empty_iff faceI\<U> polytope_def z \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close>\<open>K \<subseteq> rel_frontier C\<close>)
  3564                 using face_of_polytope_insert2
  3483               then show ?thesis
  3565                 by (metis "*" IntI \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close>\<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> aff_independent_finite ahK_C_disjoint empty_iff faceI\<U> polytope_convex_hull z)
  3484                 using True X Y \<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> \<open>convex C\<close> \<open>convex K\<close> \<open>convex L\<close> convex_hull_insert_Int_eq z by force
  3566             next
  3485             next
  3567               case False
  3486               case False
  3568               have "convex D"
  3487               have "convex D"
  3569                 by (simp add: \<open>D \<in> \<N>\<close> convex\<N>)
  3488                 by (simp add: \<open>D \<in> \<N>\<close> convex\<N>)
  3570               have "K \<subseteq> C"
  3489               have "K \<subseteq> C"
  3588                   by (metis Int_commute False face_of_aff_dim_lt inf.idem inf_le1 intface\<M> not_le poly\<M> polytope_imp_convex)
  3507                   by (metis Int_commute False face_of_aff_dim_lt inf.idem inf_le1 intface\<M> not_le poly\<M> polytope_imp_convex)
  3589               qed
  3508               qed
  3590               finally have CD: "C \<inter> (rel_interior D) = {}" .
  3509               finally have CD: "C \<inter> (rel_interior D) = {}" .
  3591               have zKC: "(convex hull insert ?z K) \<subseteq> C"
  3510               have zKC: "(convex hull insert ?z K) \<subseteq> C"
  3592                 by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed convex\<N> hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z)
  3511                 by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed convex\<N> hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z)
  3593               have eq: "convex hull (insert ?z K) \<inter> convex hull (insert ?w L) =
  3512               have "disjnt (convex hull insert (SOME z. z \<in> rel_interior C) K) (rel_interior D)"
  3594                           convex hull (insert ?z K) \<inter> convex hull L"
  3513                 using zKC CD by (force simp: disjnt_def)
  3595                 apply (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex D\<close> \<open>L \<subseteq> D\<close> w])
  3514               then have eq: "convex hull (insert ?z K) \<inter> convex hull (insert ?w L) =
  3596                 using zKC CD apply (force simp: disjnt_def)
  3515                              convex hull (insert ?z K) \<inter> convex hull L"
  3597                 done
  3516                 by (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex D\<close> \<open>L \<subseteq> D\<close> w])
  3598               have ch_id: "convex hull K = K" "convex hull L = L"
  3517               have ch_id: "convex hull K = K" "convex hull L = L"
  3599                 using "*" \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> hull_same by auto
  3518                 using "*" \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> hull_same by auto
  3600               have "convex C"
  3519               have "convex C"
  3601                 by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
  3520                 by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
  3602               have "convex hull (insert ?z K) \<inter> L = L \<inter> convex hull (insert ?z K)"
  3521               have "convex hull (insert ?z K) \<inter> L = L \<inter> convex hull (insert ?z K)"
  3717         then have "\<Union>?F \<subseteq> C"
  3636         then have "\<Union>?F \<subseteq> C"
  3718           by auto
  3637           by auto
  3719         ultimately show ?thesis
  3638         ultimately show ?thesis
  3720           by blast
  3639           by blast
  3721       qed
  3640       qed
  3722 
       
  3723       have "(\<exists>C. C \<in> \<M> \<and> L \<subseteq> C) \<and> aff_dim L \<le> int n"  if "L \<in> \<U> \<union> ?\<T>" for L
  3641       have "(\<exists>C. C \<in> \<M> \<and> L \<subseteq> C) \<and> aff_dim L \<le> int n"  if "L \<in> \<U> \<union> ?\<T>" for L
  3724         using that
  3642         using that
  3725       proof
  3643       proof
  3726         assume "L \<in> \<U>"
  3644         assume "L \<in> \<U>"
  3727         then show ?thesis
  3645         then show ?thesis
  3800   qed
  3718   qed
  3801   moreover
  3719   moreover
  3802   have "\<Union>(\<Union>C\<in>\<M>. {F. F face_of C}) = \<Union>\<M>"
  3720   have "\<Union>(\<Union>C\<in>\<M>. {F. F face_of C}) = \<Union>\<M>"
  3803     using face_of_imp_subset face by blast
  3721     using face_of_imp_subset face by blast
  3804   ultimately show ?thesis
  3722   ultimately show ?thesis
  3805     apply clarify
  3723     using face_of_imp_subset n
  3806     apply (rule that, assumption+)
  3724     by (fastforce intro!: that simp add: poly face_of_refl polytope_imp_convex)
  3807        using n apply blast
       
  3808       apply (simp_all add: poly face_of_refl polytope_imp_convex)
       
  3809     using face_of_imp_subset by fastforce
       
  3810 next
  3725 next
  3811   case False
  3726   case False
  3812   then have m1: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C = -1"
  3727   then have m1: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C = -1"
  3813     by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less)
  3728     by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less)
  3814   then have face\<M>: "\<And>F S. \<lbrakk>S \<in> \<M>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<M>"
  3729   then have face\<M>: "\<And>F S. \<lbrakk>S \<in> \<M>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<M>"