new material: Colinearity, convex sets, polytopes
authorpaulson <lp15@cam.ac.uk>
Wed, 19 Jul 2017 16:41:26 +0100
changeset 66287 005a30862ed0
parent 66286 1c977b13414f
child 66288 e5995950b98a
new material: Colinearity, convex sets, polytopes
src/HOL/Analysis/Convex_Euclidean_Space.thy
src/HOL/Analysis/Homeomorphism.thy
src/HOL/Analysis/Linear_Algebra.thy
src/HOL/Analysis/Polytope.thy
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Tue Jul 18 11:35:32 2017 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Wed Jul 19 16:41:26 2017 +0100
@@ -2704,6 +2704,13 @@
   qed
 qed
 
+lemma convex_hull_insert_alt:
+   "convex hull (insert a S) =
+      (if S = {} then {a}
+      else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> convex hull S})"
+  apply (auto simp: convex_hull_insert)
+  using diff_eq_eq apply fastforce
+  by (metis add.group_left_neutral add_le_imp_le_diff diff_add_cancel)
 
 subsubsection \<open>Explicit expression for convex hull\<close>
 
@@ -3271,13 +3278,13 @@
 
 subsection \<open>Some Properties of Affine Dependent Sets\<close>
 
-lemma affine_independent_0: "\<not> affine_dependent {}"
+lemma affine_independent_0 [simp]: "\<not> affine_dependent {}"
   by (simp add: affine_dependent_def)
 
-lemma affine_independent_1: "\<not> affine_dependent {a}"
+lemma affine_independent_1 [simp]: "\<not> affine_dependent {a}"
   by (simp add: affine_dependent_def)
 
-lemma affine_independent_2: "\<not> affine_dependent {a,b}"
+lemma affine_independent_2 [simp]: "\<not> affine_dependent {a,b}"
   by (simp add: affine_dependent_def insert_Diff_if hull_same)
 
 lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) `  S) = (\<lambda>x. a + x) ` (affine hull S)"
@@ -7806,6 +7813,7 @@
     by (metis image_comp convex_translation)
 qed
 
+
 lemmas convex_segment = convex_closed_segment convex_open_segment
 
 lemma connected_segment [iff]:
@@ -7836,6 +7844,36 @@
     by (auto intro: rel_interior_closure_convex_shrink)
 qed
 
+lemma convex_hull_insert_segments:
+   "convex hull (insert a S) =
+    (if S = {} then {a} else  \<Union>x \<in> convex hull S. closed_segment a x)"
+  by (force simp add: convex_hull_insert_alt in_segment)
+
+lemma Int_convex_hull_insert_rel_exterior:
+  fixes z :: "'a::euclidean_space"
+  assumes "convex C" "T \<subseteq> C" and z: "z \<in> rel_interior C" and dis: "disjnt S (rel_interior C)"
+  shows "S \<inter> (convex hull (insert z T)) = S \<inter> (convex hull T)" (is "?lhs = ?rhs")
+proof
+  have "T = {} \<Longrightarrow> z \<notin> S"
+    using dis z by (auto simp add: disjnt_def)
+  then show "?lhs \<subseteq> ?rhs"
+  proof (clarsimp simp add: convex_hull_insert_segments)
+    fix x y
+    assume "x \<in> S" and y: "y \<in> convex hull T" and "x \<in> closed_segment z y"
+    have "y \<in> closure C"
+      by (metis y \<open>convex C\<close> \<open>T \<subseteq> C\<close> closure_subset contra_subsetD convex_hull_eq hull_mono)
+    moreover have "x \<notin> rel_interior C"
+      by (meson \<open>x \<in> S\<close> dis disjnt_iff)
+    moreover have "x \<in> open_segment z y \<union> {z, y}"
+      using \<open>x \<in> closed_segment z y\<close> closed_segment_eq_open by blast
+    ultimately show "x \<in> convex hull T"
+      using rel_interior_closure_convex_segment [OF \<open>convex C\<close> z]
+      using y z by blast
+  qed
+  show "?rhs \<subseteq> ?lhs"
+    by (meson hull_mono inf_mono subset_insertI subset_refl)
+qed
+
 subsection\<open>More results about segments\<close>
 
 lemma dist_half_times2:
@@ -8210,6 +8248,24 @@
   shows "between (X, B) Y \<longleftrightarrow> between (A, Y) X"
 using assms by (auto simp add: between)
 
+lemma between_translation [simp]: "between (a + y,a + z) (a + x) \<longleftrightarrow> between (y,z) x"
+  by (auto simp: between_def)
+
+lemma between_trans_2:
+  fixes a :: "'a :: euclidean_space"
+  shows "\<lbrakk>between (b,c) a; between (a,b) d\<rbrakk> \<Longrightarrow> between (c,d) a"
+  by (metis between_commute between_swap between_trans)
+
+lemma between_scaleR_lift [simp]:
+  fixes v :: "'a::euclidean_space"
+  shows "between (a *\<^sub>R v, b *\<^sub>R v) (c *\<^sub>R v) \<longleftrightarrow> v = 0 \<or> between (a, b) c"
+  by (simp add: between dist_norm scaleR_left_diff_distrib [symmetric] distrib_right [symmetric])
+
+lemma between_1:
+  fixes x::real
+  shows "between (a,b) x \<longleftrightarrow> (a \<le> x \<and> x \<le> b) \<or> (b \<le> x \<and> x \<le> a)"
+  by (auto simp: between_mem_segment closed_segment_eq_real_ivl)
+
 
 subsection \<open>Shrinking towards the interior of a convex set\<close>
 
@@ -11527,6 +11583,24 @@
   by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans
     convex_hull_subset_affine_hull Diff_subset collinear_affine_hull)
 
+lemma collinear_between_cases:
+  fixes c :: "'a::euclidean_space"
+  shows "collinear {a,b,c} \<longleftrightarrow> between (b,c) a \<or> between (c,a) b \<or> between (a,b) c"
+         (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then obtain u v where uv: "\<And>x. x \<in> {a, b, c} \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v"
+    by (auto simp: collinear_alt)
+  show ?rhs
+    using uv [of a] uv [of b] uv [of c] by (auto simp: between_1)
+next
+  assume ?rhs
+  then show ?lhs
+    unfolding between_mem_convex_hull
+    by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull)
+qed
+
+
 lemma subset_continuous_image_segment_1:
   fixes f :: "'a::euclidean_space \<Rightarrow> real"
   assumes "continuous_on (closed_segment a b) f"
@@ -12401,6 +12475,145 @@
     by (simp add: continuous_on_closed * closedin_imp_subset)
 qed
 
+subsection\<open>Trivial fact: convexity equals connectedness for collinear sets\<close>
+
+lemma convex_connected_collinear:
+  fixes S :: "'a::euclidean_space set"
+  assumes "collinear S"
+    shows "convex S \<longleftrightarrow> connected S"
+proof
+  assume "convex S"
+  then show "connected S"
+    using convex_connected by blast
+next
+  assume S: "connected S"
+  show "convex S"
+  proof (cases "S = {}")
+    case True
+    then show ?thesis by simp
+  next
+    case False
+    then obtain a where "a \<in> S" by auto
+    have "collinear (affine hull S)"
+      by (simp add: assms collinear_affine_hull_collinear)
+    then obtain z where "z \<noteq> 0" "\<And>x. x \<in> affine hull S \<Longrightarrow> \<exists>c. x - a = c *\<^sub>R z"
+      by (meson \<open>a \<in> S\<close> collinear hull_inc)
+    then obtain f where f: "\<And>x. x \<in> affine hull S \<Longrightarrow> x - a = f x *\<^sub>R z"
+      by metis
+    then have inj_f: "inj_on f (affine hull S)"
+      by (metis diff_add_cancel inj_onI)
+    have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \<in> affine hull S" and y: "y \<in> affine hull S" for x y
+    proof -
+      have "f x *\<^sub>R z = x - a"
+        by (simp add: f hull_inc x)
+      moreover have "f y *\<^sub>R z = y - a"
+        by (simp add: f hull_inc y)
+      ultimately show ?thesis
+        by (simp add: scaleR_left.diff)
+    qed
+    have cont_f: "continuous_on (affine hull S) f"
+      apply (clarsimp simp: dist_norm continuous_on_iff diff)
+      by (metis \<open>z \<noteq> 0\<close> mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff)
+    then have conn_fS: "connected (f ` S)"
+      by (meson S connected_continuous_image continuous_on_subset hull_subset)
+    show ?thesis
+    proof (clarsimp simp: convex_contains_segment)
+      fix x y z
+      assume "x \<in> S" "y \<in> S" "z \<in> closed_segment x y"
+      have False if "z \<notin> S"
+      proof -
+        have "f ` (closed_segment x y) = closed_segment (f x) (f y)"
+          apply (rule continuous_injective_image_segment_1)
+          apply (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f])
+          by (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f])
+        then have fz: "f z \<in> closed_segment (f x) (f y)"
+          using \<open>z \<in> closed_segment x y\<close> by blast
+        have "z \<in> affine hull S"
+          by (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>z \<in> closed_segment x y\<close> convex_affine_hull convex_contains_segment hull_inc subset_eq)
+        then have fz_notin: "f z \<notin> f ` S"
+          using hull_subset inj_f inj_onD that by fastforce
+        moreover have "{..<f z} \<inter> f ` S \<noteq> {}" "{f z<..} \<inter> f ` S \<noteq> {}"
+        proof -
+          have "{..<f z} \<inter> f ` {x,y} \<noteq> {}"  "{f z<..} \<inter> f ` {x,y} \<noteq> {}"
+            using fz fz_notin \<open>x \<in> S\<close> \<open>y \<in> S\<close>
+             apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm)
+             apply (metis image_eqI less_eq_real_def)+
+            done
+          then show "{..<f z} \<inter> f ` S \<noteq> {}" "{f z<..} \<inter> f ` S \<noteq> {}"
+            using \<open>x \<in> S\<close> \<open>y \<in> S\<close> by blast+
+        qed
+        ultimately show False
+          using connectedD [OF conn_fS, of "{..<f z}" "{f z<..}"] by force
+      qed
+      then show "z \<in> S" by meson
+    qed
+  qed
+qed
+
+lemma compact_convex_collinear_segment_alt:
+  fixes S :: "'a::euclidean_space set"
+  assumes "S \<noteq> {}" "compact S" "connected S" "collinear S"
+  obtains a b where "S = closed_segment a b"
+proof -
+  obtain \<xi> where "\<xi> \<in> S" using \<open>S \<noteq> {}\<close> by auto
+  have "collinear (affine hull S)"
+    by (simp add: assms collinear_affine_hull_collinear)
+  then obtain z where "z \<noteq> 0" "\<And>x. x \<in> affine hull S \<Longrightarrow> \<exists>c. x - \<xi> = c *\<^sub>R z"
+    by (meson \<open>\<xi> \<in> S\<close> collinear hull_inc)
+  then obtain f where f: "\<And>x. x \<in> affine hull S \<Longrightarrow> x - \<xi> = f x *\<^sub>R z"
+    by metis
+  let ?g = "\<lambda>r. r *\<^sub>R z + \<xi>"
+  have gf: "?g (f x) = x" if "x \<in> affine hull S" for x
+    by (metis diff_add_cancel f that)
+  then have inj_f: "inj_on f (affine hull S)"
+    by (metis inj_onI)
+  have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \<in> affine hull S" and y: "y \<in> affine hull S" for x y
+  proof -
+    have "f x *\<^sub>R z = x - \<xi>"
+      by (simp add: f hull_inc x)
+    moreover have "f y *\<^sub>R z = y - \<xi>"
+      by (simp add: f hull_inc y)
+    ultimately show ?thesis
+      by (simp add: scaleR_left.diff)
+  qed
+  have cont_f: "continuous_on (affine hull S) f"
+    apply (clarsimp simp: dist_norm continuous_on_iff diff)
+    by (metis \<open>z \<noteq> 0\<close> mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff)
+  then have "connected (f ` S)"
+    by (meson \<open>connected S\<close> connected_continuous_image continuous_on_subset hull_subset)
+  moreover have "compact (f ` S)"
+    by (meson \<open>compact S\<close> compact_continuous_image_eq cont_f hull_subset inj_f)
+  ultimately obtain x y where "f ` S = {x..y}"
+    by (meson connected_compact_interval_1)
+  then have fS_eq: "f ` S = closed_segment x y"
+    using \<open>S \<noteq> {}\<close> closed_segment_eq_real_ivl by auto
+  obtain a b where "a \<in> S" "f a = x" "b \<in> S" "f b = y"
+    by (metis (full_types) ends_in_segment fS_eq imageE)
+  have "f ` (closed_segment a b) = closed_segment (f a) (f b)"
+    apply (rule continuous_injective_image_segment_1)
+     apply (meson \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f])
+    by (meson \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f])
+  then have "f ` (closed_segment a b) = f ` S"
+    by (simp add: \<open>f a = x\<close> \<open>f b = y\<close> fS_eq)
+  then have "?g ` f ` (closed_segment a b) = ?g ` f ` S"
+    by simp
+  moreover have "(\<lambda>x. f x *\<^sub>R z + \<xi>) ` closed_segment a b = closed_segment a b"
+    apply safe
+     apply (metis (mono_tags, hide_lams) \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment gf hull_inc subsetCE)
+    by (metis (mono_tags, lifting) \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE)
+  ultimately have "closed_segment a b = S"
+    using gf by (simp add: image_comp o_def hull_inc cong: image_cong)
+  then show ?thesis
+    using that by blast
+qed
+
+lemma compact_convex_collinear_segment:
+  fixes S :: "'a::euclidean_space set"
+  assumes "S \<noteq> {}" "compact S" "convex S" "collinear S"
+  obtains a b where "S = closed_segment a b"
+  using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast
+
+
 lemma proper_map_from_compact:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S"
--- a/src/HOL/Analysis/Homeomorphism.thy	Tue Jul 18 11:35:32 2017 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy	Wed Jul 19 16:41:26 2017 +0100
@@ -157,6 +157,65 @@
     by (simp add: \<open>interior S = rel_interior S\<close> frontier_def rel_frontier_def that)
 qed
 
+
+lemma segment_to_rel_frontier_aux:
+  fixes x :: "'a::euclidean_space"
+  assumes "convex S" "bounded S" and x: "x \<in> rel_interior S" and y: "y \<in> S" and xy: "x \<noteq> y"
+  obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
+                   "open_segment x z \<subseteq> rel_interior S"
+proof -
+  have "x + (y - x) \<in> affine hull S"
+    using hull_inc [OF y] by auto
+  then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \<in> rel_frontier S"
+                  and di: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (x + e *\<^sub>R (y-x)) \<in> rel_interior S"
+    by (rule ray_to_rel_frontier [OF \<open>bounded S\<close> x]) (use xy in auto)
+  show ?thesis
+  proof
+    show "x + d *\<^sub>R (y - x) \<in> rel_frontier S"
+      by (simp add: df)
+  next
+    have "open_segment x y \<subseteq> rel_interior S"
+      using rel_interior_closure_convex_segment [OF \<open>convex S\<close> x] closure_subset y by blast
+    moreover have "x + d *\<^sub>R (y - x) \<in> open_segment x y" if "d < 1"
+      using xy
+      apply (auto simp: in_segment)
+      apply (rule_tac x="d" in exI)
+      using \<open>0 < d\<close> that apply (auto simp: divide_simps algebra_simps)
+      done
+    ultimately have "1 \<le> d"
+      using df rel_frontier_def by fastforce
+    moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x"
+      by (metis \<open>0 < d\<close> add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one)
+    ultimately show "y \<in> closed_segment x (x + d *\<^sub>R (y - x))"
+      apply (auto simp: in_segment)
+      apply (rule_tac x="1/d" in exI)
+      apply (auto simp: divide_simps algebra_simps)
+      done
+  next
+    show "open_segment x (x + d *\<^sub>R (y - x)) \<subseteq> rel_interior S"
+      apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> x])
+      using df rel_frontier_def by auto
+  qed
+qed
+
+lemma segment_to_rel_frontier:
+  fixes x :: "'a::euclidean_space"
+  assumes S: "convex S" "bounded S" and x: "x \<in> rel_interior S"
+      and y: "y \<in> S" and xy: "~(x = y \<and> S = {x})"
+  obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
+                  "open_segment x z \<subseteq> rel_interior S"
+proof (cases "x=y")
+  case True
+  with xy have "S \<noteq> {x}"
+    by blast
+  with True show ?thesis
+    by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y)
+next
+  case False
+  then show ?thesis
+    using segment_to_rel_frontier_aux [OF S x y] that by blast
+qed
+
 proposition rel_frontier_not_sing:
   fixes a :: "'a::euclidean_space"
   assumes "bounded S"
--- a/src/HOL/Analysis/Linear_Algebra.thy	Tue Jul 18 11:35:32 2017 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Wed Jul 19 16:41:26 2017 +0100
@@ -3140,6 +3140,44 @@
 definition collinear :: "'a::real_vector set \<Rightarrow> bool"
   where "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)"
 
+lemma collinear_alt:
+     "collinear S \<longleftrightarrow> (\<exists>u v. \<forall>x \<in> S. \<exists>c. x = u + c *\<^sub>R v)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel)
+next
+  assume ?rhs
+  then obtain u v where *: "\<And>x. x \<in> S \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v"
+    by (auto simp: )
+  have "\<exists>c. x - y = c *\<^sub>R v" if "x \<in> S" "y \<in> S" for x y
+        by (metis *[OF \<open>x \<in> S\<close>] *[OF \<open>y \<in> S\<close>] scaleR_left.diff add_diff_cancel_left)
+  then show ?lhs
+    using collinear_def by blast
+qed
+
+lemma collinear:
+  fixes S :: "'a::{perfect_space,real_vector} set"
+  shows "collinear S \<longleftrightarrow> (\<exists>u. u \<noteq> 0 \<and> (\<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u))"
+proof -
+  have "\<exists>v. v \<noteq> 0 \<and> (\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R v)"
+    if "\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R u" "u=0" for u
+  proof -
+    have "\<forall>x\<in>S. \<forall>y\<in>S. x = y"
+      using that by auto
+    moreover
+    obtain v::'a where "v \<noteq> 0"
+      using UNIV_not_singleton [of 0] by auto
+    ultimately have "\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R v"
+      by auto
+    then show ?thesis
+      using \<open>v \<noteq> 0\<close> by blast
+  qed
+  then show ?thesis
+    apply (clarsimp simp: collinear_def)
+    by (metis real_vector.scale_zero_right vector_fraction_eq_iff)
+qed
+
 lemma collinear_subset: "\<lbrakk>collinear T; S \<subseteq> T\<rbrakk> \<Longrightarrow> collinear S"
   by (meson collinear_def subsetCE)
 
--- a/src/HOL/Analysis/Polytope.thy	Tue Jul 18 11:35:32 2017 +0200
+++ b/src/HOL/Analysis/Polytope.thy	Wed Jul 19 16:41:26 2017 +0100
@@ -74,9 +74,9 @@
   unfolding face_of_def by blast
 
 lemma face_of_imp_eq_affine_Int:
-     fixes S :: "'a::euclidean_space set"
-     assumes S: "convex S" "closed S" and T: "T face_of S"
-     shows "T = (affine hull T) \<inter> S"
+  fixes S :: "'a::euclidean_space set"
+  assumes S: "convex S"  and T: "T face_of S"
+  shows "T = (affine hull T) \<inter> S"
 proof -
   have "convex T" using T by (simp add: face_of_def)
   have *: False if x: "x \<in> affine hull T" and "x \<in> S" "x \<notin> T" and y: "y \<in> rel_interior T" for x y
@@ -269,6 +269,20 @@
     by simp
 qed
 
+lemma subset_of_face_of_affine_hull:
+    fixes S :: "'a::euclidean_space set"
+  assumes T: "T face_of S" and "convex S" "U \<subseteq> S" and dis: "~disjnt (affine hull T) (rel_interior U)"
+  shows "U \<subseteq> T"
+  apply (rule subset_of_face_of [OF T \<open>U \<subseteq> S\<close>])
+  using face_of_imp_eq_affine_Int [OF \<open>convex S\<close> T]
+  using rel_interior_subset [of U] dis
+  using \<open>U \<subseteq> S\<close> disjnt_def by fastforce
+
+lemma affine_hull_face_of_disjoint_rel_interior:
+    fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "F face_of S" "F \<noteq> S"
+  shows "affine hull F \<inter> rel_interior S = {}"
+  by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull)
 
 lemma affine_diff_divide:
     assumes "affine S" "k \<noteq> 0" "k \<noteq> 1" and xy: "x \<in> S" "y /\<^sub>R (1 - k) \<in> S"
@@ -2160,9 +2174,8 @@
 
 lemma face_of_polyhedron_polyhedron:
   fixes S :: "'a :: euclidean_space set"
-  assumes "polyhedron S" "c face_of S"
-    shows "polyhedron c"
-by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_closed polyhedron_imp_convex)
+  assumes "polyhedron S" "c face_of S" shows "polyhedron c"
+by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex)
 
 lemma finite_polyhedron_faces:
   fixes S :: "'a :: euclidean_space set"
@@ -2778,4 +2791,130 @@
   qed (auto simp: eq poly aff face \<open>finite \<F>'\<close>)
 qed
 
+
+subsection\<open>Simplicial complexes and triangulations\<close>
+
+subsubsection\<open>The notion of n-simplex for integer @{term"n \<ge> -1"}\<close>
+
+definition simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
+  where "n simplex S \<equiv> \<exists>C. ~(affine_dependent C) \<and> int(card C) = n + 1 \<and> S = convex hull C"
+
+lemma simplex:
+    "n simplex S \<longleftrightarrow> (\<exists>C. finite C \<and>
+                       ~(affine_dependent C) \<and>
+                       int(card C) = n + 1 \<and>
+                       S = convex hull C)"
+  by (auto simp add: simplex_def intro: aff_independent_finite)
+
+lemma simplex_convex_hull:
+   "~affine_dependent C \<and> int(card C) = n + 1 \<Longrightarrow> n simplex (convex hull C)"
+  by (auto simp add: simplex_def)
+
+lemma convex_simplex: "n simplex S \<Longrightarrow> convex S"
+  by (metis convex_convex_hull simplex_def)
+
+lemma compact_simplex: "n simplex S \<Longrightarrow> compact S"
+  unfolding simplex
+  using finite_imp_compact_convex_hull by blast
+
+lemma closed_simplex: "n simplex S \<Longrightarrow> closed S"
+  by (simp add: compact_imp_closed compact_simplex)
+
+lemma simplex_imp_polytope:
+   "n simplex S \<Longrightarrow> polytope S"
+  unfolding simplex_def polytope_def
+  using aff_independent_finite by blast
+
+lemma simplex_imp_polyhedron:
+   "n simplex S \<Longrightarrow> polyhedron S"
+  by (simp add: polytope_imp_polyhedron simplex_imp_polytope)
+
+lemma simplex_dim_ge: "n simplex S \<Longrightarrow> -1 \<le> n"
+  by (metis (no_types, hide_lams) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def)
+
+lemma simplex_empty [simp]: "n simplex {} \<longleftrightarrow> n = -1"
+proof
+  assume "n simplex {}"
+  then show "n = -1"
+    unfolding simplex by (metis card_empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0)
+next
+  assume "n = -1" then show "n simplex {}"
+    by (fastforce simp: simplex)
+qed
+
+lemma simplex_minus_1 [simp]: "-1 simplex S \<longleftrightarrow> S = {}"
+  by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty)
+
+
+lemma aff_dim_simplex:
+   "n simplex S \<Longrightarrow> aff_dim S = n"
+  by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card)
+
+lemma zero_simplex_sing: "0 simplex {a}"
+  apply (simp add: simplex_def)
+  by (metis affine_independent_1 card_empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI)
+
+lemma simplex_sing [simp]: "n simplex {a} \<longleftrightarrow> n = 0"
+  using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast
+
+lemma simplex_zero: "0 simplex S \<longleftrightarrow> (\<exists>a. S = {a})"
+apply (auto simp: )
+  using aff_dim_eq_0 aff_dim_simplex by blast
+
+lemma one_simplex_segment: "a \<noteq> b \<Longrightarrow> 1 simplex closed_segment a b"
+  apply (simp add: simplex_def)
+  apply (rule_tac x="{a,b}" in exI)
+  apply (auto simp: segment_convex_hull)
+  done
+
+lemma simplex_segment_cases:
+   "(if a = b then 0 else 1) simplex closed_segment a b"
+  by (auto simp: one_simplex_segment)
+
+lemma simplex_segment:
+   "\<exists>n. n simplex closed_segment a b"
+  using simplex_segment_cases by metis
+
+lemma polytope_lowdim_imp_simplex:
+  assumes "polytope P" "aff_dim P \<le> 1"
+  obtains n where "n simplex P"
+proof (cases "P = {}")
+  case True
+  then show ?thesis
+    by (simp add: that)
+next
+  case False
+  then show ?thesis
+    by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that)
+qed
+
+lemma simplex_insert_dimplus1:
+  fixes n::int
+  assumes "n simplex S" and a: "a \<notin> affine hull S"
+  shows "(n+1) simplex (convex hull (insert a S))"
+proof -
+  obtain C where C: "finite C" "~(affine_dependent C)" "int(card C) = n+1" and S: "S = convex hull C"
+    using assms unfolding simplex by force
+  show ?thesis
+    unfolding simplex
+  proof (intro exI conjI)
+      have "aff_dim S = n"
+        using aff_dim_simplex assms(1) by blast
+      moreover have "a \<notin> affine hull C"
+        using S a affine_hull_convex_hull by blast
+      moreover have "a \<notin> C"
+          using S a hull_inc by fastforce
+      ultimately show "\<not> affine_dependent (insert a C)"
+        by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card)
+  next
+    have "a \<notin> C"
+      using S a hull_inc by fastforce
+    then show "int (card (insert a C)) = n + 1 + 1"
+      by (simp add: C)
+  next
+    show "convex hull insert a S = convex hull (insert a C)"
+      by (simp add: S convex_hull_insert_segments)
+  qed (use C in auto)
+qed
+
 end