products/sources/formale sprachen/Isabelle/HOL/Analysis image not shown  

Quellcode-Bibliothek

© Kompilation durch diese Firma

[Weder Korrektheit noch Funktionsfähigkeit der Software werden zugesichert.]

Datei:   Sprache: Isabelle

Original von: Isabelle©

(* Title:      HOL/Analysis/Convex_Euclidean_Space.thy
   Author:     L C Paulson, University of Cambridge
   Author:     Robert Himmelmann, TU Muenchen
   Author:     Bogdan Grechuk, University of Edinburgh
   Author:     Armin Heller, TU Muenchen
   Author:     Johannes Hoelzl, TU Muenchen
*)


section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close>

theory Convex_Euclidean_Space
imports
  Convex
  Topology_Euclidean_Space
begin

subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>

lemma aff_dim_cball:
  fixes a :: "'n::euclidean_space"
  assumes "e > 0"
  shows "aff_dim (cball a e) = int (DIM('n))"
proof -
  have "(\x. a + x) ` (cball 0 e) \ cball a e"
    unfolding cball_def dist_norm by auto
  then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \ aff_dim (cball a e)"
    using aff_dim_translation_eq[of a "cball 0 e"]
          aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"]
    by auto
  moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
    using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"]
      centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms
    by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
  ultimately show ?thesis
    using aff_dim_le_DIM[of "cball a e"by auto
qed

lemma aff_dim_open:
  fixes S :: "'n::euclidean_space set"
  assumes "open S"
    and "S \ {}"
  shows "aff_dim S = int (DIM('n))"
proof -
  obtain x where "x \ S"
    using assms by auto
  then obtain e where e: "e > 0" "cball x e \ S"
    using open_contains_cball[of S] assms by auto
  then have "aff_dim (cball x e) \ aff_dim S"
    using aff_dim_subset by auto
  with e show ?thesis
    using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto
qed

lemma low_dim_interior:
  fixes S :: "'n::euclidean_space set"
  assumes "\ aff_dim S = int (DIM('n))"
  shows "interior S = {}"
proof -
  have "aff_dim(interior S) \ aff_dim S"
    using interior_subset aff_dim_subset[of "interior S" S] by auto
  then show ?thesis
    using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto
qed

corollary empty_interior_lowdim:
  fixes S :: "'n::euclidean_space set"
  shows "dim S < DIM ('n) \ interior S = {}"
by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV)

corollary aff_dim_nonempty_interior:
  fixes S :: "'a::euclidean_space set"
  shows "interior S \ {} \ aff_dim S = DIM('a)"
by (metis low_dim_interior)


subsection \<open>Relative interior of a set\<close>

definition\<^marker>\<open>tag important\<close> "rel_interior S =
  {x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"

lemma rel_interior_mono:
   "\S \ T; affine hull S = affine hull T\
   \<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)"
  by (auto simp: rel_interior_def)

lemma rel_interior_maximal:
   "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)"
  by (auto simp: rel_interior_def)

lemma rel_interior: "rel_interior S = {x \ S. \T. open T \ x \ T \ T \ affine hull S \ S}"
       (is "?lhs = ?rhs")
proof
  show "?lhs \ ?rhs"
    by (force simp add: rel_interior_def openin_open)
  { fix x T
    assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S"
    then have **: "x \ T \ affine hull S"
      using hull_inc by auto
    with * have "\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S"
      by (rule_tac x = "T \ (affine hull S)" in exI) auto
  }
  then show "?rhs \ ?lhs"
    by (force simp add: rel_interior_def openin_open)
qed

lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)"
  by (auto simp: rel_interior)

lemma mem_rel_interior_ball:
  "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)"
  (is "?lhs = ?rhs")
proof
  assume ?rhs then show ?lhs
  by (simp add: rel_interior) (meson Elementary_Metric_Spaces.open_ball centre_in_ball)
qed (force simp: rel_interior open_contains_ball)

lemma rel_interior_ball:
  "rel_interior S = {x \ S. \e. e > 0 \ ball x e \ affine hull S \ S}"
  using mem_rel_interior_ball [of _ S] by auto

lemma mem_rel_interior_cball:
  "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)"
  (is "?lhs = ?rhs")
proof
  assume ?rhs then obtain e where "x \ S" "e > 0" "cball x e \ affine hull S \ S"
    by (auto simp: rel_interior)
  then have "ball x e \ affine hull S \ S"
    by auto
  then show ?lhs
    using \<open>0 < e\<close> \<open>x \<in> S\<close> rel_interior_ball by auto
qed (force simp: rel_interior open_contains_cball)

lemma rel_interior_cball:
  "rel_interior S = {x \ S. \e. e > 0 \ cball x e \ affine hull S \ S}"
  using mem_rel_interior_cball [of _ S] by auto

lemma rel_interior_empty [simp]: "rel_interior {} = {}"
   by (auto simp: rel_interior_def)

lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
  by (metis affine_hull_eq affine_sing)

lemma rel_interior_sing [simp]:
  fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
proof -
  have "\x::real. 0 < x"
    using zero_less_one by blast
  then show ?thesis
    by (auto simp: rel_interior_ball)
qed

lemma subset_rel_interior:
  fixes S T :: "'n::euclidean_space set"
  assumes "S \ T"
    and "affine hull S = affine hull T"
  shows "rel_interior S \ rel_interior T"
  using assms by (auto simp: rel_interior_def)

lemma rel_interior_subset: "rel_interior S \ S"
  by (auto simp: rel_interior_def)

lemma rel_interior_subset_closure: "rel_interior S \ closure S"
  using rel_interior_subset by (auto simp: closure_def)

lemma interior_subset_rel_interior: "interior S \ rel_interior S"
  by (auto simp: rel_interior interior_def)

lemma interior_rel_interior:
  fixes S :: "'n::euclidean_space set"
  assumes "aff_dim S = int(DIM('n))"
  shows "rel_interior S = interior S"
proof -
  have "affine hull S = UNIV"
    using assms affine_hull_UNIV[of S] by auto
  then show ?thesis
    unfolding rel_interior interior_def by auto
qed

lemma rel_interior_interior:
  fixes S :: "'n::euclidean_space set"
  assumes "affine hull S = UNIV"
  shows "rel_interior S = interior S"
  using assms unfolding rel_interior interior_def by auto

lemma rel_interior_open:
  fixes S :: "'n::euclidean_space set"
  assumes "open S"
  shows "rel_interior S = S"
  by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset)

lemma interior_rel_interior_gen:
  fixes S :: "'n::euclidean_space set"
  shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
  by (metis interior_rel_interior low_dim_interior)

lemma rel_interior_nonempty_interior:
  fixes S :: "'n::euclidean_space set"
  shows "interior S \ {} \ rel_interior S = interior S"
by (metis interior_rel_interior_gen)

lemma affine_hull_nonempty_interior:
  fixes S :: "'n::euclidean_space set"
  shows "interior S \ {} \ affine hull S = UNIV"
by (metis affine_hull_UNIV interior_rel_interior_gen)

lemma rel_interior_affine_hull [simp]:
  fixes S :: "'n::euclidean_space set"
  shows "rel_interior (affine hull S) = affine hull S"
proof -
  have *: "rel_interior (affine hull S) \ affine hull S"
    using rel_interior_subset by auto
  {
    fix x
    assume x: "x \ affine hull S"
    define e :: real where "e = 1"
    then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S"
      using hull_hull[of _ S] by auto
    then have "x \ rel_interior (affine hull S)"
      using x rel_interior_ball[of "affine hull S"by auto
  }
  then show ?thesis using * by auto
qed

lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
  by (metis open_UNIV rel_interior_open)

lemma rel_interior_convex_shrink:
  fixes S :: "'a::euclidean_space set"
  assumes "convex S"
    and "c \ rel_interior S"
    and "x \ S"
    and "0 < e"
    and "e \ 1"
  shows "x - e *\<^sub>R (x - c) \ rel_interior S"
proof -
  obtain d where "d > 0" and d: "ball c d \ affine hull S \ S"
    using assms(2) unfolding  mem_rel_interior_ball by auto
  {
    fix y
    assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S"
    have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
      using \<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib)
    have "x \ affine hull S"
      using assms hull_subset[of S] by auto
    moreover have "1 / e + - ((1 - e) / e) = 1"
      using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
    ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ affine hull S"
      using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"]
      by (simp add: algebra_simps)
    have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
      using \<open>e > 0\<close>
      by (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
    then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
      unfolding dist_norm norm_scaleR[symmetric] by auto
    also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)"
      by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
    also have "\ < d"
      using as[unfolded dist_norm] and \<open>e > 0\<close>
      by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
    finally have "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ S"
      using "**"  d by auto
    then have "y \ S"
      using * convexD [OF \<open>convex S\<close>] assms(3-5)
      by (metis diff_add_cancel diff_ge_0_iff_ge le_add_same_cancel1 less_eq_real_def)
  }
  then have "ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S"
    by auto
  moreover have "e * d > 0"
    using \<open>e > 0\<close> \<open>d > 0\<close> by simp
  moreover have c: "c \ S"
    using assms rel_interior_subset by auto
  moreover from c have "x - e *\<^sub>R (x - c) \ S"
    using convexD_alt[of S x c e] assms
    by (metis  diff_add_eq diff_diff_eq2 less_eq_real_def scaleR_diff_left scaleR_one scale_right_diff_distrib)
  ultimately show ?thesis
    using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \e > 0\ by auto
qed

lemma interior_real_atLeast [simp]:
  fixes a :: real
  shows "interior {a..} = {a<..}"
proof -
  {
    fix y
    have "ball y (y - a) \ {a..}"
      by (auto simp: dist_norm)
    moreover assume "a < y"
    ultimately have "y \ interior {a..}"
      by (force simp add: mem_interior)
  }
  moreover
  {
    fix y
    assume "y \ interior {a..}"
    then obtain e where e: "e > 0" "cball y e \ {a..}"
      using mem_interior_cball[of y "{a..}"by auto
    moreover from e have "y - e \ cball y e"
      by (auto simp: cball_def dist_norm)
    ultimately have "a \ y - e" by blast
    then have "a < y" using e by auto
  }
  ultimately show ?thesis by auto
qed

lemma continuous_ge_on_Ioo:
  assumes "continuous_on {c..d} g" "\x. x \ {c<.. g x \ a" "c < d" "x \ {c..d}"
  shows "g (x::real) \ (a::real)"
proof-
  from assms(3) have "{c..d} = closure {c<.. by (rule closure_greaterThanLessThan[symmetric])
  also from assms(2) have "{c<.. (g -` {a..} \ {c..d})" by auto
  hence "closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono)
  also from assms(1) have "closed (g -` {a..} \ {c..d})"
    by (auto simp: continuous_on_closed_vimage)
  hence "closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp
  finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto
qed

lemma interior_real_atMost [simp]:
  fixes a :: real
  shows "interior {..a} = {..
proof -
  {
    fix y
    have "ball y (a - y) \ {..a}"
      by (auto simp: dist_norm)
    moreover assume "a > y"
    ultimately have "y \ interior {..a}"
      by (force simp add: mem_interior)
  }
  moreover
  {
    fix y
    assume "y \ interior {..a}"
    then obtain e where e: "e > 0" "cball y e \ {..a}"
      using mem_interior_cball[of y "{..a}"by auto
    moreover from e have "y + e \ cball y e"
      by (auto simp: cball_def dist_norm)
    ultimately have "a \ y + e" by auto
    then have "a > y" using e by auto
  }
  ultimately show ?thesis by auto
qed

lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<..
proof-
  have "{a..b} = {a..} \ {..b}" by auto
  also have "interior \ = {a<..} \ {..
    by (simp)
  also have "\ = {a<..
  finally show ?thesis .
qed

lemma interior_atLeastLessThan [simp]:
  fixes a::real shows "interior {a..
  by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_atLeast)

lemma interior_lessThanAtMost [simp]:
  fixes a::real shows "interior {a<..b} = {a<..
  by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int
            interior_interior interior_real_atLeast)

lemma interior_greaterThanLessThan_real [simp]: "interior {a<..
  by (metis interior_atLeastAtMost_real interior_interior)

lemma frontier_real_atMost [simp]:
  fixes a :: real
  shows "frontier {..a} = {a}"
  unfolding frontier_def by auto

lemma frontier_real_atLeast [simp]: "frontier {a..} = {a::real}"
  by (auto simp: frontier_def)

lemma frontier_real_greaterThan [simp]: "frontier {a<..} = {a::real}"
  by (auto simp: interior_open frontier_def)

lemma frontier_real_lessThan [simp]: "frontier {..
  by (auto simp: interior_open frontier_def)

lemma rel_interior_real_box [simp]:
  fixes a b :: real
  assumes "a < b"
  shows "rel_interior {a .. b} = {a <..< b}"
proof -
  have "box a b \ {}"
    using assms
    unfolding set_eq_iff
    by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def)
  then show ?thesis
    using interior_rel_interior_gen[of "cbox a b", symmetric]
    by (simp split: if_split_asm del: box_real add: box_real[symmetric])
qed

lemma rel_interior_real_semiline [simp]:
  fixes a :: real
  shows "rel_interior {a..} = {a<..}"
proof -
  have *: "{a<..} \ {}"
    unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"])
  then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"]
    by (auto split: if_split_asm)
qed

subsubsection \<open>Relative open sets\<close>

definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S"

lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" (is "?lhs = ?rhs")
proof
  assume ?lhs
  then show ?rhs
    unfolding rel_open_def rel_interior_def
    using openin_subopen[of "top_of_set (affine hull S)" S] by auto
qed (auto simp:  rel_open_def rel_interior_def)

lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)"
  using openin_subopen by (fastforce simp add: rel_interior_def)

lemma openin_set_rel_interior:
   "openin (top_of_set S) (rel_interior S)"
by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset])

lemma affine_rel_open:
  fixes S :: "'n::euclidean_space set"
  assumes "affine S"
  shows "rel_open S"
  unfolding rel_open_def
  using assms rel_interior_affine_hull[of S] affine_hull_eq[of S]
  by metis

lemma affine_closed:
  fixes S :: "'n::euclidean_space set"
  assumes "affine S"
  shows "closed S"
proof -
  {
    assume "S \ {}"
    then obtain L where L: "subspace L" "affine_parallel S L"
      using assms affine_parallel_subspace[of S] by auto
    then obtain a where a: "S = ((+) a ` L)"
      using affine_parallel_def[of L S] affine_parallel_commut by auto
    from L have "closed L" using closed_subspace by auto
    then have "closed S"
      using closed_translation a by auto
  }
  then show ?thesis by auto
qed

lemma closure_affine_hull:
  fixes S :: "'n::euclidean_space set"
  shows "closure S \ affine hull S"
  by (intro closure_minimal hull_subset affine_closed affine_affine_hull)

lemma closed_affine_hull [iff]:
  fixes S :: "'n::euclidean_space set"
  shows "closed (affine hull S)"
  by (metis affine_affine_hull affine_closed)

lemma closure_same_affine_hull [simp]:
  fixes S :: "'n::euclidean_space set"
  shows "affine hull (closure S) = affine hull S"
proof -
  have "affine hull (closure S) \ affine hull S"
    using hull_mono[of "closure S" "affine hull S" "affine"]
      closure_affine_hull[of S] hull_hull[of "affine" S]
    by auto
  moreover have "affine hull (closure S) \ affine hull S"
    using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
  ultimately show ?thesis by auto
qed

lemma closure_aff_dim [simp]:
  fixes S :: "'n::euclidean_space set"
  shows "aff_dim (closure S) = aff_dim S"
proof -
  have "aff_dim S \ aff_dim (closure S)"
    using aff_dim_subset closure_subset by auto
  moreover have "aff_dim (closure S) \ aff_dim (affine hull S)"
    using aff_dim_subset closure_affine_hull by blast
  moreover have "aff_dim (affine hull S) = aff_dim S"
    using aff_dim_affine_hull by auto
  ultimately show ?thesis by auto
qed

lemma rel_interior_closure_convex_shrink:
  fixes S :: "_::euclidean_space set"
  assumes "convex S"
    and "c \ rel_interior S"
    and "x \ closure S"
    and "e > 0"
    and "e \ 1"
  shows "x - e *\<^sub>R (x - c) \ rel_interior S"
proof -
  obtain d where "d > 0" and d: "ball c d \ affine hull S \ S"
    using assms(2) unfolding mem_rel_interior_ball by auto
  have "\y \ S. norm (y - x) * (1 - e) < e * d"
  proof (cases "x \ S")
    case True
    then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> by force
  next
    case False
    then have x: "x islimpt S"
      using assms(3)[unfolded closure_def] by auto
    show ?thesis
    proof (cases "e = 1")
      case True
      obtain y where "y \ S" "y \ x" "dist y x < 1"
        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
      then show ?thesis
        unfolding True using \<open>d > 0\<close> by (force simp add: )
    next
      case False
      then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
      then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)"
        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
      then show ?thesis
        unfolding dist_norm using pos_less_divide_eq[OF *] by force
    qed
  qed
  then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d"
    by auto
  define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
  have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
    unfolding z_def using \<open>e > 0\<close>
    by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
  have zball: "z \ ball c d"
    using mem_ball z_def dist_norm[of c]
    using y and assms(4,5)
    by (simp add: norm_minus_commute) (simp add: field_simps)
  have "x \ affine hull S"
    using closure_affine_hull assms by auto
  moreover have "y \ affine hull S"
    using \<open>y \<in> S\<close> hull_subset[of S] by auto
  moreover have "c \ affine hull S"
    using assms rel_interior_subset hull_subset[of S] by auto
  ultimately have "z \ affine hull S"
    using z_def affine_affine_hull[of S]
      mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
      assms
    by simp
  then have "z \ S" using d zball by auto
  obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d"
    using zball open_ball[of c d] openE[of "ball c d" z] by auto
  then have "ball z d1 \ affine hull S \ ball c d \ affine hull S"
    by auto
  then have "ball z d1 \ affine hull S \ S"
    using d by auto
  then have "z \ rel_interior S"
    using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto
  then have "y - e *\<^sub>R (y - z) \ rel_interior S"
    using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto
  then show ?thesis using * by auto
qed

lemma rel_interior_eq:
   "rel_interior s = s \ openin(top_of_set (affine hull s)) s"
using rel_open rel_open_def by blast

lemma rel_interior_openin:
   "openin(top_of_set (affine hull s)) s \ rel_interior s = s"
by (simp add: rel_interior_eq)

lemma rel_interior_affine:
  fixes S :: "'n::euclidean_space set"
  shows  "affine S \ rel_interior S = S"
using affine_rel_open rel_open_def by auto

lemma rel_interior_eq_closure:
  fixes S :: "'n::euclidean_space set"
  shows "rel_interior S = closure S \ affine S"
proof (cases "S = {}")
  case True
 then show ?thesis
    by auto
next
  case False show ?thesis
  proof
    assume eq: "rel_interior S = closure S"
    have "openin (top_of_set (affine hull S)) S"
      by (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym)
    moreover have "closedin (top_of_set (affine hull S)) S"
      by (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset)
    ultimately have "S = {} \ S = affine hull S"
      using convex_connected connected_clopen convex_affine_hull by metis
    with False have "affine hull S = S"
      by auto
    then show "affine S"
      by (metis affine_hull_eq)
  next
    assume "affine S"
    then show "rel_interior S = closure S"
      by (simp add: rel_interior_affine affine_closed)
  qed
qed


subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Relative interior preserves under linear transformations\<close>

lemma rel_interior_translation_aux:
  fixes a :: "'n::euclidean_space"
  shows "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)"
proof -
  {
    fix x
    assume x: "x \ rel_interior S"
    then obtain T where "open T" "x \ T \ S" "T \ affine hull S \ S"
      using mem_rel_interior[of x S] by auto
    then have "open ((\x. a + x) ` T)"
      and "a + x \ ((\x. a + x) ` T) \ ((\x. a + x) ` S)"
      and "((\x. a + x) ` T) \ affine hull ((\x. a + x) ` S) \ (\x. a + x) ` S"
      using affine_hull_translation[of a S] open_translation[of T a] x by auto
    then have "a + x \ rel_interior ((\x. a + x) ` S)"
      using mem_rel_interior[of "a+x" "((\x. a + x) ` S)"] by auto
  }
  then show ?thesis by auto
qed

lemma rel_interior_translation:
  fixes a :: "'n::euclidean_space"
  shows "rel_interior ((\x. a + x) ` S) = (\x. a + x) ` rel_interior S"
proof -
  have "(\x. (-a) + x) ` rel_interior ((\x. a + x) ` S) \ rel_interior S"
    using rel_interior_translation_aux[of "-a" "(\x. a + x) ` S"]
      translation_assoc[of "-a" "a"]
    by auto
  then have "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)"
    using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"]
    by auto
  then show ?thesis
    using rel_interior_translation_aux[of a S] by auto
qed


lemma affine_hull_linear_image:
  assumes "bounded_linear f"
  shows "f ` (affine hull s) = affine hull f ` s"
proof -
  interpret f: bounded_linear f by fact
  have "affine {x. f x \ affine hull f ` s}"
    unfolding affine_def
    by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
  moreover have "affine {x. x \ f ` (affine hull s)}"
    using affine_affine_hull[unfolded affine_def, of s]
    unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric])
  ultimately show ?thesis
    by (auto simp: hull_inc elim!: hull_induct)
qed 


lemma rel_interior_injective_on_span_linear_image:
  fixes f :: "'m::euclidean_space \ 'n::euclidean_space"
    and S :: "'m::euclidean_space set"
  assumes "bounded_linear f"
    and "inj_on f (span S)"
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
proof -
  {
    fix z
    assume z: "z \ rel_interior (f ` S)"
    then have "z \ f ` S"
      using rel_interior_subset[of "f ` S"by auto
    then obtain x where x: "x \ S" "f x = z" by auto
    obtain e2 where e2: "e2 > 0" "cball z e2 \ affine hull (f ` S) \ (f ` S)"
      using z rel_interior_cball[of "f ` S"by auto
    obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K"
     using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
    define e1 where "e1 = 1 / K"
    then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x"
      using K pos_le_divide_eq[of e1] by auto
    define e where "e = e1 * e2"
    then have "e > 0" using e1 e2 by auto
    {
      fix y
      assume y: "y \ cball x e \ affine hull S"
      then have h1: "f y \ affine hull (f ` S)"
        using affine_hull_linear_image[of f S] assms by auto
      from y have "norm (x-y) \ e1 * e2"
        using cball_def[of x e] dist_norm[of x y] e_def by auto
      moreover have "f x - f y = f (x - y)"
        using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto
      moreover have "e1 * norm (f (x-y)) \ norm (x - y)"
        using e1 by auto
      ultimately have "e1 * norm ((f x)-(f y)) \ e1 * e2"
        by auto
      then have "f y \ cball z e2"
        using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto
      then have "f y \ f ` S"
        using y e2 h1 by auto
      then have "y \ S"
        using assms y hull_subset[of S] affine_hull_subset_span
          inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>]
        by (metis Int_iff span_superset subsetCE)
    }
    then have "z \ f ` (rel_interior S)"
      using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto
  }
  moreover
  {
    fix x
    assume x: "x \ rel_interior S"
    then obtain e2 where e2: "e2 > 0" "cball x e2 \ affine hull S \ S"
      using rel_interior_cball[of S] by auto
    have "x \ S" using x rel_interior_subset by auto
    then have *: "f x \ f ` S" by auto
    have "\x\span S. f x = 0 \ x = 0"
      using assms subspace_span linear_conv_bounded_linear[of f]
        linear_injective_on_subspace_0[of f "span S"]
      by auto
    then obtain e1 where e1: "e1 > 0" "\x \ span S. e1 * norm x \ norm (f x)"
      using assms injective_imp_isometric[of "span S" f]
        subspace_span[of S] closed_subspace[of "span S"]
      by auto
    define e where "e = e1 * e2"
    hence "e > 0" using e1 e2 by auto
    {
      fix y
      assume y: "y \ cball (f x) e \ affine hull (f ` S)"
      then have "y \ f ` (affine hull S)"
        using affine_hull_linear_image[of f S] assms by auto
      then obtain xy where xy: "xy \ affine hull S" "f xy = y" by auto
      with y have "norm (f x - f xy) \ e1 * e2"
        using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto
      moreover have "f x - f xy = f (x - xy)"
        using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto
      moreover have *: "x - xy \ span S"
        using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy
          affine_hull_subset_span[of S] span_superset
        by auto
      moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))"
        using e1 by auto
      ultimately have "e1 * norm (x - xy) \ e1 * e2"
        by auto
      then have "xy \ cball x e2"
        using cball_def[of x e2] dist_norm[of x xy] e1 by auto
      then have "y \ f ` S"
        using xy e2 by auto
    }
    then have "f x \ rel_interior (f ` S)"
      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto
  }
  ultimately show ?thesis by auto
qed

lemma rel_interior_injective_linear_image:
  fixes f :: "'m::euclidean_space \ 'n::euclidean_space"
  assumes "bounded_linear f"
    and "inj f"
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
  using assms rel_interior_injective_on_span_linear_image[of f S]
    subset_inj_on[of f "UNIV" "span S"]
  by auto


subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness and compactness are preserved by convex hull operation\<close>

lemma open_convex_hull[intro]:
  fixes S :: "'a::real_normed_vector set"
  assumes "open S"
  shows "open (convex hull S)"
proof (clarsimp simp: open_contains_cball convex_hull_explicit)
  fix T and u :: "'a\real"
  assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1"

  from assms[unfolded open_contains_cball] obtain b
    where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis
  have "b ` T \ {}"
    using obt by auto
  define i where "i = b ` T"
  let ?\<Phi> = "\<lambda>y. \<exists>F. finite F \<and> F \<subseteq> S \<and> (\<exists>u. (\<forall>x\<in>F. 0 \<le> u x) \<and> sum u F = 1 \<and> (\<Sum>v\<in>F. u v *\<^sub>R v) = y)"
  let ?a = "\v\T. u v *\<^sub>R v"
  show "\e > 0. cball ?a e \ {y. ?\ y}"
  proof (intro exI subsetI conjI)
    show "0 < Min i"
      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>]
      using b \<open>T\<subseteq>S\<close> by auto
  next
    fix y
    assume "y \ cball ?a (Min i)"
    then have y: "norm (?a - y) \ Min i"
      unfolding dist_norm[symmetric] by auto
    { fix x
      assume "x \ T"
      then have "Min i \ b x"
        by (simp add: i_def obt(1))
      then have "x + (y - ?a) \ cball x (b x)"
        using y unfolding mem_cball dist_norm by auto
      moreover have "x \ S"
        using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto
      ultimately have "x + (y - ?a) \ S"
        using y b by blast
    }
    moreover
    have *: "inj_on (\v. v + (y - ?a)) T"
      unfolding inj_on_def by auto
    have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y"
      unfolding sum.reindex[OF *] o_def using obt(4)
      by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib)
    ultimately show "y \ {y. ?\ y}"
    proof (intro CollectI exI conjI)
      show "finite ((\v. v + (y - ?a)) ` T)"
        by (simp add: obt(1))
      show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1"
        unfolding sum.reindex[OF *] o_def using obt(4) by auto
    qed (use obt(1, 3) in auto)
  qed
qed

lemma compact_convex_combinations:
  fixes S T :: "'a::real_normed_vector set"
  assumes "compact S" "compact T"
  shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}"
proof -
  let ?X = "{0..1} \ S \ T"
  let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
  have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X"
    by force
  have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
  with assms show ?thesis
    by (simp add: * compact_Times compact_continuous_image)
qed

lemma finite_imp_compact_convex_hull:
  fixes S :: "'a::real_normed_vector set"
  assumes "finite S"
  shows "compact (convex hull S)"
proof (cases "S = {}")
  case True
  then show ?thesis by simp
next
  case False
  with assms show ?thesis
  proof (induct rule: finite_ne_induct)
    case (singleton x)
    show ?case by simp
  next
    case (insert x A)
    let ?f = "\(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y"
    let ?T = "{0..1::real} \ (convex hull A)"
    have "continuous_on ?T ?f"
      unfolding split_def continuous_on by (intro ballI tendsto_intros)
    moreover have "compact ?T"
      by (intro compact_Times compact_Icc insert)
    ultimately have "compact (?f ` ?T)"
      by (rule compact_continuous_image)
    also have "?f ` ?T = convex hull (insert x A)"
      unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
      apply safe
      apply (rule_tac x=a in exI, simp)
      apply (rule_tac x="1 - a" in exI, simp, fast)
      apply (rule_tac x="(u, b)" in image_eqI, simp_all)
      done
    finally show "compact (convex hull (insert x A))" .
  qed
qed

lemma compact_convex_hull:
  fixes S :: "'a::euclidean_space set"
  assumes "compact S"
  shows "compact (convex hull S)"
proof (cases "S = {}")
  case True
  then show ?thesis using compact_empty by simp
next
  case False
  then obtain w where "w \ S" by auto
  show ?thesis
    unfolding caratheodory[of S]
  proof (induct ("DIM('a) + 1"))
    case 0
    have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}"
      using compact_empty by auto
    from 0 show ?case unfolding * by simp
  next
    case (Suc n)
    show ?case
    proof (cases "n = 0")
      case True
      have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S"
        unfolding set_eq_iff and mem_Collect_eq
      proof (rule, rule)
        fix x
        assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T"
        then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T"
          by auto
        show "x \ S"
        proof (cases "card T = 0")
          case True
          then show ?thesis
            using T(4) unfolding card_0_eq[OF T(1)] by simp
        next
          case False
          then have "card T = Suc 0" using T(3) \<open>n=0\<close> by auto
          then obtain a where "T = {a}" unfolding card_Suc_eq by auto
          then show ?thesis using T(2,4) by simp
        qed
      next
        fix x assume "x\S"
        then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T"
          by (rule_tac x="{x}" in exI) (use convex_hull_singleton in auto)
      qed
      then show ?thesis using assms by simp
    next
      case False
      have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} =
        {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
          0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> {x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> x \<in> convex hull T}}"
        unfolding set_eq_iff and mem_Collect_eq
      proof (rule, rule)
        fix x
        assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \
          0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)"
        then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
          "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T"
          by auto
        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T"
          by (meson convexD_alt convex_convex_hull hull_inc hull_mono in_mono insertCI obt(2) obt(7) subset_insertI)
        ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T"
          by (rule_tac x="insert u T" in exI) (auto simp: card_insert_if)
      next
        fix x
        assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T"
        then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T"
          by auto
        show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \
          0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)"
        proof (cases "card T = Suc n")
          case False
          then have "card T \ n" using T(3) by auto
          then show ?thesis
            using \<open>w\<in>S\<close> and T
            by (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) auto
        next
          case True
          then obtain a u where au: "T = insert a u" "a\u"
            by (metis card_le_Suc_iff order_refl)
          show ?thesis
          proof (cases "u = {}")
            case True
            then have "x = a" using T(4)[unfolded au] by auto
            show ?thesis unfolding \<open>x = a\<close>
              using T \<open>n \<noteq> 0\<close> unfolding au              
              by (rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) force
          next
            case False
            obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1"
              "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
              using T(4)[unfolded au convex_hull_insert[OF False]]
              by auto
            have *: "1 - vx = ux" using obt(3) by auto
            show ?thesis
              using obt T(1-3) card_insert_disjoint[OF _ au(2)] unfolding au *  
              by (rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI) force
          qed
        qed
      qed
      then show ?thesis
        using compact_convex_combinations[OF assms Suc] by simp
    qed
  qed
qed


subsection\<^marker>\<open>tag unimportant\<close> \<open>Extremal points of a simplex are some vertices\<close>

lemma dist_increases_online:
  fixes a b d :: "'a::real_inner"
  assumes "d \ 0"
  shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b"
proof (cases "inner a d - inner b d > 0")
  case True
  then have "0 < inner d d + (inner a d * 2 - inner b d * 2)"
    using assms
    by (intro add_pos_pos) auto
  then show ?thesis
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
    by (simp add: algebra_simps inner_commute)
next
  case False
  then have "0 < inner d d + (inner b d * 2 - inner a d * 2)"
    using assms
    by (intro add_pos_nonneg) auto
  then show ?thesis
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
    by (simp add: algebra_simps inner_commute)
qed

lemma norm_increases_online:
  fixes d :: "'a::real_inner"
  shows "d \ 0 \ norm (a + d) > norm a \ norm(a - d) > norm a"
  using dist_increases_online[of d a 0] unfolding dist_norm by auto

lemma simplex_furthest_lt:
  fixes S :: "'a::real_inner set"
  assumes "finite S"
  shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))"
  using assms
proof induct
  fix x S
  assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))"
  show "\xa\convex hull insert x S. xa \ insert x S \
    (\<exists>y\<in>convex hull insert x S. norm (xa - a) < norm (y - a))"
  proof (intro impI ballI, cases "S = {}")
    case False
    fix y
    assume y: "y \ convex hull insert x S" "y \ insert x S"
    obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b"
      using y(1)[unfolded convex_hull_insert[OF False]] by auto
    show "\z\convex hull insert x S. norm (y - a) < norm (z - a)"
    proof (cases "y \ convex hull S")
      case True
      then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)"
        using as(3)[THEN bspec[where x=y]] and y(2) by auto
      then show ?thesis
        by (meson hull_mono subsetD subset_insertI)
    next
      case False
      show ?thesis
      proof (cases "u = 0 \ v = 0")
        case True
        with False show ?thesis
          using obt y by auto
      next
        case False
        then obtain w where w: "w>0" "w "w
          using field_lbound_gt_zero[of u v] and obt(1,2) by auto
        have "x \ b"
        proof
          assume "x = b"
          then have "y = b" unfolding obt(5)
            using obt(3) by (auto simp: scaleR_left_distrib[symmetric])
          then show False using obt(4) and False
            using \<open>x = b\<close> y(2) by blast
        qed
        then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto
        show ?thesis
          using dist_increases_online[OF *, of a y]
        proof (elim disjE)
          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
          then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
            unfolding dist_commute[of a]
            unfolding dist_norm obt(5)
            by (simp add: algebra_simps)
          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S"
            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
          proof (intro CollectI conjI exI)
            show "u + w \ 0" "v - w \ 0"
              using obt(1) w by auto
          qed (use obt in auto)
          ultimately show ?thesis by auto
        next
          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
          then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
            unfolding dist_commute[of a]
            unfolding dist_norm obt(5)
            by (simp add: algebra_simps)
          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S"
            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
          proof (intro CollectI conjI exI)
            show "u - w \ 0" "v + w \ 0"
              using obt(1) w by auto
          qed (use obt in auto)
          ultimately show ?thesis by auto
        qed
      qed
    qed
  qed auto
qed (auto simp: assms)

lemma simplex_furthest_le:
  fixes S :: "'a::real_inner set"
  assumes "finite S"
    and "S \ {}"
  shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)"
proof -
  have "convex hull S \ {}"
    using hull_subset[of S convex] and assms(2) by auto
  then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)"
    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \<open>finite S\<close>], of a]
    unfolding dist_commute[of a]
    unfolding dist_norm
    by auto
  show ?thesis
  proof (cases "x \ S")
    case False
    then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)"
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1)
      by auto
    then show ?thesis
      using x(2)[THEN bspec[where x=y]] by auto
  next
    case True
    with x show ?thesis by auto
  qed
qed

lemma simplex_furthest_le_exists:
  fixes S :: "('a::real_inner) set"
  shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)"
  using simplex_furthest_le[of S] by (cases "S = {}") auto

lemma simplex_extremal_le:
  fixes S :: "'a::real_inner set"
  assumes "finite S"
    and "S \ {}"
  shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)"
proof -
  have "convex hull S \ {}"
    using hull_subset[of S convex] and assms(2) by auto
  then obtain u v where obt: "u \ convex hull S" "v \ convex hull S"
    "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)"
    using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]]
    by (auto simp: dist_norm)
  then show ?thesis
  proof (cases "u\S \ v\S", elim disjE)
    assume "u \ S"
    then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)"
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1)
      by auto
    then show ?thesis
      using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2)
      by auto
  next
    assume "v \ S"
    then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)"
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2)
      by auto
    then show ?thesis
      using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
      by (auto simp: norm_minus_commute)
  qed auto
qed

lemma simplex_extremal_le_exists:
  fixes S :: "'a::real_inner set"
  shows "finite S \ x \ convex hull S \ y \ convex hull S \
    \<exists>u\<in>S. \<exists>v\<in>S. norm (x - y) \<le> norm (u - v)"
  using convex_hull_empty simplex_extremal_le[of S]
  by(cases "S = {}") auto


subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close>

definition\<^marker>\<open>tag important\<close> closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
  where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))"

lemma closest_point_exists:
  assumes "closed S"
    and "S \ {}"
  shows closest_point_in_set: "closest_point S a \ S"
    and "\y\S. dist a (closest_point S a) \ dist a y"
  unfolding closest_point_def
  by (rule_tac someI2_ex, auto intro: distance_attains_inf[OF assms(1,2), of a])+

lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x"
  using closest_point_exists[of S] by auto

lemma closest_point_self:
  assumes "x \ S"
  shows "closest_point S x = x"
  unfolding closest_point_def
  by (rule some1_equality, rule ex1I[of _ x]) (use assms in auto)

lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S"
  using closest_point_in_set[of S x] closest_point_self[of x S]
  by auto

lemma closer_points_lemma:
  assumes "inner y z > 0"
  shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y"
proof -
  have z: "inner z z > 0"
    unfolding inner_gt_zero_iff using assms by auto
  have "norm (v *\<^sub>R z - y) < norm y"
    if "0 < v" and "v \ inner y z / inner z z" for v
    unfolding norm_lt using z assms that
    by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
  then show ?thesis
    using assms z
    by (rule_tac x = "inner y z / inner z z" in exI) auto
qed

lemma closer_point_lemma:
  assumes "inner (y - x) (z - x) > 0"
  shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y"
proof -
  obtain u where "u > 0"
    and u: "\v. \0 u\ \ norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
    using closer_points_lemma[OF assms] by auto
  show ?thesis
    using u[of "min u 1"and \<open>u > 0\<close>
    by (metis diff_diff_add dist_commute dist_norm less_eq_real_def not_less u zero_less_one)
qed

lemma any_closest_point_dot:
  assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z"
  shows "inner (a - x) (y - x) \ 0"
proof (rule ccontr)
  assume "\ ?thesis"
  then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a"
    using closer_point_lemma[of a x y] by auto
  let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y"
  have "?z \ S"
    using convexD_alt[OF assms(1,3,4), of u] using u by auto
  then show False
    using assms(5)[THEN bspec[where x="?z"]] and u(3)
    by (auto simp: dist_commute algebra_simps)
qed

lemma any_closest_point_unique:
  fixes x :: "'a::real_inner"
  assumes "convex S" "closed S" "x \ S" "y \ S"
    "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z"
  shows "x = y"
  using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
  unfolding norm_pths(1) and norm_le_square
  by (auto simp: algebra_simps)

lemma closest_point_unique:
  assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z"
  shows "x = closest_point S a"
  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"]
  using closest_point_exists[OF assms(2)] and assms(3) by auto

lemma closest_point_dot:
  assumes "convex S" "closed S" "x \ S"
  shows "inner (a - closest_point S a) (x - closest_point S a) \ 0"
  using any_closest_point_dot[OF assms(1,2) _ assms(3)]
  by (metis assms(2) assms(3) closest_point_in_set closest_point_le empty_iff)

lemma closest_point_lt:
  assumes "convex S" "closed S" "x \ S" "x \ closest_point S a"
  shows "dist a (closest_point S a) < dist a x"
  using closest_point_unique[where a=a] closest_point_le[where a=a] assms by fastforce

lemma setdist_closest_point:
    "\closed S; S \ {}\ \ setdist {a} S = dist a (closest_point S a)"
  by (metis closest_point_exists(2) closest_point_in_set emptyE insert_iff setdist_unique)

lemma closest_point_lipschitz:
  assumes "convex S"
    and "closed S" "S \ {}"
  shows "dist (closest_point S x) (closest_point S y) \ dist x y"
proof -
  have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0"
    and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0"
    by (simp_all add: assms closest_point_dot closest_point_in_set)
  then show ?thesis unfolding dist_norm and norm_le
    using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"]
    by (simp add: inner_add inner_diff inner_commute)
qed

lemma continuous_at_closest_point:
  assumes "convex S"
    and "closed S"
    and "S \ {}"
  shows "continuous (at x) (closest_point S)"
  unfolding continuous_at_eps_delta
  using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto

lemma continuous_on_closest_point:
  assumes "convex S"
    and "closed S"
    and "S \ {}"
  shows "continuous_on t (closest_point S)"
  by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])

proposition closest_point_in_rel_interior:
  assumes "closed S" "S \ {}" and x: "x \ affine hull S"
    shows "closest_point S x \ rel_interior S \ x \ rel_interior S"
proof (cases "x \ S")
  case True
  then show ?thesis
    by (simp add: closest_point_self)
next
  case False
  then have "False" if asm: "closest_point S x \ rel_interior S"
  proof -
    obtain e where "e > 0" and clox: "closest_point S x \ S"
               and e: "cball (closest_point S x) e \ affine hull S \ S"
      using asm mem_rel_interior_cball by blast
    then have clo_notx: "closest_point S x \ x"
      using \<open>x \<notin> S\<close> by auto
    define y where "y \ closest_point S x -
                        (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)"
    have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)"
      by (simp add: y_def algebra_simps)
    then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)"
      by simp
    also have "\ < norm(x - closest_point S x)"
      using clo_notx \<open>e > 0\<close>
      by (auto simp: mult_less_cancel_right2 field_split_simps)
    finally have no_less: "norm (x - y) < norm (x - closest_point S x)" .
    have "y \ affine hull S"
      unfolding y_def
      by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x)
    moreover have "dist (closest_point S x) y \ e"
      using \<open>e > 0\<close> by (auto simp: y_def min_mult_distrib_right)
    ultimately have "y \ S"
      using subsetD [OF e] by simp
    then have "dist x (closest_point S x) \ dist x y"
      by (simp add: closest_point_le \<open>closed S\<close>)
    with no_less show False
      by (simp add: dist_norm)
  qed
  moreover have "x \ rel_interior S"
    using rel_interior_subset False by blast
  ultimately show ?thesis by blast
qed


subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Various point-to-set separating/supporting hyperplane theorems\<close>

lemma supporting_hyperplane_closed_point:
  fixes z :: "'a::{real_inner,heine_borel}"
  assumes "convex S"
    and "closed S"
    and "S \ {}"
    and "z \ S"
  shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)"
proof -
  obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x"
    by (metis distance_attains_inf[OF assms(2-3)])
  show ?thesis
  proof (intro exI bexI conjI ballI)
    show "(y - z) \ z < (y - z) \ y"
      by (metis \<open>y \<in> S\<close> assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq)
    show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x
    proof (rule ccontr)
      have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
        using assms(1)[unfolded convex_alt] and y and \<open>x\<in>S\<close> and \<open>y\<in>S\<close> by auto
      assume "\ (y - z) \ y \ (y - z) \ x"
      then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
        using closer_point_lemma[of z y x] by (auto simp: inner_diff)
      then show False
        using *[of v] by (auto simp: dist_commute algebra_simps)
    qed
  qed (use \<open>y \<in> S\<close> in auto)
qed

lemma separating_hyperplane_closed_point:
  fixes z :: "'a::{real_inner,heine_borel}"
  assumes "convex S"
    and "closed S"
    and "z \ S"
  shows "\a b. inner a z < b \ (\x\S. inner a x > b)"
proof (cases "S = {}")
  case True
  then show ?thesis
    by (simp add: gt_ex)
next
  case False
  obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x"
    by (metis distance_attains_inf[OF assms(2) False])
  show ?thesis
  proof (intro exI conjI ballI)
    show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2"
      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
  next
    fix x
    assume "x \ S"
    have "False" if *: "0 < inner (z - y) (x - y)"
    proof -
      obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
        using * closer_point_lemma by blast
      then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\]
        using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: dist_commute algebra_simps)
    qed
    moreover have "0 < (norm (y - z))\<^sup>2"
      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
    then have "0 < inner (y - z) (y - z)"
      unfolding power2_norm_eq_inner by simp
    ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x"
      by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff)
  qed 
qed

lemma separating_hyperplane_closed_0:
  assumes "convex (S::('a::euclidean_space) set)"
    and "closed S"
    and "0 \ S"
  shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)"
proof (cases "S = {}")
  case True
  have "(SOME i. i\Basis) \ (0::'a)"
    by (metis Basis_zero SOME_Basis)
  then show ?thesis
    using True zero_less_one by blast
next
  case False
  then show ?thesis
    using False using separating_hyperplane_closed_point[OF assms]
    by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le)
qed


subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Now set-to-set for closed/compact sets\<close>

lemma separating_hyperplane_closed_compact:
  fixes S :: "'a::euclidean_space set"
  assumes "convex S"
    and "closed S"
    and "convex T"
    and "compact T"
    and "T \ {}"
    and "S \ T = {}"
  shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)"
proof (cases "S = {}")
  case True
  obtain b where b: "b > 0" "\x\T. norm x \ b"
    using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
  obtain z :: 'a where z: "norm z = b + 1"
    using vector_choose_size[of "b + 1"and b(1) by auto
  then have "z \ T" using b(2)[THEN bspec[where x=z]] by auto
  then obtain a b where ab: "inner a z < b" "\x\T. b < inner a x"
    using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z]
    by auto
  then show ?thesis
    using True by auto
next
  case False
  then obtain y where "y \ S" by auto
  obtain a b where "0 < b" and \<section>: "\<And>x. x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}) \<Longrightarrow> b < inner a x"
    using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
    using closed_compact_differences assms by fastforce 
  have ab: "b + inner a y < inner a x" if "x\S" "y\T" for x y
    using \<section> [of "x-y"] that by (auto simp add: inner_diff_right less_diff_eq)
  define k where "k = (SUP x\T. a \ x)"
  have "k + b / 2 < a \ x" if "x \ S" for x
  proof -
    have "k \ inner a x - b"
      unfolding k_def
      using \<open>T \<noteq> {}\<close> ab that by (fastforce intro: cSUP_least)
    then show ?thesis
      using \<open>0 < b\<close> by auto
  qed
  moreover
  have "- (k + b / 2) < - a \ x" if "x \ T" for x
  proof -
    have "inner a x - b / 2 < k"
      unfolding k_def
    proof (subst less_cSUP_iff)
      show "T \ {}" by fact
      show "bdd_above ((\) a ` T)"
        using ab[rule_format, of y] \<open>y \<in> S\<close>
        by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
      show "\y\T. a \ x - b / 2 < a \ y"
        using \<open>0 < b\<close> that by force
    qed 
    then show ?thesis
      by auto
  qed
  ultimately show ?thesis
    by (metis inner_minus_left neg_less_iff_less)
qed

lemma separating_hyperplane_compact_closed:
  fixes S :: "'a::euclidean_space set"
  assumes "convex S"
    and "compact S"
    and "S \ {}"
    and "convex T"
    and "closed T"
    and "S \ T = {}"
  shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)"
proof -
  obtain a b where "(\x\T. inner a x < b) \ (\x\S. b < inner a x)"
    by (metis disjoint_iff_not_equal separating_hyperplane_closed_compact assms)
  then show ?thesis
    by (metis inner_minus_left neg_less_iff_less)
qed


subsubsection\<^marker>\<open>tag unimportant\<close> \<open>General case without assuming closure and getting non-strict separation\<close>

lemma separating_hyperplane_set_0:
  assumes "convex S" "(0::'a::euclidean_space) \ S"
  shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)"
proof -
  let ?k = "\c. {x::'a. 0 \ inner c x}"
  have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f
  proof -
    obtain c where c: "f = ?k ` c" "c \ S" "finite c"
      using finite_subset_image[OF as(2,1)] by auto
    then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x"
      using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
      using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2)
      using subset_hull[of convex, OF assms(1), symmetric, of c]
      by force
    have "norm (a /\<^sub>R norm a) = 1"
      by (simp add: ab(1))
    moreover have "(\y\c. 0 \ y \ (a /\<^sub>R norm a))"
      using hull_subset[of c convex] ab by (force simp: inner_commute)
    ultimately have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)"
      by blast
    then show "frontier (cball 0 1) \ \f \ {}"
      unfolding c(1) frontier_cball sphere_def dist_norm by auto
  qed
  have "frontier (cball 0 1) \ (\(?k ` S)) \ {}"
    by (rule compact_imp_fip) (use * closed_halfspace_ge in auto)
  then obtain x where "norm x = 1" "\y\S. x\?k y"
    unfolding frontier_cball dist_norm sphere_def by auto
  then show ?thesis
    by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one)
qed

lemma separating_hyperplane_sets:
  fixes S T :: "'a::euclidean_space set"
  assumes "convex S"
    and "convex T"
    and "S \ {}"
    and "T \ {}"
    and "S \ T = {}"
  shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)"
proof -
  from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
  obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x"
    using assms(3-5) by force
  then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x"
    by (force simp: inner_diff)
  then have bdd: "bdd_above (((\) a)`S)"
    using \<open>T \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
  show ?thesis
    using \<open>a\<noteq>0\<close>
    by (intro exI[of _ a] exI[of _ "SUP x\S. a \ x"])
       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *)
qed


subsection\<^marker>\<open>tag unimportant\<close> \<open>More convexity generalities\<close>

lemma convex_closure [intro,simp]:
  fixes S :: "'a::real_normed_vector set"
  assumes "convex S"
  shows "convex (closure S)"
  apply (rule convexI)
  unfolding closure_sequential
  apply (elim exE)
  subgoal for x y u v f g
    by (rule_tac x="\n. u *\<^sub>R f n + v *\<^sub>R g n" in exI) (force intro: tendsto_intros dest: convexD [OF assms])
  done

lemma convex_interior [intro,simp]:
  fixes S :: "'a::real_normed_vector set"
  assumes "convex S"
  shows "convex (interior S)"
  unfolding convex_alt Ball_def mem_interior
proof clarify
  fix x y u
  assume u: "0 \ u" "u \ (1::real)"
  fix e d
  assume ed: "ball x e \ S" "ball y d \ S" "0
  show "\e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S"
  proof (intro exI conjI subsetI)
    fix z
    assume z: "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
    have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S"
    proof (rule_tac assms[unfolded convex_alt, rule_format])
       show "z - u *\<^sub>R (y - x) \ S" "z + (1 - u) *\<^sub>R (y - x) \ S"
         using ed z u by (auto simp add: algebra_simps dist_norm)
    qed (use u in auto)
    then show "z \ S"
      using u by (auto simp: algebra_simps)
  qed(use u ed in auto)
qed

lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}"
  using hull_subset[of S convex] convex_hull_empty by auto


subsection\<^marker>\<open>tag unimportant\<close> \<open>Convex set as intersection of halfspaces\<close>

lemma convex_halfspace_intersection:
  fixes S :: "('a::euclidean_space) set"
  assumes "closed S" "convex S"
  shows "S = \{h. S \ h \ (\a b. h = {x. inner a x \ b})}"
proof -
  { fix z
    assume "\T. S \ T \ (\a b. T = {x. inner a x \ b}) \ z \ T" "z \ S"
    then have \<section>: "\<And>a b. S \<subseteq> {x. inner a x \<le> b} \<Longrightarrow> z \<in> {x. inner a x \<le> b}"
      by blast
    obtain a b where "inner a z < b" "(\x\S. inner a x > b)"
      using \<open>z \<notin> S\<close> assms separating_hyperplane_closed_point by blast
    then have False
--> --------------------

--> maximum size reached

--> --------------------

¤ Dauer der Verarbeitung: 0.68 Sekunden  (vorverarbeitet)  ¤





zum Wurzelverzeichnis wechseln
Diese Quellcodebibliothek enthält Beispiele in vielen Programmiersprachen. Man kann per Verzeichnistruktur darin navigieren. Der Code wird farblich markiert angezeigt.
zum Wurzelverzeichnis wechseln
sprechenden Kalenders

in der Quellcodebibliothek suchen




Laden

Fehler beim Verzeichnis:


in der Quellcodebibliothek suchen

Die farbliche Syntaxdarstellung ist noch experimentell.


Verzeichnis aufwärts